src/HOL/ZF/Games.thy
author haftmann
Mon, 20 Feb 2012 08:01:08 +0100
changeset 46555 c2b5900988e2
parent 45694 4a8743618257
child 46557 ae926869a311
permissions -rw-r--r--
tuned proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41777
1f7cbe39d425 more precise headers;
wenzelm
parents: 41528
diff changeset
     1
(*  Title:      HOL/ZF/Games.thy
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     2
    Author:     Steven Obua
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     3
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
     4
An application of HOLZF: Partizan Games.  See "Partizan Games in
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
     5
Isabelle/HOLZF", available from http://www4.in.tum.de/~obua/partizan
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     6
*)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     7
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     8
theory Games 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
     9
imports MainZF
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    10
begin
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    11
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    12
definition fixgames :: "ZF set \<Rightarrow> ZF set" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    13
  "fixgames A \<equiv> { Opair l r | l r. explode l \<subseteq> A & explode r \<subseteq> A}"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    14
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    15
definition games_lfp :: "ZF set" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    16
  "games_lfp \<equiv> lfp fixgames"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    17
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    18
definition games_gfp :: "ZF set" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    19
  "games_gfp \<equiv> gfp fixgames"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    20
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    21
lemma mono_fixgames: "mono (fixgames)"
46555
c2b5900988e2 tuned proof
haftmann
parents: 45694
diff changeset
    22
  by (auto simp add: mono_def fixgames_def)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    23
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    24
lemma games_lfp_unfold: "games_lfp = fixgames games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    25
  by (auto simp add: def_lfp_unfold games_lfp_def mono_fixgames)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    26
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    27
lemma games_gfp_unfold: "games_gfp = fixgames games_gfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    28
  by (auto simp add: def_gfp_unfold games_gfp_def mono_fixgames)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    29
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    30
lemma games_lfp_nonempty: "Opair Empty Empty \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    31
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    32
  have "fixgames {} \<subseteq> games_lfp" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    33
    apply (subst games_lfp_unfold)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    34
    apply (simp add: mono_fixgames[simplified mono_def, rule_format])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    35
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    36
  moreover have "fixgames {} = {Opair Empty Empty}"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    37
    by (simp add: fixgames_def explode_Empty)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    38
  finally show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    39
    by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    40
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    41
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    42
definition left_option :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    43
  "left_option g opt \<equiv> (Elem opt (Fst g))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    44
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    45
definition right_option :: "ZF \<Rightarrow> ZF \<Rightarrow> bool" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    46
  "right_option g opt \<equiv> (Elem opt (Snd g))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    47
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
    48
definition is_option_of :: "(ZF * ZF) set" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    49
  "is_option_of \<equiv> { (opt, g) | opt g. g \<in> games_gfp \<and> (left_option g opt \<or> right_option g opt) }"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    50
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    51
lemma games_lfp_subset_gfp: "games_lfp \<subseteq> games_gfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    52
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    53
  have "games_lfp \<subseteq> fixgames games_lfp" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    54
    by (simp add: games_lfp_unfold[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    55
  then show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    56
    by (simp add: games_gfp_def gfp_upperbound)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    57
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    58
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    59
lemma games_option_stable: 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    60
  assumes fixgames: "games = fixgames games"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    61
  and g: "g \<in> games"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    62
  and opt: "left_option g opt \<or> right_option g opt"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    63
  shows "opt \<in> games"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    64
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    65
  from g fixgames have "g \<in> fixgames games" by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    66
  then have "\<exists> l r. g = Opair l r \<and> explode l \<subseteq> games \<and> explode r \<subseteq> games"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    67
    by (simp add: fixgames_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    68
  then obtain l where "\<exists> r. g = Opair l r \<and> explode l \<subseteq> games \<and> explode r \<subseteq> games" ..
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    69
  then obtain r where lr: "g = Opair l r \<and> explode l \<subseteq> games \<and> explode r \<subseteq> games" ..
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    70
  with opt show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    71
    by (auto intro: Elem_explode_in simp add: left_option_def right_option_def Fst Snd)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    72
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    73
    
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    74
lemma option2elem: "(opt,g) \<in> is_option_of  \<Longrightarrow> \<exists> u v. Elem opt u \<and> Elem u v \<and> Elem v g"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    75
  apply (simp add: is_option_of_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    76
  apply (subgoal_tac "(g \<in> games_gfp) = (g \<in> (fixgames games_gfp))")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    77
  prefer 2
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    78
  apply (simp add: games_gfp_unfold[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    79
  apply (auto simp add: fixgames_def left_option_def right_option_def Fst Snd)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    80
  apply (rule_tac x=l in exI, insert Elem_Opair_exists, blast)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    81
  apply (rule_tac x=r in exI, insert Elem_Opair_exists, blast) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    82
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    83
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    84
lemma is_option_of_subset_is_Elem_of: "is_option_of \<subseteq> (is_Elem_of^+)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    85
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    86
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    87
    fix opt
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    88
    fix g
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    89
    assume "(opt, g) \<in> is_option_of"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    90
    then have "\<exists> u v. (opt, u) \<in> (is_Elem_of^+) \<and> (u,v) \<in> (is_Elem_of^+) \<and> (v,g) \<in> (is_Elem_of^+)" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    91
      apply -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    92
      apply (drule option2elem)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    93
      apply (auto simp add: r_into_trancl' is_Elem_of_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    94
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    95
    then have "(opt, g) \<in> (is_Elem_of^+)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    96
      by (blast intro: trancl_into_rtrancl trancl_rtrancl_trancl)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    97
  } 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    98
  then show ?thesis by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
    99
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   100
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   101
lemma wfzf_is_option_of: "wfzf is_option_of"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   102
proof - 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   103
  have "wfzf (is_Elem_of^+)" by (simp add: wfzf_trancl wfzf_is_Elem_of)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   104
  then show ?thesis 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   105
    apply (rule wfzf_subset)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   106
    apply (rule is_option_of_subset_is_Elem_of)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   107
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   108
  qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   109
  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   110
lemma games_gfp_imp_lfp: "g \<in> games_gfp \<longrightarrow> g \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   111
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   112
  have unfold_gfp: "\<And> x. x \<in> games_gfp \<Longrightarrow> x \<in> (fixgames games_gfp)" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   113
    by (simp add: games_gfp_unfold[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   114
  have unfold_lfp: "\<And> x. (x \<in> games_lfp) = (x \<in> (fixgames games_lfp))"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   115
    by (simp add: games_lfp_unfold[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   116
  show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   117
    apply (rule wf_induct[OF wfzf_implies_wf[OF wfzf_is_option_of]])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   118
    apply (auto simp add: is_option_of_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   119
    apply (drule_tac unfold_gfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   120
    apply (simp add: fixgames_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   121
    apply (auto simp add: left_option_def Fst right_option_def Snd)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   122
    apply (subgoal_tac "explode l \<subseteq> games_lfp")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   123
    apply (subgoal_tac "explode r \<subseteq> games_lfp")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   124
    apply (subst unfold_lfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   125
    apply (auto simp add: fixgames_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   126
    apply (simp_all add: explode_Elem Elem_explode_in)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   127
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   128
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   129
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   130
theorem games_lfp_eq_gfp: "games_lfp = games_gfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   131
  apply (auto simp add: games_gfp_imp_lfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   132
  apply (insert games_lfp_subset_gfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   133
  apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   134
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   135
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   136
theorem unique_games: "(g = fixgames g) = (g = games_lfp)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   137
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   138
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   139
    fix g 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   140
    assume g: "g = fixgames g"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   141
    from g have "fixgames g \<subseteq> g" by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   142
    then have l:"games_lfp \<subseteq> g" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   143
      by (simp add: games_lfp_def lfp_lowerbound)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   144
    from g have "g \<subseteq> fixgames g" by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   145
    then have u:"g \<subseteq> games_gfp" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   146
      by (simp add: games_gfp_def gfp_upperbound)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   147
    from l u games_lfp_eq_gfp[symmetric] have "g = games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   148
      by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   149
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   150
  note games = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   151
  show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   152
    apply (rule iff[rule_format])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   153
    apply (erule games)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   154
    apply (simp add: games_lfp_unfold[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   155
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   156
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   157
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   158
lemma games_lfp_option_stable: 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   159
  assumes g: "g \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   160
  and opt: "left_option g opt \<or> right_option g opt"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   161
  shows "opt \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   162
  apply (rule games_option_stable[where g=g])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   163
  apply (simp add: games_lfp_unfold[symmetric])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   164
  apply (simp_all add: assms)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   165
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   166
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   167
lemma is_option_of_imp_games:
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   168
  assumes hyp: "(opt, g) \<in> is_option_of"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   169
  shows "opt \<in> games_lfp \<and> g \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   170
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   171
  from hyp have g_game: "g \<in> games_lfp" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   172
    by (simp add: is_option_of_def games_lfp_eq_gfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   173
  from hyp have "left_option g opt \<or> right_option g opt"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   174
    by (auto simp add: is_option_of_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   175
  with g_game games_lfp_option_stable[OF g_game, OF this] show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   176
    by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   177
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   178
 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   179
lemma games_lfp_represent: "x \<in> games_lfp \<Longrightarrow> \<exists> l r. x = Opair l r"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   180
  apply (rule exI[where x="Fst x"])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   181
  apply (rule exI[where x="Snd x"])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   182
  apply (subgoal_tac "x \<in> (fixgames games_lfp)")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   183
  apply (simp add: fixgames_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   184
  apply (auto simp add: Fst Snd)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   185
  apply (simp add: games_lfp_unfold[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   186
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   187
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   188
definition "game = games_lfp"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   189
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   190
typedef (open) game = game
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   191
  unfolding game_def by (blast intro: games_lfp_nonempty)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   192
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   193
definition left_options :: "game \<Rightarrow> game zet" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   194
  "left_options g \<equiv> zimage Abs_game (zexplode (Fst (Rep_game g)))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   195
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   196
definition right_options :: "game \<Rightarrow> game zet" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   197
  "right_options g \<equiv> zimage Abs_game (zexplode (Snd (Rep_game g)))"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   198
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   199
definition options :: "game \<Rightarrow> game zet" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   200
  "options g \<equiv> zunion (left_options g) (right_options g)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   201
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   202
definition Game :: "game zet \<Rightarrow> game zet \<Rightarrow> game" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   203
  "Game L R \<equiv> Abs_game (Opair (zimplode (zimage Rep_game L)) (zimplode (zimage Rep_game R)))"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   204
  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   205
lemma Repl_Rep_game_Abs_game: "\<forall> e. Elem e z \<longrightarrow> e \<in> games_lfp \<Longrightarrow> Repl z (Rep_game o Abs_game) = z"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   206
  apply (subst Ext)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   207
  apply (simp add: Repl)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   208
  apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   209
  apply (subst Abs_game_inverse, simp_all add: game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   210
  apply (rule_tac x=za in exI)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   211
  apply (subst Abs_game_inverse, simp_all add: game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   212
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   213
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   214
lemma game_split: "g = Game (left_options g) (right_options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   215
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   216
  have "\<exists> l r. Rep_game g = Opair l r"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   217
    apply (insert Rep_game[of g])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   218
    apply (simp add: game_def games_lfp_represent)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   219
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   220
  then obtain l r where lr: "Rep_game g = Opair l r" by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   221
  have partizan_g: "Rep_game g \<in> games_lfp" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   222
    apply (insert Rep_game[of g])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   223
    apply (simp add: game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   224
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   225
  have "\<forall> e. Elem e l \<longrightarrow> left_option (Rep_game g) e"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   226
    by (simp add: lr left_option_def Fst)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   227
  then have partizan_l: "\<forall> e. Elem e l \<longrightarrow> e \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   228
    apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   229
    apply (rule games_lfp_option_stable[where g="Rep_game g", OF partizan_g])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   230
    apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   231
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   232
  have "\<forall> e. Elem e r \<longrightarrow> right_option (Rep_game g) e"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   233
    by (simp add: lr right_option_def Snd)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   234
  then have partizan_r: "\<forall> e. Elem e r \<longrightarrow> e \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   235
    apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   236
    apply (rule games_lfp_option_stable[where g="Rep_game g", OF partizan_g])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   237
    apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   238
    done   
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   239
  let ?L = "zimage (Abs_game) (zexplode l)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   240
  let ?R = "zimage (Abs_game) (zexplode r)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   241
  have L:"?L = left_options g"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   242
    by (simp add: left_options_def lr Fst)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   243
  have R:"?R = right_options g"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   244
    by (simp add: right_options_def lr Snd)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   245
  have "g = Game ?L ?R"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   246
    apply (simp add: Game_def Rep_game_inject[symmetric] comp_zimage_eq zimage_zexplode_eq zimplode_zexplode)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   247
    apply (simp add: Repl_Rep_game_Abs_game partizan_l partizan_r)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   248
    apply (subst Abs_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   249
    apply (simp_all add: lr[symmetric] Rep_game) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   250
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   251
  then show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   252
    by (simp add: L R)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   253
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   254
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   255
lemma Opair_in_games_lfp: 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   256
  assumes l: "explode l \<subseteq> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   257
  and r: "explode r \<subseteq> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   258
  shows "Opair l r \<in> games_lfp"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   259
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   260
  note f = unique_games[of games_lfp, simplified]
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   261
  show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   262
    apply (subst f)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   263
    apply (simp add: fixgames_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   264
    apply (rule exI[where x=l])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   265
    apply (rule exI[where x=r])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   266
    apply (auto simp add: l r)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   267
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   268
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   269
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   270
lemma left_options[simp]: "left_options (Game l r) = l"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   271
  apply (simp add: left_options_def Game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   272
  apply (subst Abs_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   273
  apply (simp add: game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   274
  apply (rule Opair_in_games_lfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   275
  apply (auto simp add: explode_Elem Elem_zimplode zimage_iff Rep_game[simplified game_def])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   276
  apply (simp add: Fst zexplode_zimplode comp_zimage_eq)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   277
  apply (simp add: zet_ext_eq zimage_iff Rep_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   278
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   279
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   280
lemma right_options[simp]: "right_options (Game l r) = r"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   281
  apply (simp add: right_options_def Game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   282
  apply (subst Abs_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   283
  apply (simp add: game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   284
  apply (rule Opair_in_games_lfp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   285
  apply (auto simp add: explode_Elem Elem_zimplode zimage_iff Rep_game[simplified game_def])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   286
  apply (simp add: Snd zexplode_zimplode comp_zimage_eq)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   287
  apply (simp add: zet_ext_eq zimage_iff Rep_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   288
  done  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   289
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   290
lemma Game_ext: "(Game l1 r1 = Game l2 r2) = ((l1 = l2) \<and> (r1 = r2))"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   291
  apply auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   292
  apply (subst left_options[where l=l1 and r=r1,symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   293
  apply (subst left_options[where l=l2 and r=r2,symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   294
  apply simp
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   295
  apply (subst right_options[where l=l1 and r=r1,symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   296
  apply (subst right_options[where l=l2 and r=r2,symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   297
  apply simp
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   298
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   299
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   300
definition option_of :: "(game * game) set" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   301
  "option_of \<equiv> image (\<lambda> (option, g). (Abs_game option, Abs_game g)) is_option_of"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   302
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   303
lemma option_to_is_option_of: "((option, g) \<in> option_of) = ((Rep_game option, Rep_game g) \<in> is_option_of)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   304
  apply (auto simp add: option_of_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   305
  apply (subst Abs_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   306
  apply (simp add: is_option_of_imp_games game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   307
  apply (subst Abs_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   308
  apply (simp add: is_option_of_imp_games game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   309
  apply simp
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   310
  apply (auto simp add: Bex_def image_def)  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   311
  apply (rule exI[where x="Rep_game option"])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   312
  apply (rule exI[where x="Rep_game g"])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   313
  apply (simp add: Rep_game_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   314
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   315
  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   316
lemma wf_is_option_of: "wf is_option_of"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   317
  apply (rule wfzf_implies_wf)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   318
  apply (simp add: wfzf_is_option_of)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   319
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   320
44011
f67c93f52d13 eliminated obsolete recdef/wfrec related declarations
krauss
parents: 41777
diff changeset
   321
lemma wf_option_of[simp, intro]: "wf option_of"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   322
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   323
  have option_of: "option_of = inv_image is_option_of Rep_game"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 35440
diff changeset
   324
    apply (rule set_eqI)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   325
    apply (case_tac "x")
19769
c40ce2de2020 Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
krauss
parents: 19203
diff changeset
   326
    by (simp add: option_to_is_option_of) 
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   327
  show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   328
    apply (simp add: option_of)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   329
    apply (auto intro: wf_inv_image wf_is_option_of)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   330
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   331
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   332
  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   333
lemma right_option_is_option[simp, intro]: "zin x (right_options g) \<Longrightarrow> zin x (options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   334
  by (simp add: options_def zunion)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   335
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   336
lemma left_option_is_option[simp, intro]: "zin x (left_options g) \<Longrightarrow> zin x (options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   337
  by (simp add: options_def zunion)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   338
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   339
lemma zin_options[simp, intro]: "zin x (options g) \<Longrightarrow> (x, g) \<in> option_of"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   340
  apply (simp add: options_def zunion left_options_def right_options_def option_of_def 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   341
    image_def is_option_of_def zimage_iff zin_zexplode_eq) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   342
  apply (cases g)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   343
  apply (cases x)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   344
  apply (auto simp add: Abs_game_inverse games_lfp_eq_gfp[symmetric] game_def 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   345
    right_option_def[symmetric] left_option_def[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   346
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   347
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   348
function
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   349
  neg_game :: "game \<Rightarrow> game"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   350
where
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   351
  [simp del]: "neg_game g = Game (zimage neg_game (right_options g)) (zimage neg_game (left_options g))"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   352
by auto
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   353
termination by (relation "option_of") auto
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   354
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   355
lemma "neg_game (neg_game g) = g"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   356
  apply (induct g rule: neg_game.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   357
  apply (subst neg_game.simps)+
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   358
  apply (simp add: right_options left_options comp_zimage_eq)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   359
  apply (subgoal_tac "zimage (neg_game o neg_game) (left_options g) = left_options g")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   360
  apply (subgoal_tac "zimage (neg_game o neg_game) (right_options g) = right_options g")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   361
  apply (auto simp add: game_split[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   362
  apply (auto simp add: zet_ext_eq zimage_iff)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   363
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   364
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   365
function
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   366
  ge_game :: "(game * game) \<Rightarrow> bool" 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   367
where
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   368
  [simp del]: "ge_game (G, H) = (\<forall> x. if zin x (right_options G) then (
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   369
                          if zin x (left_options H) then \<not> (ge_game (H, x) \<or> (ge_game (x, G))) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   370
                                                    else \<not> (ge_game (H, x)))
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   371
                          else (if zin x (left_options H) then \<not> (ge_game (x, G)) else True))"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   372
by auto
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   373
termination by (relation "(gprod_2_1 option_of)") 
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   374
 (simp, auto simp: gprod_2_1_def)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   375
26304
02fbd0e7954a avoid rebinding of existing facts;
wenzelm
parents: 25764
diff changeset
   376
lemma ge_game_eq: "ge_game (G, H) = (\<forall> x. (zin x (right_options G) \<longrightarrow> \<not> ge_game (H, x)) \<and> (zin x (left_options H) \<longrightarrow> \<not> ge_game (x, G)))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   377
  apply (subst ge_game.simps[where G=G and H=H])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   378
  apply (auto)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   379
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   380
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   381
lemma ge_game_leftright_refl[rule_format]: 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   382
  "\<forall> y. (zin y (right_options x) \<longrightarrow> \<not> ge_game (x, y)) \<and> (zin y (left_options x) \<longrightarrow> \<not> (ge_game (y, x))) \<and> ge_game (x, x)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   383
proof (induct x rule: wf_induct[OF wf_option_of]) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   384
  case (1 "g")
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   385
  { 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   386
    fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   387
    assume y: "zin y (right_options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   388
    have "\<not> ge_game (g, y)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   389
    proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   390
      have "(y, g) \<in> option_of" by (auto intro: y)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   391
      with 1 have "ge_game (y, y)" by auto
26304
02fbd0e7954a avoid rebinding of existing facts;
wenzelm
parents: 25764
diff changeset
   392
      with y show ?thesis by (subst ge_game_eq, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   393
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   394
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   395
  note right = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   396
  { 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   397
    fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   398
    assume y: "zin y (left_options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   399
    have "\<not> ge_game (y, g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   400
    proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   401
      have "(y, g) \<in> option_of" by (auto intro: y)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   402
      with 1 have "ge_game (y, y)" by auto
26304
02fbd0e7954a avoid rebinding of existing facts;
wenzelm
parents: 25764
diff changeset
   403
      with y show ?thesis by (subst ge_game_eq, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   404
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   405
  } 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   406
  note left = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   407
  from left right show ?case
26304
02fbd0e7954a avoid rebinding of existing facts;
wenzelm
parents: 25764
diff changeset
   408
    by (auto, subst ge_game_eq, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   409
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   410
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   411
lemma ge_game_refl: "ge_game (x,x)" by (simp add: ge_game_leftright_refl)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   412
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   413
lemma "\<forall> y. (zin y (right_options x) \<longrightarrow> \<not> ge_game (x, y)) \<and> (zin y (left_options x) \<longrightarrow> \<not> (ge_game (y, x))) \<and> ge_game (x, x)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   414
proof (induct x rule: wf_induct[OF wf_option_of]) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   415
  case (1 "g")  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   416
  show ?case
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   417
  proof (auto)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   418
    {case (goal1 y) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   419
      from goal1 have "(y, g) \<in> option_of" by (auto)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   420
      with 1 have "ge_game (y, y)" by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   421
      with goal1 have "\<not> ge_game (g, y)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   422
        by (subst ge_game_eq, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   423
      with goal1 show ?case by auto}
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   424
    note right = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   425
    {case (goal2 y)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   426
      from goal2 have "(y, g) \<in> option_of" by (auto)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   427
      with 1 have "ge_game (y, y)" by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   428
      with goal2 have "\<not> ge_game (y, g)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   429
        by (subst ge_game_eq, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   430
      with goal2 show ?case by auto}
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   431
    note left = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   432
    {case goal3
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   433
      from left right show ?case
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   434
        by (subst ge_game_eq, auto)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   435
    }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   436
  qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   437
qed
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   438
        
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   439
definition eq_game :: "game \<Rightarrow> game \<Rightarrow> bool" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   440
  "eq_game G H \<equiv> ge_game (G, H) \<and> ge_game (H, G)" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   441
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   442
lemma eq_game_sym: "(eq_game G H) = (eq_game H G)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   443
  by (auto simp add: eq_game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   444
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   445
lemma eq_game_refl: "eq_game G G"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   446
  by (simp add: ge_game_refl eq_game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   447
23771
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 22282
diff changeset
   448
lemma induct_game: "(\<And>x. \<forall>y. (y, x) \<in> lprod option_of \<longrightarrow> P y \<Longrightarrow> P x) \<Longrightarrow> P a"
bde6db239efa Restored set notation in Multiset theory.
berghofe
parents: 22282
diff changeset
   449
  by (erule wf_induct[OF wf_lprod[OF wf_option_of]])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   450
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   451
lemma ge_game_trans:
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   452
  assumes "ge_game (x, y)" "ge_game (y, z)" 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   453
  shows "ge_game (x, z)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   454
proof -  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   455
  { 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   456
    fix a
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   457
    have "\<forall> x y z. a = [x,y,z] \<longrightarrow> ge_game (x,y) \<longrightarrow> ge_game (y,z) \<longrightarrow> ge_game (x, z)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   458
    proof (induct a rule: induct_game)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   459
      case (1 a)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   460
      show ?case
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   461
      proof (rule allI | rule impI)+
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   462
        case (goal1 x y z)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   463
        show ?case
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   464
        proof -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   465
          { fix xr
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   466
            assume xr:"zin xr (right_options x)"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   467
            assume a: "ge_game (z, xr)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   468
            have "ge_game (y, xr)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   469
              apply (rule 1[rule_format, where y="[y,z,xr]"])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   470
              apply (auto intro: xr lprod_3_1 simp add: goal1 a)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   471
              done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   472
            moreover from xr have "\<not> ge_game (y, xr)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   473
              by (simp add: goal1(2)[simplified ge_game_eq[of x y], rule_format, of xr, simplified xr])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   474
            ultimately have "False" by auto      
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   475
          }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   476
          note xr = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   477
          { fix zl
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   478
            assume zl:"zin zl (left_options z)"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   479
            assume a: "ge_game (zl, x)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   480
            have "ge_game (zl, y)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   481
              apply (rule 1[rule_format, where y="[zl,x,y]"])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   482
              apply (auto intro: zl lprod_3_2 simp add: goal1 a)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   483
              done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   484
            moreover from zl have "\<not> ge_game (zl, y)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   485
              by (simp add: goal1(3)[simplified ge_game_eq[of y z], rule_format, of zl, simplified zl])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   486
            ultimately have "False" by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   487
          }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   488
          note zl = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   489
          show ?thesis
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   490
            by (auto simp add: ge_game_eq[of x z] intro: xr zl)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   491
        qed
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   492
      qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   493
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   494
  } 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   495
  note trans = this[of "[x, y, z]", simplified, rule_format]    
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   496
  with assms show ?thesis by blast
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   497
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   498
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   499
lemma eq_game_trans: "eq_game a b \<Longrightarrow> eq_game b c \<Longrightarrow> eq_game a c"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   500
  by (auto simp add: eq_game_def intro: ge_game_trans)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   501
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   502
definition zero_game :: game
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   503
 where  "zero_game \<equiv> Game zempty zempty"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   504
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   505
function 
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   506
  plus_game :: "game \<Rightarrow> game \<Rightarrow> game"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   507
where
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   508
  [simp del]: "plus_game G H = Game (zunion (zimage (\<lambda> g. plus_game g H) (left_options G))
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   509
                                   (zimage (\<lambda> h. plus_game G h) (left_options H)))
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   510
                           (zunion (zimage (\<lambda> g. plus_game g H) (right_options G))
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   511
                                   (zimage (\<lambda> h. plus_game G h) (right_options H)))"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   512
by auto
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   513
termination by (relation "gprod_2_2 option_of")
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   514
  (simp, auto simp: gprod_2_2_def)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   515
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   516
lemma plus_game_comm: "plus_game G H = plus_game H G"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   517
proof (induct G H rule: plus_game.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   518
  case (1 G H)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   519
  show ?case
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   520
    by (auto simp add: 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   521
      plus_game.simps[where G=G and H=H] 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   522
      plus_game.simps[where G=H and H=G]
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   523
      Game_ext zet_ext_eq zunion zimage_iff 1)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   524
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   525
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   526
lemma game_ext_eq: "(G = H) = (left_options G = left_options H \<and> right_options G = right_options H)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   527
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   528
  have "(G = H) = (Game (left_options G) (right_options G) = Game (left_options H) (right_options H))"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   529
    by (simp add: game_split[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   530
  then show ?thesis by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   531
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   532
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   533
lemma left_zero_game[simp]: "left_options (zero_game) = zempty"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   534
  by (simp add: zero_game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   535
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   536
lemma right_zero_game[simp]: "right_options (zero_game) = zempty"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   537
  by (simp add: zero_game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   538
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   539
lemma plus_game_zero_right[simp]: "plus_game G zero_game = G"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   540
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   541
  { 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   542
    fix G H
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   543
    have "H = zero_game \<longrightarrow> plus_game G H = G "
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   544
    proof (induct G H rule: plus_game.induct, rule impI)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   545
      case (goal1 G H)
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   546
      note induct_hyp = this[simplified goal1, simplified] and this
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   547
      show ?case
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   548
        apply (simp only: plus_game.simps[where G=G and H=H])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   549
        apply (simp add: game_ext_eq goal1)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   550
        apply (auto simp add: 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   551
          zimage_cong[where f = "\<lambda> g. plus_game g zero_game" and g = "id"] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   552
          induct_hyp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   553
        done
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   554
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   555
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   556
  then show ?thesis by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   557
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   558
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   559
lemma plus_game_zero_left: "plus_game zero_game G = G"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   560
  by (simp add: plus_game_comm)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   561
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   562
lemma left_imp_options[simp]: "zin opt (left_options g) \<Longrightarrow> zin opt (options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   563
  by (simp add: options_def zunion)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   564
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   565
lemma right_imp_options[simp]: "zin opt (right_options g) \<Longrightarrow> zin opt (options g)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   566
  by (simp add: options_def zunion)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   567
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   568
lemma left_options_plus: 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   569
  "left_options (plus_game u v) =  zunion (zimage (\<lambda>g. plus_game g v) (left_options u)) (zimage (\<lambda>h. plus_game u h) (left_options v))" 
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   570
  by (subst plus_game.simps, simp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   571
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   572
lemma right_options_plus:
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   573
  "right_options (plus_game u v) =  zunion (zimage (\<lambda>g. plus_game g v) (right_options u)) (zimage (\<lambda>h. plus_game u h) (right_options v))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   574
  by (subst plus_game.simps, simp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   575
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   576
lemma left_options_neg: "left_options (neg_game u) = zimage neg_game (right_options u)"  
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   577
  by (subst neg_game.simps, simp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   578
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   579
lemma right_options_neg: "right_options (neg_game u) = zimage neg_game (left_options u)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   580
  by (subst neg_game.simps, simp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   581
  
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   582
lemma plus_game_assoc: "plus_game (plus_game F G) H = plus_game F (plus_game G H)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   583
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   584
  { 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   585
    fix a
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   586
    have "\<forall> F G H. a = [F, G, H] \<longrightarrow> plus_game (plus_game F G) H = plus_game F (plus_game G H)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   587
    proof (induct a rule: induct_game, (rule impI | rule allI)+)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   588
      case (goal1 x F G H)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   589
      let ?L = "plus_game (plus_game F G) H"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   590
      let ?R = "plus_game F (plus_game G H)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   591
      note options_plus = left_options_plus right_options_plus
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   592
      {
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   593
        fix opt
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   594
        note hyp = goal1(1)[simplified goal1(2), rule_format] 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   595
        have F: "zin opt (options F)  \<Longrightarrow> plus_game (plus_game opt G) H = plus_game opt (plus_game G H)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   596
          by (blast intro: hyp lprod_3_3)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   597
        have G: "zin opt (options G) \<Longrightarrow> plus_game (plus_game F opt) H = plus_game F (plus_game opt H)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   598
          by (blast intro: hyp lprod_3_4)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   599
        have H: "zin opt (options H) \<Longrightarrow> plus_game (plus_game F G) opt = plus_game F (plus_game G opt)" 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   600
          by (blast intro: hyp lprod_3_5)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   601
        note F and G and H
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   602
      }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   603
      note induct_hyp = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   604
      have "left_options ?L = left_options ?R \<and> right_options ?L = right_options ?R"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   605
        by (auto simp add: 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   606
          plus_game.simps[where G="plus_game F G" and H=H]
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   607
          plus_game.simps[where G="F" and H="plus_game G H"] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   608
          zet_ext_eq zunion zimage_iff options_plus
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   609
          induct_hyp left_imp_options right_imp_options)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   610
      then show ?case
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   611
        by (simp add: game_ext_eq)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   612
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   613
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   614
  then show ?thesis by auto
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   615
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   616
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   617
lemma neg_plus_game: "neg_game (plus_game G H) = plus_game (neg_game G) (neg_game H)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   618
proof (induct G H rule: plus_game.induct)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   619
  case (1 G H)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   620
  note opt_ops = 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   621
    left_options_plus right_options_plus 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   622
    left_options_neg right_options_neg  
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   623
  show ?case
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   624
    by (auto simp add: opt_ops
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   625
      neg_game.simps[of "plus_game G H"]
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   626
      plus_game.simps[of "neg_game G" "neg_game H"]
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   627
      Game_ext zet_ext_eq zunion zimage_iff 1)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   628
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   629
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   630
lemma eq_game_plus_inverse: "eq_game (plus_game x (neg_game x)) zero_game"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   631
proof (induct x rule: wf_induct[OF wf_option_of])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   632
  case (goal1 x)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   633
  { fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   634
    assume "zin y (options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   635
    then have "eq_game (plus_game y (neg_game y)) zero_game"
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   636
      by (auto simp add: goal1)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   637
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   638
  note ihyp = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   639
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   640
    fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   641
    assume y: "zin y (right_options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   642
    have "\<not> (ge_game (zero_game, plus_game y (neg_game x)))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   643
      apply (subst ge_game.simps, simp)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   644
      apply (rule exI[where x="plus_game y (neg_game y)"])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   645
      apply (auto simp add: ihyp[of y, simplified y right_imp_options eq_game_def])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   646
      apply (auto simp add: left_options_plus left_options_neg zunion zimage_iff intro: y)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   647
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   648
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   649
  note case1 = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   650
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   651
    fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   652
    assume y: "zin y (left_options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   653
    have "\<not> (ge_game (zero_game, plus_game x (neg_game y)))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   654
      apply (subst ge_game.simps, simp)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   655
      apply (rule exI[where x="plus_game y (neg_game y)"])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   656
      apply (auto simp add: ihyp[of y, simplified y left_imp_options eq_game_def])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   657
      apply (auto simp add: left_options_plus zunion zimage_iff intro: y)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   658
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   659
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   660
  note case2 = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   661
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   662
    fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   663
    assume y: "zin y (left_options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   664
    have "\<not> (ge_game (plus_game y (neg_game x), zero_game))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   665
      apply (subst ge_game.simps, simp)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   666
      apply (rule exI[where x="plus_game y (neg_game y)"])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   667
      apply (auto simp add: ihyp[of y, simplified y left_imp_options eq_game_def])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   668
      apply (auto simp add: right_options_plus right_options_neg zunion zimage_iff intro: y)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   669
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   670
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   671
  note case3 = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   672
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   673
    fix y
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   674
    assume y: "zin y (right_options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   675
    have "\<not> (ge_game (plus_game x (neg_game y), zero_game))"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   676
      apply (subst ge_game.simps, simp)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   677
      apply (rule exI[where x="plus_game y (neg_game y)"])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   678
      apply (auto simp add: ihyp[of y, simplified y right_imp_options eq_game_def])
41528
276078f01ada eliminated global prems;
wenzelm
parents: 40824
diff changeset
   679
      apply (auto simp add: right_options_plus zunion zimage_iff intro: y)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   680
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   681
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   682
  note case4 = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   683
  show ?case
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   684
    apply (simp add: eq_game_def)
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   685
    apply (simp add: ge_game.simps[of "plus_game x (neg_game x)" "zero_game"])
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   686
    apply (simp add: ge_game.simps[of "zero_game" "plus_game x (neg_game x)"])
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   687
    apply (simp add: right_options_plus left_options_plus right_options_neg left_options_neg zunion zimage_iff)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   688
    apply (auto simp add: case1 case2 case3 case4)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   689
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   690
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   691
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   692
lemma ge_plus_game_left: "ge_game (y,z) = ge_game (plus_game x y, plus_game x z)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   693
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   694
  { fix a
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   695
    have "\<forall> x y z. a = [x,y,z] \<longrightarrow> ge_game (y,z) = ge_game (plus_game x y, plus_game x z)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   696
    proof (induct a rule: induct_game, (rule impI | rule allI)+)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   697
      case (goal1 a x y z)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   698
      note induct_hyp = goal1(1)[rule_format, simplified goal1(2)]
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   699
      { 
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   700
        assume hyp: "ge_game(plus_game x y, plus_game x z)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   701
        have "ge_game (y, z)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   702
        proof -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   703
          { fix yr
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   704
            assume yr: "zin yr (right_options y)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   705
            from hyp have "\<not> (ge_game (plus_game x z, plus_game x yr))"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   706
              by (auto simp add: ge_game_eq[of "plus_game x y" "plus_game x z"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   707
                right_options_plus zunion zimage_iff intro: yr)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   708
            then have "\<not> (ge_game (z, yr))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   709
              apply (subst induct_hyp[where y="[x, z, yr]", of "x" "z" "yr"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   710
              apply (simp_all add: yr lprod_3_6)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   711
              done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   712
          }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   713
          note yr = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   714
          { fix zl
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   715
            assume zl: "zin zl (left_options z)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   716
            from hyp have "\<not> (ge_game (plus_game x zl, plus_game x y))"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   717
              by (auto simp add: ge_game_eq[of "plus_game x y" "plus_game x z"]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   718
                left_options_plus zunion zimage_iff intro: zl)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   719
            then have "\<not> (ge_game (zl, y))"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   720
              apply (subst goal1(1)[rule_format, where y="[x, zl, y]", of "x" "zl" "y"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   721
              apply (simp_all add: goal1(2) zl lprod_3_7)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   722
              done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   723
          }     
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   724
          note zl = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   725
          show "ge_game (y, z)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   726
            apply (subst ge_game_eq)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   727
            apply (auto simp add: yr zl)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   728
            done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   729
        qed      
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   730
      }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   731
      note right_imp_left = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   732
      {
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   733
        assume yz: "ge_game (y, z)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   734
        {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   735
          fix x'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   736
          assume x': "zin x' (right_options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   737
          assume hyp: "ge_game (plus_game x z, plus_game x' y)"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   738
          then have n: "\<not> (ge_game (plus_game x' y, plus_game x' z))"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   739
            by (auto simp add: ge_game_eq[of "plus_game x z" "plus_game x' y"] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   740
              right_options_plus zunion zimage_iff intro: x')
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   741
          have t: "ge_game (plus_game x' y, plus_game x' z)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   742
            apply (subst induct_hyp[symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   743
            apply (auto intro: lprod_3_3 x' yz)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   744
            done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   745
          from n t have "False" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   746
        }    
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   747
        note case1 = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   748
        {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   749
          fix x'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   750
          assume x': "zin x' (left_options x)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   751
          assume hyp: "ge_game (plus_game x' z, plus_game x y)"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   752
          then have n: "\<not> (ge_game (plus_game x' y, plus_game x' z))"
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   753
            by (auto simp add: ge_game_eq[of "plus_game x' z" "plus_game x y"] 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   754
              left_options_plus zunion zimage_iff intro: x')
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   755
          have t: "ge_game (plus_game x' y, plus_game x' z)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   756
            apply (subst induct_hyp[symmetric])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   757
            apply (auto intro: lprod_3_3 x' yz)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   758
            done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   759
          from n t have "False" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   760
        }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   761
        note case3 = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   762
        {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   763
          fix y'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   764
          assume y': "zin y' (right_options y)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   765
          assume hyp: "ge_game (plus_game x z, plus_game x y')"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   766
          then have "ge_game(z, y')"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   767
            apply (subst induct_hyp[of "[x, z, y']" "x" "z" "y'"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   768
            apply (auto simp add: hyp lprod_3_6 y')
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   769
            done
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   770
          with yz have "ge_game (y, y')"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   771
            by (blast intro: ge_game_trans)      
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   772
          with y' have "False" by (auto simp add: ge_game_leftright_refl)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   773
        }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   774
        note case2 = this
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   775
        {
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   776
          fix z'
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   777
          assume z': "zin z' (left_options z)"
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   778
          assume hyp: "ge_game (plus_game x z', plus_game x y)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   779
          then have "ge_game(z', y)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   780
            apply (subst induct_hyp[of "[x, z', y]" "x" "z'" "y"])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   781
            apply (auto simp add: hyp lprod_3_7 z')
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   782
            done    
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   783
          with yz have "ge_game (z', z)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   784
            by (blast intro: ge_game_trans)      
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   785
          with z' have "False" by (auto simp add: ge_game_leftright_refl)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   786
        }
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   787
        note case4 = this   
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   788
        have "ge_game(plus_game x y, plus_game x z)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   789
          apply (subst ge_game_eq)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   790
          apply (auto simp add: right_options_plus left_options_plus zunion zimage_iff)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   791
          apply (auto intro: case1 case2 case3 case4)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   792
          done
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   793
      }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   794
      note left_imp_right = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   795
      show ?case by (auto intro: right_imp_left left_imp_right)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   796
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   797
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   798
  note a = this[of "[x, y, z]"]
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   799
  then show ?thesis by blast
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   800
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   801
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   802
lemma ge_plus_game_right: "ge_game (y,z) = ge_game(plus_game y x, plus_game z x)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   803
  by (simp add: ge_plus_game_left plus_game_comm)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   804
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   805
lemma ge_neg_game: "ge_game (neg_game x, neg_game y) = ge_game (y, x)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   806
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   807
  { fix a
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   808
    have "\<forall> x y. a = [x, y] \<longrightarrow> ge_game (neg_game x, neg_game y) = ge_game (y, x)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   809
    proof (induct a rule: induct_game, (rule impI | rule allI)+)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   810
      case (goal1 a x y)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   811
      note ihyp = goal1(1)[rule_format, simplified goal1(2)]
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   812
      { fix xl
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   813
        assume xl: "zin xl (left_options x)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   814
        have "ge_game (neg_game y, neg_game xl) = ge_game (xl, y)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   815
          apply (subst ihyp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   816
          apply (auto simp add: lprod_2_1 xl)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   817
          done
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   818
      }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   819
      note xl = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   820
      { fix yr
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   821
        assume yr: "zin yr (right_options y)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   822
        have "ge_game (neg_game yr, neg_game x) = ge_game (x, yr)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   823
          apply (subst ihyp)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   824
          apply (auto simp add: lprod_2_2 yr)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   825
          done
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   826
      }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   827
      note yr = this
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   828
      show ?case
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   829
        by (auto simp add: ge_game_eq[of "neg_game x" "neg_game y"] ge_game_eq[of "y" "x"]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 30198
diff changeset
   830
          right_options_neg left_options_neg zimage_iff  xl yr)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   831
    qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   832
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   833
  note a = this[of "[x,y]"]
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   834
  then show ?thesis by blast
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   835
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   836
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35028
diff changeset
   837
definition eq_game_rel :: "(game * game) set" where
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   838
  "eq_game_rel \<equiv> { (p, q) . eq_game p q }"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   839
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   840
definition "Pg = UNIV//eq_game_rel"
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   841
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   842
typedef (open) Pg = Pg
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 44011
diff changeset
   843
  unfolding Pg_def by (auto simp add: quotient_def)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   844
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   845
lemma equiv_eq_game[simp]: "equiv UNIV eq_game_rel"
30198
922f944f03b2 name changes
nipkow
parents: 27679
diff changeset
   846
  by (auto simp add: equiv_def refl_on_def sym_def trans_def eq_game_rel_def
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   847
    eq_game_sym intro: eq_game_refl eq_game_trans)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   848
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   849
instantiation Pg :: "{ord, zero, plus, minus, uminus}"
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   850
begin
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   851
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   852
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   853
  Pg_zero_def: "0 = Abs_Pg (eq_game_rel `` {zero_game})"
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   854
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   855
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   856
  Pg_le_def: "G \<le> H \<longleftrightarrow> (\<exists> g h. g \<in> Rep_Pg G \<and> h \<in> Rep_Pg H \<and> ge_game (h, g))"
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   857
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   858
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   859
  Pg_less_def: "G < H \<longleftrightarrow> G \<le> H \<and> G \<noteq> (H::Pg)"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   860
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   861
definition
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   862
  Pg_minus_def: "- G = the_elem (\<Union> g \<in> Rep_Pg G. {Abs_Pg (eq_game_rel `` {neg_game g})})"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   863
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   864
definition
39910
10097e0a9dbd constant `contents` renamed to `the_elem`
haftmann
parents: 39302
diff changeset
   865
  Pg_plus_def: "G + H = the_elem (\<Union> g \<in> Rep_Pg G. \<Union> h \<in> Rep_Pg H. {Abs_Pg (eq_game_rel `` {plus_game g h})})"
25764
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   866
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   867
definition
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   868
  Pg_diff_def: "G - H = G + (- (H::Pg))"
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   869
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   870
instance ..
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   871
878c37886eed removed some legacy instantiations
haftmann
parents: 23771
diff changeset
   872
end
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   873
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   874
lemma Rep_Abs_eq_Pg[simp]: "Rep_Pg (Abs_Pg (eq_game_rel `` {g})) = eq_game_rel `` {g}"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   875
  apply (subst Abs_Pg_inverse)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   876
  apply (auto simp add: Pg_def quotient_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   877
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   878
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   879
lemma char_Pg_le[simp]: "(Abs_Pg (eq_game_rel `` {g}) \<le> Abs_Pg (eq_game_rel `` {h})) = (ge_game (h, g))"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   880
  apply (simp add: Pg_le_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   881
  apply (auto simp add: eq_game_rel_def eq_game_def intro: ge_game_trans ge_game_refl)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   882
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   883
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   884
lemma char_Pg_eq[simp]: "(Abs_Pg (eq_game_rel `` {g}) = Abs_Pg (eq_game_rel `` {h})) = (eq_game g h)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   885
  apply (simp add: Rep_Pg_inject [symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   886
  apply (subst eq_equiv_class_iff[of UNIV])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   887
  apply (simp_all)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   888
  apply (simp add: eq_game_rel_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   889
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   890
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   891
lemma char_Pg_plus[simp]: "Abs_Pg (eq_game_rel `` {g}) + Abs_Pg (eq_game_rel `` {h}) = Abs_Pg (eq_game_rel `` {plus_game g h})"
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   892
proof -
35440
bdf8ad377877 killed more recdefs
krauss
parents: 35416
diff changeset
   893
  have "(\<lambda> g h. {Abs_Pg (eq_game_rel `` {plus_game g h})}) respects2 eq_game_rel" 
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   894
    apply (simp add: congruent2_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   895
    apply (auto simp add: eq_game_rel_def eq_game_def)
40824
f5a0cb45d2a5 adapted fragile proof
haftmann
parents: 39910
diff changeset
   896
    apply (rule_tac y="plus_game a ba" in ge_game_trans)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   897
    apply (simp add: ge_plus_game_left[symmetric] ge_plus_game_right[symmetric])+
40824
f5a0cb45d2a5 adapted fragile proof
haftmann
parents: 39910
diff changeset
   898
    apply (rule_tac y="plus_game b aa" in ge_game_trans)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   899
    apply (simp add: ge_plus_game_left[symmetric] ge_plus_game_right[symmetric])+
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   900
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   901
  then show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   902
    by (simp add: Pg_plus_def UN_equiv_class2[OF equiv_eq_game equiv_eq_game]) 
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   903
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   904
    
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   905
lemma char_Pg_minus[simp]: "- Abs_Pg (eq_game_rel `` {g}) = Abs_Pg (eq_game_rel `` {neg_game g})"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   906
proof -
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   907
  have "(\<lambda> g. {Abs_Pg (eq_game_rel `` {neg_game g})}) respects eq_game_rel"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   908
    apply (simp add: congruent_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   909
    apply (auto simp add: eq_game_rel_def eq_game_def ge_neg_game)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   910
    done    
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   911
  then show ?thesis
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   912
    by (simp add: Pg_minus_def UN_equiv_class[OF equiv_eq_game])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   913
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   914
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   915
lemma eq_Abs_Pg[rule_format, cases type: Pg]: "(\<forall> g. z = Abs_Pg (eq_game_rel `` {g}) \<longrightarrow> P) \<longrightarrow> P"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   916
  apply (cases z, simp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   917
  apply (simp add: Rep_Pg_inject[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   918
  apply (subst Abs_Pg_inverse, simp)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   919
  apply (auto simp add: Pg_def quotient_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   920
  done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   921
35028
108662d50512 more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents: 32960
diff changeset
   922
instance Pg :: ordered_ab_group_add 
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   923
proof
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   924
  fix a b c :: Pg
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   925
  show "a - b = a + (- b)" by (simp add: Pg_diff_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   926
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   927
    assume ab: "a \<le> b"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   928
    assume ba: "b \<le> a"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   929
    from ab ba show "a = b"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   930
      apply (cases a, cases b)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   931
      apply (simp add: eq_game_def)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   932
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   933
  }
27679
haftmann
parents: 26304
diff changeset
   934
  then show "(a < b) = (a \<le> b \<and> \<not> b \<le> a)" by (auto simp add: Pg_less_def)
19203
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   935
  show "a + b = b + a"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   936
    apply (cases a, cases b)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   937
    apply (simp add: eq_game_def plus_game_comm)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   938
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   939
  show "a + b + c = a + (b + c)"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   940
    apply (cases a, cases b, cases c)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   941
    apply (simp add: eq_game_def plus_game_assoc)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   942
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   943
  show "0 + a = a"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   944
    apply (cases a)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   945
    apply (simp add: Pg_zero_def plus_game_zero_left)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   946
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   947
  show "- a + a = 0"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   948
    apply (cases a)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   949
    apply (simp add: Pg_zero_def eq_game_plus_inverse plus_game_comm)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   950
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   951
  show "a \<le> a"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   952
    apply (cases a)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   953
    apply (simp add: ge_game_refl)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   954
    done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   955
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   956
    assume ab: "a \<le> b"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   957
    assume bc: "b \<le> c"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   958
    from ab bc show "a \<le> c"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   959
      apply (cases a, cases b, cases c)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   960
      apply (auto intro: ge_game_trans)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   961
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   962
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   963
  {
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   964
    assume ab: "a \<le> b"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   965
    from ab show "c + a \<le> c + b"
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   966
      apply (cases a, cases b, cases c)
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   967
      apply (simp add: ge_plus_game_left[symmetric])
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   968
      done
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   969
  }
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   970
qed
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   971
778507520684 Added HOL-ZF to Isabelle.
obua
parents:
diff changeset
   972
end
46555
c2b5900988e2 tuned proof
haftmann
parents: 45694
diff changeset
   973