Theory Equivalence_Games

section ‹Reachability Games›

theory Equivalence_Games
  imports Strong_Equivalences
begin

text ‹
  A game is an unlabeled graph where vertices are partitioned into
  defender and attacker positions.
›
locale game =
  fixes
    game_move :: 'g  'g  bool (infix  80) and
    defender_position :: 'g  bool
begin

abbreviation attacker_position g  ¬defender_position g
abbreviation options g  {g'. g  g'}

text ‹Each player loses at a position if it were their turn but they are stuck.›
definition defender_loses g  defender_position g  options g = {}
definition attacker_loses g  attacker_position g  options g = {}

text ‹
  A (positional) strategy is a function to select among the options at a
  position. That only possible moves are valid choices cannot be expressed in
  a HOL type. We express this by soundness predicates for attacker/defender
  strategies.
›
type_synonym ('g0) strategy = 'g0  'g0

definition sound_defender_strategy strat g 
  defender_position g  options g  {}  strat g  options g
definition sound_attacker_strategy strat g 
  attacker_position g  options g  {}  strat g  options g

text ‹
  A play (fragment) is a sequence of game positions that follow a path of game
  moves.
›
inductive play :: 'g list  bool where
  init: play [g] |
  step: play (g # (g' # gg)) if g  g' play (g' # gg)

text ‹We have defined plays in a way where there are no empty plays.›
lemma no_empty_play:
  assumes play []
  shows False
  using assms play.cases by blast

text ‹
  A play follows a defender strategy if every every defender-controlled move
  obeys the strategy. (The type here, does not really ensure the position
  sequences to be plays and the strategies to be sound. This information
  should come from the context.)›
fun play_follows_defender_strategy :: 'g list  ('g  'g)  bool
  where
  play_follows_defender_strategy (g0 # g1 # pl) strat =
    ((if defender_position g0 then strat g0 = g1 else True)
     play_follows_defender_strategy (g1 # pl) strat) |
  play_follows_defender_strategy _ _ = True

lemma play_extension:
  assumes
    last pl  g'
    play pl
  shows
    play (pl @ [g'])
  using assms no_empty_play 
  by (induct pl, auto,
      smt (verit, del_insts) append_Cons append_self_conv2
        list.sel(3) play.simps)

lemma play_follows_defender_strategy_extension_atk:
  assumes
    play_follows_defender_strategy pl strat
    last pl  g'
    attacker_position (last pl)
  shows
    play_follows_defender_strategy (pl @ [g']) strat
  using assms
  by (induct pl, auto, smt (z3) append_self_conv2 hd_append list.distinct(1)
        list.sel(1,3) play_follows_defender_strategy.elims(1))

lemma play_follows_defender_strategy_extension_dfn:
  assumes
    play_follows_defender_strategy pl strat
    defender_position (last pl)
  shows
    play_follows_defender_strategy (pl @ [strat (last pl)]) strat
  using assms
  by (induct pl, force, smt (z3) append_Cons append_Nil
    game.play_follows_defender_strategy.simps(3) last.simps
    list.inject play_follows_defender_strategy.elims(1))

fun play_follows_attacker_strategy :: 'g list  ('g  'g)  bool
  where
  play_follows_attacker_strategy (g0 # g1 # pl) strat =
    ((if attacker_position g0 then strat g0 = g1 else True)
     play_follows_attacker_strategy (g1 # pl) strat)
| play_follows_attacker_strategy _ _ = True

text ‹
  A defender strategy is winning from a position if all plays following the
  strategy from there only lead to positions where the defender has moves
  and the strategy is sound. (In particular, the defender also wins if the
  game goes on forever.)
›
definition defender_winning_strategy :: 'g strategy  'g  bool
  where defender_winning_strategy def_strat g 
  (pl. play pl  hd pl = g  play_follows_defender_strategy pl def_strat
   sound_defender_strategy def_strat (last pl)  ¬defender_loses (last pl))

end

section ‹The Bisimulation Game›

datatype ('a, 's) bisim_game_pos =
  Bisim_Attack 's 's
| Bisim_Defense 'a 's 's

fun (in lts) bisim_game_move ::
    ('a, 's) bisim_game_pos  ('a, 's) bisim_game_pos  bool (infix ↣B 80)
  where
  Bisim_Attack p q ↣B Bisim_Attack p' q' =
    (p' = q  q' = p)
| Bisim_Attack p q ↣B Bisim_Defense α p' q' =
    (p  α p'  q' = q)
| Bisim_Defense α p q ↣B Bisim_Attack p' q' =
    (q  α q'  p' = p)
| _ ↣B _ = False

primrec bisim_defender_position where
  bisim_defender_position (Bisim_Defense _ _ _) = True |
  bisim_defender_position (Bisim_Attack _ _) = False

locale bisim_game =
  lts step +
  game (↣B) bisim_defender_position
  for step :: 's  'a  's  bool (‹_ _ _› [70, 70, 70] 80) 
begin

fun strategy_from_bisim ::
    ('s  's  bool)  ('a, 's) bisim_game_pos strategy where
  strategy_from_bisim R (Bisim_Defense α p' q) =
    Bisim_Attack p' (SOME q'. R p' q'  q  α q')
| strategy_from_bisim _ _ = undefined

lemma bisim_implies_defender_winning_strategy:
  assumes simulation R symp R R p q
  shows defender_winning_strategy (strategy_from_bisim R) (Bisim_Attack p q)
  unfolding defender_winning_strategy_def
proof (clarify)
  fix pl
  assume play pl hd pl = Bisim_Attack p q
    play_follows_defender_strategy pl (strategy_from_bisim R)
  thus sound_defender_strategy (strategy_from_bisim R) (last pl)
       ¬ defender_loses (last pl)
    using R p q
  proof (induct pl arbitrary: p q rule: induct_list012)
    case (1 p q)
    then show ?case using no_empty_play by fastforce
  next
    case (2 g p q)
    then show ?case
      unfolding defender_loses_def sound_defender_strategy_def by simp
  next
    case (3 g g' pl p q)
    then have g = Bisim_Attack p q g ↣B g' using play.cases
      by (auto, fastforce)
    then obtain α p' where
      g'_spec: g' = Bisim_Defense α p' q  p  α p'  g' = Bisim_Attack q p
      by (smt (verit, best) bisim_game_move.elims(1) bisim_game_pos.distinct(1)
          bisim_game_pos.inject(1))
    then show ?case
    proof (cases pl, safe)
      assume g'_def: g' = Bisim_Defense α p' q p α p'
      then obtain q' where q'_spec: q α q' R p' q'
        using simulation R R p q unfolding simulation_def by blast
      then show sound_defender_strategy (strategy_from_bisim R)
                  (last [g, Bisim_Defense α p' q])
        unfolding sound_defender_strategy_def
        by (auto, metis (mono_tags, lifting) someI2_ex)
      show defender_loses (last [g, Bisim_Defense α p' q])  False
        using q'_spec unfolding defender_loses_def
          using bisim_game_move.simps(3) by (auto, blast)
      have play (g' # pl) using play (g # g' # pl) using play.cases by blast
      fix g'' plrt
      assume pl_def: pl = g'' # plrt
      hence play pl using play (g' # pl) play.cases by auto
      have g'' = (strategy_from_bisim R) g' using g'_def(1) pl_def 3(5) by simp
      then obtain q'' where q''_spec:
          g'' = Bisim_Attack p' q'' q α q'' R p' q''
        using q'_spec g'_def by (auto, metis (mono_tags, lifting) someI_ex)
      then show
          sound_defender_strategy (strategy_from_bisim R)
            (last (g # Bisim_Defense α p' q # g'' # plrt))
          defender_loses (last (g # Bisim_Defense α p' q # g'' # plrt))  False
        using 3 play pl unfolding defender_loses_def pl_def by auto
    next
      assume g' = Bisim_Attack q p
      then show
          sound_defender_strategy (strategy_from_bisim R)
            (last [g, Bisim_Attack q p])
          defender_loses (last [g, Bisim_Attack q p])  False
        unfolding sound_defender_strategy_def defender_loses_def by auto
      fix g'' plrt
      assume pl_def: pl = g'' # plrt
      hence play pl using play (g # g' # pl) play.cases by auto
      have R q p using R p q symp R by (meson sympE)
      thus sound_defender_strategy (strategy_from_bisim R)
            (last (g # Bisim_Attack q p # g'' # plrt))
           defender_loses (last (g # Bisim_Attack q p # g'' # plrt))  False
        using 3 play.cases
        unfolding sound_defender_strategy_def defender_loses_def pl_def
          g' = Bisim_Attack q p
        by (auto) blast+
    qed
  qed
qed

lemma defender_winning_strategy_implies_bisim:
  assumes
    defender_winning_strategy strat (Bisim_Attack p0 q0)
  defines
    R == λp q. (pl.
      hd pl = (Bisim_Attack p0 q0)
       play pl
       play_follows_defender_strategy pl strat
       last pl = (Bisim_Attack p q))
  shows
    simulation R symp R R p0 q0
proof -
  show symp R
    unfolding symp_def R_def
    using no_empty_play
    by (metis bisim_defender_position.simps(2) bisim_game_move.simps(1)
      game.play_extension game.play_follows_defender_strategy_extension_atk
      hd_append2 last_snoc)
  show R p0 q0
    using assms(1) unfolding defender_winning_strategy_def R_def
    using play.init by force
  show simulation R
    unfolding simulation_def
  proof safe
    fix p q a p'
    assume p a p' R p q
    then obtain pl where pl_spec:
        hd pl = Bisim_Attack p0 q0
        play_follows_defender_strategy pl strat
        last pl = (Bisim_Attack p q)
        play pl
      unfolding R_def by blast
    from p a p' have Bisim_Attack p q ↣B Bisim_Defense a p' q by simp
    hence defender_extension:
        play_follows_defender_strategy (pl @ [Bisim_Defense a p' q]) strat
        play (pl @ [Bisim_Defense a p' q])
      using play_follows_defender_strategy_extension_atk 
        pl_spec play_extension[of pl Bisim_Defense a p' q] by auto
    hence defender_extension_hd:
      hd (pl @ [Bisim_Defense a p' q]) = Bisim_Attack p0 q0
      by (metis hd_append no_empty_play pl_spec(1,4))
    hence options (Bisim_Defense a p' q)  {}
      using assms(1) pl_spec defender_extension
      unfolding defender_winning_strategy_def defender_loses_def
      by force
    then obtain q' where q'_spec:
        strat (Bisim_Defense a p' q) = Bisim_Attack p' q'
        Bisim_Attack p' q'  options (Bisim_Defense a p' q)
      using defender_extension assms(1) defender_extension_hd
        bisim_defender_position.simps(1) bisim_game_move.simps(3,5)
      unfolding defender_winning_strategy_def sound_defender_strategy_def
      by (metis last_snoc mem_Collect_eq bisim_game_pos.exhaust)
    define long_play where
      long_play  pl @ [Bisim_Defense a p' q, Bisim_Attack p' q']
    hence hd long_play = Bisim_Attack p0 q0  play long_play
         play_follows_defender_strategy long_play strat
         last long_play = Bisim_Attack p' q'
      using play_extension play_follows_defender_strategy_extension_dfn
        defender_extension q'_spec pl_spec no_empty_play last_snoc
      by (auto simp add: hd_append) fastforce+
    thus q'. q a q'  R p' q'
      using q'_spec unfolding long_play_def R_def by force
  qed
qed

theorem bisim_game_characterization:
  shows
    (strat. defender_winning_strategy strat (Bisim_Attack p q)) =
      bisimilar p q  
  using defender_winning_strategy_implies_bisim
    bisim_implies_defender_winning_strategy
  unfolding bisimilar_def
  by blast

end

end