1.1 --- a/doc-src/Intro/bool.thy Sat Oct 17 01:05:59 2009 +0200
1.2 +++ b/doc-src/Intro/bool.thy Sat Oct 17 14:43:18 2009 +0200
1.3 @@ -1,5 +1,5 @@
1.4 Bool = FOL +
1.5 -types bool 0
1.6 -arities bool :: term
1.7 -consts tt,ff :: "bool"
1.8 +types bool 0
1.9 +arities bool :: term
1.10 +consts tt,ff :: "bool"
1.11 end
2.1 --- a/doc-src/Intro/list.thy Sat Oct 17 01:05:59 2009 +0200
2.2 +++ b/doc-src/Intro/list.thy Sat Oct 17 14:43:18 2009 +0200
2.3 @@ -1,6 +1,6 @@
2.4 List = FOL +
2.5 -types list 1
2.6 -arities list :: (term)term
2.7 -consts Nil :: "'a list"
2.8 - Cons :: "['a, 'a list] => 'a list"
2.9 +types list 1
2.10 +arities list :: (term)term
2.11 +consts Nil :: "'a list"
2.12 + Cons :: "['a, 'a list] => 'a list"
2.13 end
3.1 --- a/doc-src/Locales/Locales/Examples.thy Sat Oct 17 01:05:59 2009 +0200
3.2 +++ b/doc-src/Locales/Locales/Examples.thy Sat Oct 17 14:43:18 2009 +0200
3.3 @@ -171,13 +171,13 @@
3.4 proof (rule anti_sym)
3.5 from inf' show "i \<sqsubseteq> i'"
3.6 proof (rule is_inf_greatest)
3.7 - from inf show "i \<sqsubseteq> x" ..
3.8 - from inf show "i \<sqsubseteq> y" ..
3.9 + from inf show "i \<sqsubseteq> x" ..
3.10 + from inf show "i \<sqsubseteq> y" ..
3.11 qed
3.12 from inf show "i' \<sqsubseteq> i"
3.13 proof (rule is_inf_greatest)
3.14 - from inf' show "i' \<sqsubseteq> x" ..
3.15 - from inf' show "i' \<sqsubseteq> y" ..
3.16 + from inf' show "i' \<sqsubseteq> x" ..
3.17 + from inf' show "i' \<sqsubseteq> y" ..
3.18 qed
3.19 qed
3.20 qed
3.21 @@ -213,13 +213,13 @@
3.22 proof (rule anti_sym)
3.23 from sup show "s \<sqsubseteq> s'"
3.24 proof (rule is_sup_least)
3.25 - from sup' show "x \<sqsubseteq> s'" ..
3.26 - from sup' show "y \<sqsubseteq> s'" ..
3.27 + from sup' show "x \<sqsubseteq> s'" ..
3.28 + from sup' show "y \<sqsubseteq> s'" ..
3.29 qed
3.30 from sup' show "s' \<sqsubseteq> s"
3.31 proof (rule is_sup_least)
3.32 - from sup show "x \<sqsubseteq> s" ..
3.33 - from sup show "y \<sqsubseteq> s" ..
3.34 + from sup show "x \<sqsubseteq> s" ..
3.35 + from sup show "y \<sqsubseteq> s" ..
3.36 qed
3.37 qed
3.38 qed
3.39 @@ -346,9 +346,9 @@
3.40 show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> x" ..
3.41 show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y"
3.42 proof -
3.43 - have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
3.44 - also have "\<dots> \<sqsubseteq> y" ..
3.45 - finally show ?thesis .
3.46 + have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
3.47 + also have "\<dots> \<sqsubseteq> y" ..
3.48 + finally show ?thesis .
3.49 qed
3.50 qed
3.51 show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> z"
3.52 @@ -362,19 +362,19 @@
3.53 proof
3.54 show "w \<sqsubseteq> x"
3.55 proof -
3.56 - have "w \<sqsubseteq> x \<sqinter> y" by fact
3.57 - also have "\<dots> \<sqsubseteq> x" ..
3.58 - finally show ?thesis .
3.59 + have "w \<sqsubseteq> x \<sqinter> y" by fact
3.60 + also have "\<dots> \<sqsubseteq> x" ..
3.61 + finally show ?thesis .
3.62 qed
3.63 show "w \<sqsubseteq> y \<sqinter> z"
3.64 proof
3.65 - show "w \<sqsubseteq> y"
3.66 - proof -
3.67 - have "w \<sqsubseteq> x \<sqinter> y" by fact
3.68 - also have "\<dots> \<sqsubseteq> y" ..
3.69 - finally show ?thesis .
3.70 - qed
3.71 - show "w \<sqsubseteq> z" by fact
3.72 + show "w \<sqsubseteq> y"
3.73 + proof -
3.74 + have "w \<sqsubseteq> x \<sqinter> y" by fact
3.75 + also have "\<dots> \<sqsubseteq> y" ..
3.76 + finally show ?thesis .
3.77 + qed
3.78 + show "w \<sqsubseteq> z" by fact
3.79 qed
3.80 qed
3.81 qed
3.82 @@ -402,9 +402,9 @@
3.83 show "x \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
3.84 show "y \<sqsubseteq> x \<squnion> (y \<squnion> z)"
3.85 proof -
3.86 - have "y \<sqsubseteq> y \<squnion> z" ..
3.87 - also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
3.88 - finally show ?thesis .
3.89 + have "y \<sqsubseteq> y \<squnion> z" ..
3.90 + also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
3.91 + finally show ?thesis .
3.92 qed
3.93 qed
3.94 show "z \<sqsubseteq> x \<squnion> (y \<squnion> z)"
3.95 @@ -418,19 +418,19 @@
3.96 proof
3.97 show "x \<sqsubseteq> w"
3.98 proof -
3.99 - have "x \<sqsubseteq> x \<squnion> y" ..
3.100 - also have "\<dots> \<sqsubseteq> w" by fact
3.101 - finally show ?thesis .
3.102 + have "x \<sqsubseteq> x \<squnion> y" ..
3.103 + also have "\<dots> \<sqsubseteq> w" by fact
3.104 + finally show ?thesis .
3.105 qed
3.106 show "y \<squnion> z \<sqsubseteq> w"
3.107 proof
3.108 - show "y \<sqsubseteq> w"
3.109 - proof -
3.110 - have "y \<sqsubseteq> x \<squnion> y" ..
3.111 - also have "... \<sqsubseteq> w" by fact
3.112 - finally show ?thesis .
3.113 - qed
3.114 - show "z \<sqsubseteq> w" by fact
3.115 + show "y \<sqsubseteq> w"
3.116 + proof -
3.117 + have "y \<sqsubseteq> x \<squnion> y" ..
3.118 + also have "... \<sqsubseteq> w" by fact
3.119 + finally show ?thesis .
3.120 + qed
3.121 + show "z \<sqsubseteq> w" by fact
3.122 qed
3.123 qed
3.124 qed
3.125 @@ -650,17 +650,17 @@
3.126 txt {* Jacobson I, p.\ 462 *}
3.127 proof -
3.128 { assume c: "y \<sqsubseteq> x" "z \<sqsubseteq> x"
3.129 - from c have "?l = y \<squnion> z"
3.130 - by (metis c join_connection2 join_related2 meet_related2 total)
3.131 - also from c have "... = ?r" by (metis meet_related2)
3.132 - finally have "?l = ?r" . }
3.133 + from c have "?l = y \<squnion> z"
3.134 + by (metis c join_connection2 join_related2 meet_related2 total)
3.135 + also from c have "... = ?r" by (metis meet_related2)
3.136 + finally have "?l = ?r" . }
3.137 moreover
3.138 { assume c: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z"
3.139 - from c have "?l = x"
3.140 - by (metis join_connection2 join_related2 meet_connection total trans)
3.141 - also from c have "... = ?r"
3.142 - by (metis join_commute join_related2 meet_connection meet_related2 total)
3.143 - finally have "?l = ?r" . }
3.144 + from c have "?l = x"
3.145 + by (metis join_connection2 join_related2 meet_connection total trans)
3.146 + also from c have "... = ?r"
3.147 + by (metis join_commute join_related2 meet_connection meet_related2 total)
3.148 + finally have "?l = ?r" . }
3.149 moreover note total
3.150 ultimately show ?thesis by blast
3.151 qed
4.1 --- a/doc-src/TutorialI/Overview/LNCS/Sets.thy Sat Oct 17 01:05:59 2009 +0200
4.2 +++ b/doc-src/TutorialI/Overview/LNCS/Sets.thy Sat Oct 17 14:43:18 2009 +0200
4.3 @@ -149,7 +149,7 @@
4.4 proof (rule converse_rtrancl_induct)
4.5 show "t \<in> lfp ?F"
4.6 proof (subst lfp_unfold[OF mono_ef])
4.7 - show "t \<in> ?F(lfp ?F)" using tA by blast
4.8 + show "t \<in> ?F(lfp ?F)" using tA by blast
4.9 qed
4.10 next
4.11 fix s s'
4.12 @@ -157,7 +157,7 @@
4.13 and IH: "s' \<in> lfp ?F"
4.14 show "s \<in> lfp ?F"
4.15 proof (subst lfp_unfold[OF mono_ef])
4.16 - show "s \<in> ?F(lfp ?F)" using prems by blast
4.17 + show "s \<in> ?F(lfp ?F)" using prems by blast
4.18 qed
4.19 qed
4.20 qed
5.1 --- a/doc-src/TutorialI/Protocol/Event.thy Sat Oct 17 01:05:59 2009 +0200
5.2 +++ b/doc-src/TutorialI/Protocol/Event.thy Sat Oct 17 14:43:18 2009 +0200
5.3 @@ -22,7 +22,7 @@
5.4 | Notes agent msg
5.5
5.6 consts
5.7 - bad :: "agent set" (*compromised agents*)
5.8 + bad :: "agent set" -- {* compromised agents *}
5.9 knows :: "agent => event list => msg set"
5.10
5.11
5.12 @@ -43,19 +43,19 @@
5.13 knows_Cons:
5.14 "knows A (ev # evs) =
5.15 (if A = Spy then
5.16 - (case ev of
5.17 - Says A' B X => insert X (knows Spy evs)
5.18 - | Gets A' X => knows Spy evs
5.19 - | Notes A' X =>
5.20 - if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
5.21 - else
5.22 - (case ev of
5.23 - Says A' B X =>
5.24 - if A'=A then insert X (knows A evs) else knows A evs
5.25 - | Gets A' X =>
5.26 - if A'=A then insert X (knows A evs) else knows A evs
5.27 - | Notes A' X =>
5.28 - if A'=A then insert X (knows A evs) else knows A evs))"
5.29 + (case ev of
5.30 + Says A' B X => insert X (knows Spy evs)
5.31 + | Gets A' X => knows Spy evs
5.32 + | Notes A' X =>
5.33 + if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
5.34 + else
5.35 + (case ev of
5.36 + Says A' B X =>
5.37 + if A'=A then insert X (knows A evs) else knows A evs
5.38 + | Gets A' X =>
5.39 + if A'=A then insert X (knows A evs) else knows A evs
5.40 + | Notes A' X =>
5.41 + if A'=A then insert X (knows A evs) else knows A evs))"
5.42
5.43 (*
5.44 Case A=Spy on the Gets event
5.45 @@ -71,10 +71,10 @@
5.46 primrec
5.47 used_Nil: "used [] = (UN B. parts (initState B))"
5.48 used_Cons: "used (ev # evs) =
5.49 - (case ev of
5.50 - Says A B X => parts {X} \<union> used evs
5.51 - | Gets A X => used evs
5.52 - | Notes A X => parts {X} \<union> used evs)"
5.53 + (case ev of
5.54 + Says A B X => parts {X} \<union> used evs
5.55 + | Gets A X => used evs
5.56 + | Notes A X => parts {X} \<union> used evs)"
5.57 --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
5.58 follows @{term Says} in real protocols. Seems difficult to change.
5.59 See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
6.1 --- a/doc-src/TutorialI/Protocol/Message.thy Sat Oct 17 01:05:59 2009 +0200
6.2 +++ b/doc-src/TutorialI/Protocol/Message.thy Sat Oct 17 14:43:18 2009 +0200
6.3 @@ -61,8 +61,8 @@
6.4 msg = Agent agent
6.5 | Nonce nat
6.6 | Key key
6.7 - | MPair msg msg
6.8 - | Crypt key msg
6.9 + | MPair msg msg
6.10 + | Crypt key msg
6.11
6.12 text {*
6.13 \noindent
6.14 @@ -855,8 +855,8 @@
6.15 (Fake_insert_simp_tac ss 1
6.16 THEN
6.17 IF_UNSOLVED (Blast.depth_tac
6.18 - (cs addIs [analz_insertI,
6.19 - impOfSubs analz_subset_parts]) 4 1))
6.20 + (cs addIs [analz_insertI,
6.21 + impOfSubs analz_subset_parts]) 4 1))
6.22
6.23 fun spy_analz_tac (cs,ss) i =
6.24 DETERM
7.1 --- a/doc-src/TutorialI/Protocol/NS_Public.thy Sat Oct 17 01:05:59 2009 +0200
7.2 +++ b/doc-src/TutorialI/Protocol/NS_Public.thy Sat Oct 17 14:43:18 2009 +0200
7.3 @@ -221,8 +221,8 @@
7.4 lemma A_trusts_NS2_lemma [rule_format]:
7.5 "\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
7.6 \<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
7.7 - Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
7.8 - Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
7.9 + Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
7.10 + Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
7.11 apply (erule ns_public.induct, simp_all)
7.12 (*Fake, NS1*)
7.13 apply (blast dest: Spy_not_see_NA)+
7.14 @@ -240,8 +240,8 @@
7.15 lemma B_trusts_NS1 [rule_format]:
7.16 "evs \<in> ns_public
7.17 \<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
7.18 - Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
7.19 - Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
7.20 + Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
7.21 + Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
7.22 apply (erule ns_public.induct, simp_all)
7.23 (*Fake*)
7.24 apply (blast intro!: analz_insertI)
8.1 --- a/doc-src/TutorialI/Protocol/Public.thy Sat Oct 17 01:05:59 2009 +0200
8.2 +++ b/doc-src/TutorialI/Protocol/Public.thy Sat Oct 17 14:43:18 2009 +0200
8.3 @@ -25,11 +25,11 @@
8.4 primrec
8.5 (*Agents know their private key and all public keys*)
8.6 initState_Server: "initState Server =
8.7 - insert (Key (priK Server)) (Key ` range pubK)"
8.8 + insert (Key (priK Server)) (Key ` range pubK)"
8.9 initState_Friend: "initState (Friend i) =
8.10 - insert (Key (priK (Friend i))) (Key ` range pubK)"
8.11 + insert (Key (priK (Friend i))) (Key ` range pubK)"
8.12 initState_Spy: "initState Spy =
8.13 - (Key`invKey`pubK`bad) Un (Key ` range pubK)"
8.14 + (Key`invKey`pubK`bad) Un (Key ` range pubK)"
8.15 (*>*)
8.16
8.17 text {*
9.1 --- a/doc-src/TutorialI/Rules/Basic.thy Sat Oct 17 01:05:59 2009 +0200
9.2 +++ b/doc-src/TutorialI/Rules/Basic.thy Sat Oct 17 14:43:18 2009 +0200
9.3 @@ -1,4 +1,3 @@
9.4 -(* ID: $Id$ *)
9.5 theory Basic imports Main begin
9.6
9.7 lemma conj_rule: "\<lbrakk> P; Q \<rbrakk> \<Longrightarrow> P \<and> (Q \<and> P)"
9.8 @@ -149,9 +148,9 @@
9.9
9.10 lemma "\<lbrakk>\<not>(P\<longrightarrow>Q); \<not>(R\<longrightarrow>Q)\<rbrakk> \<Longrightarrow> R"
9.11 apply (erule_tac Q="R\<longrightarrow>Q" in contrapos_np)
9.12 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.13 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.14 apply (intro impI)
9.15 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.16 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.17 by (erule notE)
9.18
9.19 text {*
9.20 @@ -161,11 +160,11 @@
9.21
9.22 lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
9.23 apply (intro disjCI conjI)
9.24 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.25 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.26
9.27 apply (elim conjE disjE)
9.28 apply assumption
9.29 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.30 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.31
9.32 by (erule contrapos_np, rule conjI)
9.33 text{*
9.34 @@ -241,18 +240,18 @@
9.35 text{*rename_tac*}
9.36 lemma "x < y \<Longrightarrow> \<forall>x y. P x (f y)"
9.37 apply (intro allI)
9.38 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.39 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.40 apply (rename_tac v w)
9.41 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.42 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.43 oops
9.44
9.45
9.46 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
9.47 apply (frule spec)
9.48 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.49 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.50 apply (drule mp, assumption)
9.51 apply (drule spec)
9.52 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.53 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.54 by (drule mp)
9.55
9.56 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
9.57 @@ -276,11 +275,11 @@
9.58
9.59 lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
9.60 apply (frule spec)
9.61 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.62 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.63 apply (drule mp, assumption)
9.64 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.65 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.66 apply (drule_tac x = "h a" in spec)
9.67 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.68 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.69 by (drule mp)
9.70
9.71 text {*
9.72 @@ -290,15 +289,15 @@
9.73
9.74 lemma mult_dvd_mono: "\<lbrakk>i dvd m; j dvd n\<rbrakk> \<Longrightarrow> i*j dvd (m*n :: nat)"
9.75 apply (simp add: dvd_def)
9.76 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.77 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.78 apply (erule exE)
9.79 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.80 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.81 apply (erule exE)
9.82 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.83 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.84 apply (rename_tac l)
9.85 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.86 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.87 apply (rule_tac x="k*l" in exI)
9.88 - --{* @{subgoals[display,indent=0,margin=65]} *}
9.89 + --{* @{subgoals[display,indent=0,margin=65]} *}
9.90 apply simp
9.91 done
9.92
10.1 --- a/src/FOL/FOL.thy Sat Oct 17 01:05:59 2009 +0200
10.2 +++ b/src/FOL/FOL.thy Sat Oct 17 14:43:18 2009 +0200
10.3 @@ -174,13 +174,13 @@
10.4 structure Blast = Blast
10.5 (
10.6 val thy = @{theory}
10.7 - type claset = Cla.claset
10.8 + type claset = Cla.claset
10.9 val equality_name = @{const_name "op ="}
10.10 val not_name = @{const_name Not}
10.11 - val notE = @{thm notE}
10.12 - val ccontr = @{thm ccontr}
10.13 + val notE = @{thm notE}
10.14 + val ccontr = @{thm ccontr}
10.15 val contr_tac = Cla.contr_tac
10.16 - val dup_intr = Cla.dup_intr
10.17 + val dup_intr = Cla.dup_intr
10.18 val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
10.19 val rep_cs = Cla.rep_cs
10.20 val cla_modifiers = Cla.cla_modifiers
11.1 --- a/src/FOL/ex/Classical.thy Sat Oct 17 01:05:59 2009 +0200
11.2 +++ b/src/FOL/ex/Classical.thy Sat Oct 17 14:43:18 2009 +0200
11.3 @@ -418,7 +418,7 @@
11.4 by fast
11.5
11.6 text{*Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
11.7 - author U. Egly*}
11.8 + author U. Egly*}
11.9 lemma "((\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z)))) -->
11.10 (\<exists>w. C(w) & (\<forall>y. C(y) --> (\<forall>z. D(w,y,z)))))
11.11 &
12.1 --- a/src/HOL/Algebra/Coset.thy Sat Oct 17 01:05:59 2009 +0200
12.2 +++ b/src/HOL/Algebra/Coset.thy Sat Oct 17 14:43:18 2009 +0200
12.3 @@ -609,7 +609,7 @@
12.4 proof (simp add: r_congruent_def sym_def, clarify)
12.5 fix x y
12.6 assume [simp]: "x \<in> carrier G" "y \<in> carrier G"
12.7 - and "inv x \<otimes> y \<in> H"
12.8 + and "inv x \<otimes> y \<in> H"
12.9 hence "inv (inv x \<otimes> y) \<in> H" by (simp add: m_inv_closed)
12.10 thus "inv y \<otimes> x \<in> H" by (simp add: inv_mult_group)
12.11 qed
12.12 @@ -618,10 +618,10 @@
12.13 proof (simp add: r_congruent_def trans_def, clarify)
12.14 fix x y z
12.15 assume [simp]: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
12.16 - and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
12.17 + and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
12.18 hence "(inv x \<otimes> y) \<otimes> (inv y \<otimes> z) \<in> H" by simp
12.19 hence "inv x \<otimes> (y \<otimes> inv y) \<otimes> z \<in> H"
12.20 - by (simp add: m_assoc del: r_inv Units_r_inv)
12.21 + by (simp add: m_assoc del: r_inv Units_r_inv)
12.22 thus "inv x \<otimes> z \<in> H" by simp
12.23 qed
12.24 qed
13.1 --- a/src/HOL/Algebra/Divisibility.thy Sat Oct 17 01:05:59 2009 +0200
13.2 +++ b/src/HOL/Algebra/Divisibility.thy Sat Oct 17 14:43:18 2009 +0200
13.3 @@ -2579,33 +2579,33 @@
13.4 by force
13.5
13.6 from this obtain p'
13.7 - where "p' \<in> set (as@bs)"
13.8 - and pp': "p \<sim> p'" by auto
13.9 + where "p' \<in> set (as@bs)"
13.10 + and pp': "p \<sim> p'" by auto
13.11
13.12 hence "p' \<in> set as \<or> p' \<in> set bs" by simp
13.13 moreover
13.14 {
13.15 - assume p'elem: "p' \<in> set as"
13.16 - with ascarr have [simp]: "p' \<in> carrier G" by fast
13.17 -
13.18 - note pp'
13.19 - also from afac p'elem
13.20 - have "p' divides a" by (rule factors_dividesI) fact+
13.21 - finally
13.22 - have "p divides a" by simp
13.23 + assume p'elem: "p' \<in> set as"
13.24 + with ascarr have [simp]: "p' \<in> carrier G" by fast
13.25 +
13.26 + note pp'
13.27 + also from afac p'elem
13.28 + have "p' divides a" by (rule factors_dividesI) fact+
13.29 + finally
13.30 + have "p divides a" by simp
13.31 }
13.32 moreover
13.33 {
13.34 - assume p'elem: "p' \<in> set bs"
13.35 - with bscarr have [simp]: "p' \<in> carrier G" by fast
13.36 -
13.37 - note pp'
13.38 - also from bfac
13.39 - have "p' divides b" by (rule factors_dividesI) fact+
13.40 - finally have "p divides b" by simp
13.41 + assume p'elem: "p' \<in> set bs"
13.42 + with bscarr have [simp]: "p' \<in> carrier G" by fast
13.43 +
13.44 + note pp'
13.45 + also from bfac
13.46 + have "p' divides b" by (rule factors_dividesI) fact+
13.47 + finally have "p divides b" by simp
13.48 }
13.49 ultimately
13.50 - show "p divides a \<or> p divides b" by fast
13.51 + show "p divides a \<or> p divides b" by fast
13.52 qed
13.53 qed
13.54 qed
13.55 @@ -3176,7 +3176,7 @@
13.56 have "c = c \<otimes> \<one>" by simp
13.57 also from abrelprime[symmetric]
13.58 have "\<dots> \<sim> c \<otimes> somegcd G a b"
13.59 - by (rule assoc_subst) (simp add: mult_cong_r)+
13.60 + by (rule assoc_subst) (simp add: mult_cong_r)+
13.61 also have "\<dots> \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by (rule gcd_mult) fact+
13.62 finally
13.63 have c: "c \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by simp
13.64 @@ -3188,13 +3188,13 @@
13.65 have "somegcd G a (b \<otimes> c) \<sim> somegcd G a (c \<otimes> b)" by (simp add: m_comm)
13.66 also from a
13.67 have "\<dots> \<sim> somegcd G (somegcd G a (c \<otimes> a)) (c \<otimes> b)"
13.68 - by (rule assoc_subst) (simp add: gcd_cong_l)+
13.69 + by (rule assoc_subst) (simp add: gcd_cong_l)+
13.70 also from gcd_assoc
13.71 have "\<dots> \<sim> somegcd G a (somegcd G (c \<otimes> a) (c \<otimes> b))"
13.72 by (rule assoc_subst) simp+
13.73 also from c[symmetric]
13.74 have "\<dots> \<sim> somegcd G a c"
13.75 - by (rule assoc_subst) (simp add: gcd_cong_r)+
13.76 + by (rule assoc_subst) (simp add: gcd_cong_r)+
13.77 also note acrelprime
13.78 finally
13.79 show "somegcd G a (b \<otimes> c) \<sim> \<one>" by simp
14.1 --- a/src/HOL/Algebra/FiniteProduct.thy Sat Oct 17 01:05:59 2009 +0200
14.2 +++ b/src/HOL/Algebra/FiniteProduct.thy Sat Oct 17 14:43:18 2009 +0200
14.3 @@ -423,17 +423,17 @@
14.4 then have "finprod G f A = finprod G f (insert x B)" by simp
14.5 also from insert have "... = f x \<otimes> finprod G f B"
14.6 proof (intro finprod_insert)
14.7 - show "finite B" by fact
14.8 + show "finite B" by fact
14.9 next
14.10 - show "x ~: B" by fact
14.11 + show "x ~: B" by fact
14.12 next
14.13 - assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
14.14 - "g \<in> insert x B \<rightarrow> carrier G"
14.15 - thus "f \<in> B -> carrier G" by fastsimp
14.16 + assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
14.17 + "g \<in> insert x B \<rightarrow> carrier G"
14.18 + thus "f \<in> B -> carrier G" by fastsimp
14.19 next
14.20 - assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
14.21 - "g \<in> insert x B \<rightarrow> carrier G"
14.22 - thus "f x \<in> carrier G" by fastsimp
14.23 + assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
14.24 + "g \<in> insert x B \<rightarrow> carrier G"
14.25 + thus "f x \<in> carrier G" by fastsimp
14.26 qed
14.27 also from insert have "... = g x \<otimes> finprod G g B" by fastsimp
14.28 also from insert have "... = finprod G g (insert x B)"
15.1 --- a/src/HOL/Algebra/Group.thy Sat Oct 17 01:05:59 2009 +0200
15.2 +++ b/src/HOL/Algebra/Group.thy Sat Oct 17 14:43:18 2009 +0200
15.3 @@ -785,16 +785,16 @@
15.4 assume H: "H \<in> A"
15.5 with L have subgroupH: "subgroup H G" by auto
15.6 from subgroupH have groupH: "group (G (| carrier := H |))" (is "group ?H")
15.7 - by (rule subgroup_imp_group)
15.8 + by (rule subgroup_imp_group)
15.9 from groupH have monoidH: "monoid ?H"
15.10 - by (rule group.is_monoid)
15.11 + by (rule group.is_monoid)
15.12 from H have Int_subset: "?Int \<subseteq> H" by fastsimp
15.13 then show "le ?L ?Int H" by simp
15.14 next
15.15 fix H
15.16 assume H: "H \<in> Lower ?L A"
15.17 with L Int_subgroup show "le ?L H ?Int"
15.18 - by (fastsimp simp: Lower_def intro: Inter_greatest)
15.19 + by (fastsimp simp: Lower_def intro: Inter_greatest)
15.20 next
15.21 show "A \<subseteq> carrier ?L" by (rule L)
15.22 next
16.1 --- a/src/HOL/Algebra/Lattice.thy Sat Oct 17 01:05:59 2009 +0200
16.2 +++ b/src/HOL/Algebra/Lattice.thy Sat Oct 17 14:43:18 2009 +0200
16.3 @@ -533,22 +533,22 @@
16.4 assume y: "y \<in> Upper L (insert x A)"
16.5 show "s \<sqsubseteq> y"
16.6 proof (rule least_le [OF least_s], rule Upper_memI)
16.7 - fix z
16.8 - assume z: "z \<in> {a, x}"
16.9 - then show "z \<sqsubseteq> y"
16.10 - proof
16.11 + fix z
16.12 + assume z: "z \<in> {a, x}"
16.13 + then show "z \<sqsubseteq> y"
16.14 + proof
16.15 have y': "y \<in> Upper L A"
16.16 apply (rule subsetD [where A = "Upper L (insert x A)"])
16.17 apply (rule Upper_antimono)
16.18 - apply blast
16.19 - apply (rule y)
16.20 + apply blast
16.21 + apply (rule y)
16.22 done
16.23 assume "z = a"
16.24 with y' least_a show ?thesis by (fast dest: least_le)
16.25 - next
16.26 - assume "z \<in> {x}" (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
16.27 + next
16.28 + assume "z \<in> {x}" (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
16.29 with y L show ?thesis by blast
16.30 - qed
16.31 + qed
16.32 qed (rule Upper_closed [THEN subsetD, OF y])
16.33 next
16.34 from L show "insert x A \<subseteq> carrier L" by simp
16.35 @@ -569,9 +569,9 @@
16.36 case True
16.37 with insert show ?thesis
16.38 by simp (simp add: least_cong [OF weak_sup_of_singleton]
16.39 - sup_of_singleton_closed sup_of_singletonI)
16.40 - (* The above step is hairy; least_cong can make simp loop.
16.41 - Would want special version of simp to apply least_cong. *)
16.42 + sup_of_singleton_closed sup_of_singletonI)
16.43 + (* The above step is hairy; least_cong can make simp loop.
16.44 + Would want special version of simp to apply least_cong. *)
16.45 next
16.46 case False
16.47 with insert have "least L (\<Squnion>A) (Upper L A)" by simp
16.48 @@ -774,22 +774,22 @@
16.49 assume y: "y \<in> Lower L (insert x A)"
16.50 show "y \<sqsubseteq> i"
16.51 proof (rule greatest_le [OF greatest_i], rule Lower_memI)
16.52 - fix z
16.53 - assume z: "z \<in> {a, x}"
16.54 - then show "y \<sqsubseteq> z"
16.55 - proof
16.56 + fix z
16.57 + assume z: "z \<in> {a, x}"
16.58 + then show "y \<sqsubseteq> z"
16.59 + proof
16.60 have y': "y \<in> Lower L A"
16.61 apply (rule subsetD [where A = "Lower L (insert x A)"])
16.62 apply (rule Lower_antimono)
16.63 - apply blast
16.64 - apply (rule y)
16.65 + apply blast
16.66 + apply (rule y)
16.67 done
16.68 assume "z = a"
16.69 with y' greatest_a show ?thesis by (fast dest: greatest_le)
16.70 - next
16.71 + next
16.72 assume "z \<in> {x}"
16.73 with y L show ?thesis by blast
16.74 - qed
16.75 + qed
16.76 qed (rule Lower_closed [THEN subsetD, OF y])
16.77 next
16.78 from L show "insert x A \<subseteq> carrier L" by simp
16.79 @@ -809,7 +809,7 @@
16.80 case True
16.81 with insert show ?thesis
16.82 by simp (simp add: greatest_cong [OF weak_inf_of_singleton]
16.83 - inf_of_singleton_closed inf_of_singletonI)
16.84 + inf_of_singleton_closed inf_of_singletonI)
16.85 next
16.86 case False
16.87 from insert show ?thesis
16.88 @@ -1291,7 +1291,7 @@
16.89 proof
16.90 from B show "greatest ?L (\<Inter> B \<inter> A) (Lower ?L B)"
16.91 txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
16.92 - @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
16.93 + @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
16.94 by (fastsimp intro!: greatest_LowerI simp: Lower_def)
16.95 qed
16.96 qed
17.1 --- a/src/HOL/Algebra/UnivPoly.thy Sat Oct 17 01:05:59 2009 +0200
17.2 +++ b/src/HOL/Algebra/UnivPoly.thy Sat Oct 17 14:43:18 2009 +0200
17.3 @@ -349,19 +349,19 @@
17.4 fix nn assume Succ: "n = Suc nn"
17.5 have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)"
17.6 proof -
17.7 - have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
17.8 - also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
17.9 - using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
17.10 - also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
17.11 - proof -
17.12 - have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
17.13 - using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R
17.14 - unfolding Pi_def by simp
17.15 - also have "\<dots> = \<zero>" by simp
17.16 - finally show ?thesis using r_zero R by simp
17.17 - qed
17.18 - also have "\<dots> = coeff P p (Suc nn)" using R by simp
17.19 - finally show ?thesis by simp
17.20 + have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = (\<Oplus>i\<in>{..Suc nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" using R by simp
17.21 + also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>) \<oplus> (\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))"
17.22 + using finsum_Suc [of "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "nn"] unfolding Pi_def using R by simp
17.23 + also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
17.24 + proof -
17.25 + have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
17.26 + using finsum_cong [of "{..nn}" "{..nn}" "(\<lambda>i::nat. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>))" "(\<lambda>i::nat. \<zero>)"] using R
17.27 + unfolding Pi_def by simp
17.28 + also have "\<dots> = \<zero>" by simp
17.29 + finally show ?thesis using r_zero R by simp
17.30 + qed
17.31 + also have "\<dots> = coeff P p (Suc nn)" using R by simp
17.32 + finally show ?thesis by simp
17.33 qed
17.34 then show ?thesis using Succ by simp
17.35 }
17.36 @@ -627,11 +627,11 @@
17.37 then show "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k)"
17.38 proof -
17.39 have lhs: "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
17.40 - unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
17.41 + unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
17.42 note cl = monom_closed [OF R.one_closed, of 1]
17.43 note clk = monom_closed [OF R.one_closed, of k]
17.44 have rhs: "monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> k \<otimes>\<^bsub>P\<^esub> monom P \<one> 1"
17.45 - unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc [OF cl clk cl]] ..
17.46 + unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc [OF cl clk cl]] ..
17.47 from lhs rhs show ?thesis by simp
17.48 qed
17.49 }
17.50 @@ -670,25 +670,25 @@
17.51 case True
17.52 {
17.53 show ?thesis
17.54 - unfolding True [symmetric]
17.55 - coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"]
17.56 - coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
17.57 - using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))"
17.58 - "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
17.59 - a_in_R b_in_R
17.60 - unfolding simp_implies_def
17.61 - using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
17.62 - unfolding Pi_def by auto
17.63 + unfolding True [symmetric]
17.64 + coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"]
17.65 + coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
17.66 + using R.finsum_cong [of "{.. n + m}" "{.. n + m}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = n + m - i then b else \<zero>))"
17.67 + "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
17.68 + a_in_R b_in_R
17.69 + unfolding simp_implies_def
17.70 + using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
17.71 + unfolding Pi_def by auto
17.72 }
17.73 next
17.74 case False
17.75 {
17.76 show ?thesis
17.77 - unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
17.78 - unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
17.79 - unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
17.80 - using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
17.81 - unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
17.82 + unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
17.83 + unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
17.84 + unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
17.85 + using R.finsum_cong [of "{..k}" "{..k}" "(\<lambda>i. (if n = i then a else \<zero>) \<otimes> (if m = k - i then b else \<zero>))" "(\<lambda>i. \<zero>)"]
17.86 + unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
17.87 }
17.88 qed
17.89 qed (simp_all add: a_in_R b_in_R)
17.90 @@ -1517,7 +1517,7 @@
17.91 then have max_sl: "max (deg R p) (deg R q) < m" by simp
17.92 then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith
17.93 with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m]
17.94 - using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp
17.95 + using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp
17.96 qed
17.97 qed (simp add: p_in_P q_in_P)
17.98 moreover have deg_ne: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r"
17.99 @@ -1582,114 +1582,114 @@
17.100 (*JE: we now apply the induction hypothesis with some additional facts required*)
17.101 from f_in_P deg_g_le_deg_f show ?thesis
17.102 proof (induct n \<equiv> "deg R f" arbitrary: "f" rule: nat_less_induct)
17.103 - fix n f
17.104 - assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
17.105 + fix n f
17.106 + assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
17.107 deg R g \<le> deg R x \<longrightarrow>
17.108 - m = deg R x \<longrightarrow>
17.109 - (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
17.110 - and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
17.111 - and deg_g_le_deg_f: "deg R g \<le> deg R f"
17.112 - let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
17.113 - and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
17.114 - show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
17.115 - proof -
17.116 - (*JE: we first extablish the existence of a triple satisfying the previous equation.
17.117 - Then we will have to prove the second part of the predicate.*)
17.118 - have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
17.119 - using minus_add
17.120 - using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
17.121 - using r_neg by auto
17.122 - show ?thesis
17.123 - proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
17.124 - (*JE: if the degree of the remainder satisfies the statement property we are done*)
17.125 - case True
17.126 - {
17.127 - show ?thesis
17.128 - proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
17.129 - show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
17.130 - show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
17.131 - qed (simp_all)
17.132 - }
17.133 - next
17.134 - case False note n_deg_r_l_deg_g = False
17.135 - {
17.136 - (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
17.137 - show ?thesis
17.138 - proof (cases "deg R f = 0")
17.139 - (*JE: the solutions are different if the degree of f is zero or not*)
17.140 - case True
17.141 - {
17.142 - have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
17.143 - have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
17.144 - unfolding deg_g apply simp
17.145 - unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
17.146 - using deg_zero_impl_monom [OF g_in_P deg_g] by simp
17.147 - then show ?thesis using f_in_P by blast
17.148 - }
17.149 - next
17.150 - case False note deg_f_nzero = False
17.151 - {
17.152 - (*JE: now it only remains the case where the induction hypothesis can be used.*)
17.153 - (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
17.154 - have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
17.155 - proof -
17.156 - have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
17.157 - also have "\<dots> < deg R f"
17.158 - proof (rule deg_lcoeff_cancel)
17.159 - show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
17.160 - using deg_smult_ring [of "lcoeff g" f] using prem
17.161 - using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
17.162 - show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
17.163 - using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
17.164 - by simp
17.165 - show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
17.166 - unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
17.167 - unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
17.168 - using R.finsum_cong' [of "{..deg R f}" "{..deg R f}"
17.169 - "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))"
17.170 - "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
17.171 - using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
17.172 - unfolding Pi_def using deg_g_le_deg_f by force
17.173 - qed (simp_all add: deg_f_nzero)
17.174 - finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
17.175 - qed
17.176 - moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
17.177 - moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
17.178 - moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
17.179 - (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
17.180 - ultimately obtain q' r' k'
17.181 - where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
17.182 - and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
17.183 - using hypo by blast
17.184 - (*JE: we now prove that the new quotient, remainder and exponent can be used to get
17.185 - the quotient, remainder and exponent of the long division theorem*)
17.186 - show ?thesis
17.187 - proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
17.188 - show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
17.189 - proof -
17.190 - have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)"
17.191 - using smult_assoc1 exist by simp
17.192 - also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
17.193 - using UP_smult_r_distr by simp
17.194 - also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
17.195 - using rem_desc by simp
17.196 - also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
17.197 - using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
17.198 - using q'_in_carrier r'_in_carrier by simp
17.199 - also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
17.200 - using q'_in_carrier by (auto simp add: m_comm)
17.201 - also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
17.202 - using smult_assoc2 q'_in_carrier by auto
17.203 - also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
17.204 - using sym [OF l_distr] and q'_in_carrier by auto
17.205 - finally show ?thesis using m_comm q'_in_carrier by auto
17.206 - qed
17.207 - qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
17.208 - }
17.209 - qed
17.210 - }
17.211 - qed
17.212 - qed
17.213 + m = deg R x \<longrightarrow>
17.214 + (\<exists>q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> x = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g))"
17.215 + and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
17.216 + and deg_g_le_deg_f: "deg R g \<le> deg R f"
17.217 + let ?k = "1::nat" and ?r = "(g \<otimes>\<^bsub>P\<^esub> (monom P (lcoeff f) (deg R f - deg R g))) \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)"
17.218 + and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
17.219 + show "\<exists> q r (k::nat). q \<in> carrier P \<and> r \<in> carrier P \<and> lcoeff g (^) k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> q \<oplus>\<^bsub>P\<^esub> r & (r = \<zero>\<^bsub>P\<^esub> | deg R r < deg R g)"
17.220 + proof -
17.221 + (*JE: we first extablish the existence of a triple satisfying the previous equation.
17.222 + Then we will have to prove the second part of the predicate.*)
17.223 + have exist: "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r"
17.224 + using minus_add
17.225 + using sym [OF a_assoc [of "g \<otimes>\<^bsub>P\<^esub> ?q" "\<ominus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "lcoeff g \<odot>\<^bsub>P\<^esub> f"]]
17.226 + using r_neg by auto
17.227 + show ?thesis
17.228 + proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
17.229 + (*JE: if the degree of the remainder satisfies the statement property we are done*)
17.230 + case True
17.231 + {
17.232 + show ?thesis
17.233 + proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
17.234 + show "lcoeff g (^) ?k \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r" using exist by simp
17.235 + show "\<ominus>\<^bsub>P\<^esub> ?r = \<zero>\<^bsub>P\<^esub> \<or> deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g" using True by simp
17.236 + qed (simp_all)
17.237 + }
17.238 + next
17.239 + case False note n_deg_r_l_deg_g = False
17.240 + {
17.241 + (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
17.242 + show ?thesis
17.243 + proof (cases "deg R f = 0")
17.244 + (*JE: the solutions are different if the degree of f is zero or not*)
17.245 + case True
17.246 + {
17.247 + have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
17.248 + have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
17.249 + unfolding deg_g apply simp
17.250 + unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
17.251 + using deg_zero_impl_monom [OF g_in_P deg_g] by simp
17.252 + then show ?thesis using f_in_P by blast
17.253 + }
17.254 + next
17.255 + case False note deg_f_nzero = False
17.256 + {
17.257 + (*JE: now it only remains the case where the induction hypothesis can be used.*)
17.258 + (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
17.259 + have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
17.260 + proof -
17.261 + have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
17.262 + also have "\<dots> < deg R f"
17.263 + proof (rule deg_lcoeff_cancel)
17.264 + show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
17.265 + using deg_smult_ring [of "lcoeff g" f] using prem
17.266 + using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
17.267 + show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
17.268 + using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
17.269 + by simp
17.270 + show "coeff P (g \<otimes>\<^bsub>P\<^esub> ?q) (deg R f) = \<ominus> coeff P (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) (deg R f)"
17.271 + unfolding coeff_mult [OF g_in_P monom_closed [OF lcoeff_closed [OF f_in_P], of "deg R f - deg R g"], of "deg R f"]
17.272 + unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
17.273 + using R.finsum_cong' [of "{..deg R f}" "{..deg R f}"
17.274 + "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))"
17.275 + "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
17.276 + using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
17.277 + unfolding Pi_def using deg_g_le_deg_f by force
17.278 + qed (simp_all add: deg_f_nzero)
17.279 + finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
17.280 + qed
17.281 + moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
17.282 + moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
17.283 + moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
17.284 + (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
17.285 + ultimately obtain q' r' k'
17.286 + where rem_desc: "lcoeff g (^) (k'::nat) \<odot>\<^bsub>P\<^esub> (\<ominus>\<^bsub>P\<^esub> ?r) = g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"and rem_deg: "(r' = \<zero>\<^bsub>P\<^esub> \<or> deg R r' < deg R g)"
17.287 + and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
17.288 + using hypo by blast
17.289 + (*JE: we now prove that the new quotient, remainder and exponent can be used to get
17.290 + the quotient, remainder and exponent of the long division theorem*)
17.291 + show ?thesis
17.292 + proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
17.293 + show "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<oplus>\<^bsub>P\<^esub> r'"
17.294 + proof -
17.295 + have "(lcoeff g (^) (Suc k')) \<odot>\<^bsub>P\<^esub> f = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> \<ominus>\<^bsub>P\<^esub> ?r)"
17.296 + using smult_assoc1 exist by simp
17.297 + also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ( \<ominus>\<^bsub>P\<^esub> ?r))"
17.298 + using UP_smult_r_distr by simp
17.299 + also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r')"
17.300 + using rem_desc by simp
17.301 + also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q) \<oplus>\<^bsub>P\<^esub> g \<otimes>\<^bsub>P\<^esub> q' \<oplus>\<^bsub>P\<^esub> r'"
17.302 + using sym [OF a_assoc [of "lcoeff g (^) k' \<odot>\<^bsub>P\<^esub> (g \<otimes>\<^bsub>P\<^esub> ?q)" "g \<otimes>\<^bsub>P\<^esub> q'" "r'"]]
17.303 + using q'_in_carrier r'_in_carrier by simp
17.304 + also have "\<dots> = (lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> (?q \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
17.305 + using q'_in_carrier by (auto simp add: m_comm)
17.306 + also have "\<dots> = (((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q) \<otimes>\<^bsub>P\<^esub> g) \<oplus>\<^bsub>P\<^esub> q' \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
17.307 + using smult_assoc2 q'_in_carrier by auto
17.308 + also have "\<dots> = ((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q') \<otimes>\<^bsub>P\<^esub> g \<oplus>\<^bsub>P\<^esub> r'"
17.309 + using sym [OF l_distr] and q'_in_carrier by auto
17.310 + finally show ?thesis using m_comm q'_in_carrier by auto
17.311 + qed
17.312 + qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
17.313 + }
17.314 + qed
17.315 + }
17.316 + qed
17.317 + qed
17.318 qed
17.319 }
17.320 qed
18.1 --- a/src/HOL/Algebra/poly/UnivPoly2.thy Sat Oct 17 01:05:59 2009 +0200
18.2 +++ b/src/HOL/Algebra/poly/UnivPoly2.thy Sat Oct 17 14:43:18 2009 +0200
18.3 @@ -100,7 +100,7 @@
18.4 begin
18.5
18.6 definition
18.7 - up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
18.8 + up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
18.9
18.10 instance ..
18.11
18.12 @@ -133,7 +133,7 @@
18.13
18.14 definition
18.15 up_mult_def: "p * q = Abs_UP (%n::nat. setsum
18.16 - (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
18.17 + (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
18.18
18.19 instance ..
18.20
18.21 @@ -201,18 +201,18 @@
18.22 have "(%i. f i + g i) : UP"
18.23 proof -
18.24 from fup obtain n where boundn: "bound n f"
18.25 - by (unfold UP_def) fast
18.26 + by (unfold UP_def) fast
18.27 from gup obtain m where boundm: "bound m g"
18.28 - by (unfold UP_def) fast
18.29 + by (unfold UP_def) fast
18.30 have "bound (max n m) (%i. (f i + g i))"
18.31 proof
18.32 - fix i
18.33 - assume "max n m < i"
18.34 - with boundn and boundm show "f i + g i = 0"
18.35 + fix i
18.36 + assume "max n m < i"
18.37 + with boundn and boundm show "f i + g i = 0"
18.38 by (fastsimp simp add: algebra_simps)
18.39 qed
18.40 then show "(%i. (f i + g i)) : UP"
18.41 - by (unfold UP_def) fast
18.42 + by (unfold UP_def) fast
18.43 qed
18.44 }
18.45 then show ?thesis
18.46 @@ -228,30 +228,30 @@
18.47 have "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
18.48 proof -
18.49 from fup obtain n where "bound n f"
18.50 - by (unfold UP_def) fast
18.51 + by (unfold UP_def) fast
18.52 from gup obtain m where "bound m g"
18.53 - by (unfold UP_def) fast
18.54 + by (unfold UP_def) fast
18.55 have "bound (n + m) (%n. setsum (%i. f i * g (n-i)) {..n})"
18.56 proof
18.57 - fix k
18.58 - assume bound: "n + m < k"
18.59 - {
18.60 - fix i
18.61 - have "f i * g (k-i) = 0"
18.62 - proof cases
18.63 - assume "n < i"
18.64 - with `bound n f` show ?thesis by (auto simp add: algebra_simps)
18.65 - next
18.66 - assume "~ (n < i)"
18.67 - with bound have "m < k-i" by arith
18.68 - with `bound m g` show ?thesis by (auto simp add: algebra_simps)
18.69 - qed
18.70 - }
18.71 - then show "setsum (%i. f i * g (k-i)) {..k} = 0"
18.72 - by (simp add: algebra_simps)
18.73 + fix k
18.74 + assume bound: "n + m < k"
18.75 + {
18.76 + fix i
18.77 + have "f i * g (k-i) = 0"
18.78 + proof cases
18.79 + assume "n < i"
18.80 + with `bound n f` show ?thesis by (auto simp add: algebra_simps)
18.81 + next
18.82 + assume "~ (n < i)"
18.83 + with bound have "m < k-i" by arith
18.84 + with `bound m g` show ?thesis by (auto simp add: algebra_simps)
18.85 + qed
18.86 + }
18.87 + then show "setsum (%i. f i * g (k-i)) {..k} = 0"
18.88 + by (simp add: algebra_simps)
18.89 qed
18.90 then show "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
18.91 - by (unfold UP_def) fast
18.92 + by (unfold UP_def) fast
18.93 qed
18.94 }
18.95 then show ?thesis
18.96 @@ -290,17 +290,17 @@
18.97 {
18.98 fix k and a b c :: "nat=>'a::ring"
18.99 have "k <= n ==>
18.100 - setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} =
18.101 - setsum (%j. a j * setsum (%i. b i * c (n-j-i)) {..k-j}) {..k}"
18.102 - (is "_ ==> ?eq k")
18.103 + setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} =
18.104 + setsum (%j. a j * setsum (%i. b i * c (n-j-i)) {..k-j}) {..k}"
18.105 + (is "_ ==> ?eq k")
18.106 proof (induct k)
18.107 - case 0 show ?case by simp
18.108 + case 0 show ?case by simp
18.109 next
18.110 - case (Suc k)
18.111 - then have "k <= n" by arith
18.112 - then have "?eq k" by (rule Suc)
18.113 - then show ?case
18.114 - by (simp add: Suc_diff_le natsum_ldistr)
18.115 + case (Suc k)
18.116 + then have "k <= n" by arith
18.117 + then have "?eq k" by (rule Suc)
18.118 + then show ?case
18.119 + by (simp add: Suc_diff_le natsum_ldistr)
18.120 qed
18.121 }
18.122 then show "coeff ((p * q) * r) n = coeff (p * (q * r)) n"
18.123 @@ -326,13 +326,13 @@
18.124 fix k
18.125 fix a b :: "nat=>'a::ring"
18.126 have "k <= n ==>
18.127 - setsum (%i. a i * b (n-i)) {..k} =
18.128 - setsum (%i. a (k-i) * b (i+n-k)) {..k}"
18.129 - (is "_ ==> ?eq k")
18.130 + setsum (%i. a i * b (n-i)) {..k} =
18.131 + setsum (%i. a (k-i) * b (i+n-k)) {..k}"
18.132 + (is "_ ==> ?eq k")
18.133 proof (induct k)
18.134 - case 0 show ?case by simp
18.135 + case 0 show ?case by simp
18.136 next
18.137 - case (Suc k) then show ?case by (subst natsum_Suc2) simp
18.138 + case (Suc k) then show ?case by (subst natsum_Suc2) simp
18.139 qed
18.140 }
18.141 then show "coeff (p * q) n = coeff (q * p) n"
19.1 --- a/src/HOL/Auth/CertifiedEmail.thy Sat Oct 17 01:05:59 2009 +0200
19.2 +++ b/src/HOL/Auth/CertifiedEmail.thy Sat Oct 17 14:43:18 2009 +0200
19.3 @@ -43,7 +43,7 @@
19.4
19.5 | FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
19.6 equipped with the necessary credentials to serve as an SSL server.*}
19.7 - "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
19.8 + "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
19.9 ==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
19.10
19.11 | CM1: --{*The sender approaches the recipient. The message is a number.*}
19.12 @@ -54,14 +54,14 @@
19.13 hs = Hash{|Number cleartext, Nonce q, response S R q, Crypt K (Number m)|};
19.14 S2TTP = Crypt(pubEK TTP) {|Agent S, Number BothAuth, Key K, Agent R, hs|}|]
19.15 ==> Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number BothAuth,
19.16 - Number cleartext, Nonce q, S2TTP|} # evs1
19.17 - \<in> certified_mail"
19.18 + Number cleartext, Nonce q, S2TTP|} # evs1
19.19 + \<in> certified_mail"
19.20
19.21 | CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
19.22 password to @{term TTP} over an SSL channel.*}
19.23 "[|evs2 \<in> certified_mail;
19.24 Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext,
19.25 - Nonce q, S2TTP|} \<in> set evs2;
19.26 + Nonce q, S2TTP|} \<in> set evs2;
19.27 TTP \<noteq> R;
19.28 hr = Hash {|Number cleartext, Nonce q, response S R q, em|} |]
19.29 ==>
19.30 @@ -289,14 +289,14 @@
19.31 "evs \<in> certified_mail ==>
19.32 Key K \<notin> analz (spies evs) -->
19.33 (\<forall>AO. Crypt (pubEK TTP)
19.34 - {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
19.35 + {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
19.36 (\<exists>m ctxt q.
19.37 hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
19.38 - Says S R
19.39 - {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
19.40 - Number ctxt, Nonce q,
19.41 - Crypt (pubEK TTP)
19.42 - {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))"
19.43 + Says S R
19.44 + {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
19.45 + Number ctxt, Nonce q,
19.46 + Crypt (pubEK TTP)
19.47 + {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))"
19.48 apply (erule certified_mail.induct, analz_mono_contra)
19.49 apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp)
19.50 apply (simp add: used_Nil Crypt_notin_initState, simp_all)
19.51 @@ -322,11 +322,11 @@
19.52 evs \<in> certified_mail|]
19.53 ==> \<exists>m ctxt q.
19.54 hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
19.55 - Says S R
19.56 - {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
19.57 - Number ctxt, Nonce q,
19.58 - Crypt (pubEK TTP)
19.59 - {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs"
19.60 + Says S R
19.61 + {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
19.62 + Number ctxt, Nonce q,
19.63 + Crypt (pubEK TTP)
19.64 + {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs"
19.65 by (blast intro: S2TTP_sender_lemma)
19.66
19.67
19.68 @@ -401,7 +401,7 @@
19.69 Number cleartext, Nonce q, S2TTP|} \<in> set evs;
19.70 S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
19.71 Key K \<in> analz (spies evs);
19.72 - evs \<in> certified_mail;
19.73 + evs \<in> certified_mail;
19.74 S\<noteq>Spy|]
19.75 ==> R \<in> bad & Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
19.76 apply (erule rev_mp)
19.77 @@ -428,7 +428,7 @@
19.78 "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
19.79 Number cleartext, Nonce q, S2TTP|} \<in> set evs;
19.80 S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
19.81 - evs \<in> certified_mail;
19.82 + evs \<in> certified_mail;
19.83 S\<noteq>Spy; R \<notin> bad|]
19.84 ==> Key K \<notin> analz(spies evs)"
19.85 by (blast dest: S_fairness_bad_R)
19.86 @@ -438,10 +438,10 @@
19.87 until @{term S} has access to the return receipt.*}
19.88 theorem S_guarantee:
19.89 "[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
19.90 - Number cleartext, Nonce q, S2TTP|} \<in> set evs;
19.91 - S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
19.92 - Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
19.93 - S\<noteq>Spy; evs \<in> certified_mail|]
19.94 + Number cleartext, Nonce q, S2TTP|} \<in> set evs;
19.95 + S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
19.96 + Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
19.97 + S\<noteq>Spy; evs \<in> certified_mail|]
19.98 ==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
19.99 apply (erule rev_mp)
19.100 apply (erule ssubst)
20.1 --- a/src/HOL/Auth/Event.thy Sat Oct 17 01:05:59 2009 +0200
20.2 +++ b/src/HOL/Auth/Event.thy Sat Oct 17 14:43:18 2009 +0200
20.3 @@ -22,7 +22,7 @@
20.4 | Notes agent msg
20.5
20.6 consts
20.7 - bad :: "agent set" (*compromised agents*)
20.8 + bad :: "agent set" -- {* compromised agents *}
20.9 knows :: "agent => event list => msg set"
20.10
20.11
20.12 @@ -43,19 +43,19 @@
20.13 knows_Cons:
20.14 "knows A (ev # evs) =
20.15 (if A = Spy then
20.16 - (case ev of
20.17 - Says A' B X => insert X (knows Spy evs)
20.18 - | Gets A' X => knows Spy evs
20.19 - | Notes A' X =>
20.20 - if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
20.21 - else
20.22 - (case ev of
20.23 - Says A' B X =>
20.24 - if A'=A then insert X (knows A evs) else knows A evs
20.25 - | Gets A' X =>
20.26 - if A'=A then insert X (knows A evs) else knows A evs
20.27 - | Notes A' X =>
20.28 - if A'=A then insert X (knows A evs) else knows A evs))"
20.29 + (case ev of
20.30 + Says A' B X => insert X (knows Spy evs)
20.31 + | Gets A' X => knows Spy evs
20.32 + | Notes A' X =>
20.33 + if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
20.34 + else
20.35 + (case ev of
20.36 + Says A' B X =>
20.37 + if A'=A then insert X (knows A evs) else knows A evs
20.38 + | Gets A' X =>
20.39 + if A'=A then insert X (knows A evs) else knows A evs
20.40 + | Notes A' X =>
20.41 + if A'=A then insert X (knows A evs) else knows A evs))"
20.42
20.43 (*
20.44 Case A=Spy on the Gets event
20.45 @@ -71,10 +71,10 @@
20.46 primrec
20.47 used_Nil: "used [] = (UN B. parts (initState B))"
20.48 used_Cons: "used (ev # evs) =
20.49 - (case ev of
20.50 - Says A B X => parts {X} \<union> used evs
20.51 - | Gets A X => used evs
20.52 - | Notes A X => parts {X} \<union> used evs)"
20.53 + (case ev of
20.54 + Says A B X => parts {X} \<union> used evs
20.55 + | Gets A X => used evs
20.56 + | Notes A X => parts {X} \<union> used evs)"
20.57 --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
20.58 follows @{term Says} in real protocols. Seems difficult to change.
20.59 See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
21.1 --- a/src/HOL/Auth/KerberosIV.thy Sat Oct 17 01:05:59 2009 +0200
21.2 +++ b/src/HOL/Auth/KerberosIV.thy Sat Oct 17 14:43:18 2009 +0200
21.3 @@ -170,18 +170,18 @@
21.4 B \<noteq> Tgs; authK \<in> symKeys;
21.5 Says A' Tgs \<lbrace>
21.6 (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
21.7 - Number Ta\<rbrace>),
21.8 + Number Ta\<rbrace>),
21.9 (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
21.10 - \<in> set evs4;
21.11 + \<in> set evs4;
21.12 \<not> expiredAK Ta evs4;
21.13 \<not> expiredA T2 evs4;
21.14 servKlife + (CT evs4) <= authKlife + Ta
21.15 \<rbrakk>
21.16 \<Longrightarrow> Says Tgs A
21.17 (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
21.18 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
21.19 - Number (CT evs4)\<rbrace> \<rbrace>)
21.20 - # evs4 \<in> kerbIV"
21.21 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
21.22 + Number (CT evs4)\<rbrace> \<rbrace>)
21.23 + # evs4 \<in> kerbIV"
21.24 (* Tgs creates a new session key per each request for a service, without
21.25 checking if there is still a fresh one for that service.
21.26 The cipher under Tgs' key is the authTicket, the cipher under B's key
21.27 @@ -196,14 +196,14 @@
21.28 | K5: "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
21.29 Says A Tgs
21.30 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
21.31 - Agent B\<rbrace>
21.32 + Agent B\<rbrace>
21.33 \<in> set evs5;
21.34 Says Tgs' A
21.35 (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
21.36 \<in> set evs5;
21.37 valid Ts wrt T2 \<rbrakk>
21.38 \<Longrightarrow> Says A B \<lbrace>servTicket,
21.39 - Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
21.40 + Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
21.41 # evs5 \<in> kerbIV"
21.42 (* Checks similar to those in K3. *)
21.43
21.44 @@ -609,7 +609,7 @@
21.45 evs \<in> kerbIV \<rbrakk>
21.46 \<Longrightarrow> servK \<notin> range shrK &
21.47 (\<exists>A. servTicket =
21.48 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
21.49 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
21.50 | servTicket \<in> analz (spies evs)"
21.51 apply (frule Says_imp_spies [THEN analz.Inj], auto)
21.52 apply (force dest!: servTicket_form)
21.53 @@ -1336,15 +1336,15 @@
21.54
21.55 lemma Confidentiality_lemma [rule_format]:
21.56 "\<lbrakk> Says Tgs A
21.57 - (Crypt authK
21.58 - \<lbrace>Key servK, Agent B, Number Ts,
21.59 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
21.60 - \<in> set evs;
21.61 - Key authK \<notin> analz (spies evs);
21.62 + (Crypt authK
21.63 + \<lbrace>Key servK, Agent B, Number Ts,
21.64 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
21.65 + \<in> set evs;
21.66 + Key authK \<notin> analz (spies evs);
21.67 servK \<in> symKeys;
21.68 - A \<notin> bad; B \<notin> bad; evs \<in> kerbIV \<rbrakk>
21.69 + A \<notin> bad; B \<notin> bad; evs \<in> kerbIV \<rbrakk>
21.70 \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
21.71 - expiredSK Ts evs"
21.72 + expiredSK Ts evs"
21.73 apply (erule rev_mp)
21.74 apply (erule rev_mp)
21.75 apply (erule kerbIV.induct)
22.1 --- a/src/HOL/Auth/KerberosIV_Gets.thy Sat Oct 17 01:05:59 2009 +0200
22.2 +++ b/src/HOL/Auth/KerberosIV_Gets.thy Sat Oct 17 14:43:18 2009 +0200
22.3 @@ -161,18 +161,18 @@
22.4 B \<noteq> Tgs; authK \<in> symKeys;
22.5 Gets Tgs \<lbrace>
22.6 (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
22.7 - Number Ta\<rbrace>),
22.8 + Number Ta\<rbrace>),
22.9 (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
22.10 - \<in> set evs4;
22.11 + \<in> set evs4;
22.12 \<not> expiredAK Ta evs4;
22.13 \<not> expiredA T2 evs4;
22.14 servKlife + (CT evs4) <= authKlife + Ta
22.15 \<rbrakk>
22.16 \<Longrightarrow> Says Tgs A
22.17 (Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
22.18 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
22.19 - Number (CT evs4)\<rbrace> \<rbrace>)
22.20 - # evs4 \<in> kerbIV_gets"
22.21 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
22.22 + Number (CT evs4)\<rbrace> \<rbrace>)
22.23 + # evs4 \<in> kerbIV_gets"
22.24 (* Tgs creates a new session key per each request for a service, without
22.25 checking if there is still a fresh one for that service.
22.26 The cipher under Tgs' key is the authTicket, the cipher under B's key
22.27 @@ -187,14 +187,14 @@
22.28 | K5: "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
22.29 Says A Tgs
22.30 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
22.31 - Agent B\<rbrace>
22.32 + Agent B\<rbrace>
22.33 \<in> set evs5;
22.34 Gets A
22.35 (Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
22.36 \<in> set evs5;
22.37 valid Ts wrt T2 \<rbrakk>
22.38 \<Longrightarrow> Says A B \<lbrace>servTicket,
22.39 - Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
22.40 + Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
22.41 # evs5 \<in> kerbIV_gets"
22.42 (* Checks similar to those in K3. *)
22.43
22.44 @@ -495,7 +495,7 @@
22.45 evs \<in> kerbIV_gets \<rbrakk>
22.46 \<Longrightarrow> servK \<notin> range shrK &
22.47 (\<exists>A. servTicket =
22.48 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
22.49 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
22.50 | servTicket \<in> analz (spies evs)"
22.51 apply (frule Gets_imp_knows_Spy [THEN analz.Inj], auto)
22.52 apply (force dest!: servTicket_form)
22.53 @@ -1231,15 +1231,15 @@
22.54
22.55 lemma Confidentiality_lemma [rule_format]:
22.56 "\<lbrakk> Says Tgs A
22.57 - (Crypt authK
22.58 - \<lbrace>Key servK, Agent B, Number Ts,
22.59 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
22.60 - \<in> set evs;
22.61 - Key authK \<notin> analz (spies evs);
22.62 + (Crypt authK
22.63 + \<lbrace>Key servK, Agent B, Number Ts,
22.64 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
22.65 + \<in> set evs;
22.66 + Key authK \<notin> analz (spies evs);
22.67 servK \<in> symKeys;
22.68 - A \<notin> bad; B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
22.69 + A \<notin> bad; B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
22.70 \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
22.71 - expiredSK Ts evs"
22.72 + expiredSK Ts evs"
22.73 apply (erule rev_mp)
22.74 apply (erule rev_mp)
22.75 apply (erule kerbIV_gets.induct)
23.1 --- a/src/HOL/Auth/KerberosV.thy Sat Oct 17 01:05:59 2009 +0200
23.2 +++ b/src/HOL/Auth/KerberosV.thy Sat Oct 17 14:43:18 2009 +0200
23.3 @@ -137,9 +137,9 @@
23.4 B \<noteq> Tgs; authK \<in> symKeys;
23.5 Says A' Tgs \<lbrace>
23.6 (Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
23.7 - Number Ta\<rbrace>),
23.8 + Number Ta\<rbrace>),
23.9 (Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
23.10 - \<in> set evs4;
23.11 + \<in> set evs4;
23.12 \<not> expiredAK Ta evs4;
23.13 \<not> expiredA T2 evs4;
23.14 servKlife + (CT evs4) <= authKlife + Ta
23.15 @@ -155,14 +155,14 @@
23.16 A \<noteq> Kas; A \<noteq> Tgs;
23.17 Says A Tgs
23.18 \<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
23.19 - Agent B\<rbrace>
23.20 + Agent B\<rbrace>
23.21 \<in> set evs5;
23.22 Says Tgs' A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
23.23 servTicket\<rbrace>
23.24 \<in> set evs5;
23.25 valid Ts wrt T2 \<rbrakk>
23.26 \<Longrightarrow> Says A B \<lbrace>servTicket,
23.27 - Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
23.28 + Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
23.29 # evs5 \<in> kerbV"
23.30
23.31 | KV6: "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
23.32 @@ -1081,14 +1081,14 @@
23.33
23.34 lemma Confidentiality_lemma [rule_format]:
23.35 "\<lbrakk> Says Tgs A
23.36 - \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
23.37 - Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
23.38 - \<in> set evs;
23.39 - Key authK \<notin> analz (spies evs);
23.40 + \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
23.41 + Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
23.42 + \<in> set evs;
23.43 + Key authK \<notin> analz (spies evs);
23.44 servK \<in> symKeys;
23.45 - A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
23.46 + A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
23.47 \<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
23.48 - expiredSK Ts evs"
23.49 + expiredSK Ts evs"
23.50 apply (erule rev_mp)
23.51 apply (erule rev_mp)
23.52 apply (erule kerbV.induct)
24.1 --- a/src/HOL/Auth/Kerberos_BAN.thy Sat Oct 17 01:05:59 2009 +0200
24.2 +++ b/src/HOL/Auth/Kerberos_BAN.thy Sat Oct 17 14:43:18 2009 +0200
24.3 @@ -77,45 +77,45 @@
24.4 Nil: "[] \<in> bankerberos"
24.5
24.6 | Fake: "\<lbrakk> evsf \<in> bankerberos; X \<in> synth (analz (spies evsf)) \<rbrakk>
24.7 - \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
24.8 + \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
24.9
24.10
24.11 | BK1: "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
24.12 - \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
24.13 - \<in> bankerberos"
24.14 + \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
24.15 + \<in> bankerberos"
24.16
24.17
24.18 | BK2: "\<lbrakk> evs2 \<in> bankerberos; Key K \<notin> used evs2; K \<in> symKeys;
24.19 - Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
24.20 - \<Longrightarrow> Says Server A
24.21 - (Crypt (shrK A)
24.22 - \<lbrace>Number (CT evs2), Agent B, Key K,
24.23 - (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
24.24 - # evs2 \<in> bankerberos"
24.25 + Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
24.26 + \<Longrightarrow> Says Server A
24.27 + (Crypt (shrK A)
24.28 + \<lbrace>Number (CT evs2), Agent B, Key K,
24.29 + (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
24.30 + # evs2 \<in> bankerberos"
24.31
24.32
24.33 | BK3: "\<lbrakk> evs3 \<in> bankerberos;
24.34 - Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
24.35 - \<in> set evs3;
24.36 - Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
24.37 - \<not> expiredK Tk evs3 \<rbrakk>
24.38 - \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
24.39 - # evs3 \<in> bankerberos"
24.40 + Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
24.41 + \<in> set evs3;
24.42 + Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
24.43 + \<not> expiredK Tk evs3 \<rbrakk>
24.44 + \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
24.45 + # evs3 \<in> bankerberos"
24.46
24.47
24.48 | BK4: "\<lbrakk> evs4 \<in> bankerberos;
24.49 - Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
24.50 - (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
24.51 - \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
24.52 - \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
24.53 - \<in> bankerberos"
24.54 + Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
24.55 + (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
24.56 + \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
24.57 + \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
24.58 + \<in> bankerberos"
24.59
24.60 - (*Old session keys may become compromised*)
24.61 + (*Old session keys may become compromised*)
24.62 | Oops: "\<lbrakk> evso \<in> bankerberos;
24.63 Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
24.64 - \<in> set evso;
24.65 - expiredK Tk evso \<rbrakk>
24.66 - \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
24.67 + \<in> set evso;
24.68 + expiredK Tk evso \<rbrakk>
24.69 + \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
24.70
24.71
24.72 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
25.1 --- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Sat Oct 17 01:05:59 2009 +0200
25.2 +++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Sat Oct 17 14:43:18 2009 +0200
25.3 @@ -69,47 +69,47 @@
25.4 Nil: "[] \<in> bankerb_gets"
25.5
25.6 | Fake: "\<lbrakk> evsf \<in> bankerb_gets; X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
25.7 - \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
25.8 + \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
25.9
25.10 | Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
25.11 \<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"
25.12
25.13 | BK1: "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
25.14 - \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
25.15 - \<in> bankerb_gets"
25.16 + \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
25.17 + \<in> bankerb_gets"
25.18
25.19
25.20 | BK2: "\<lbrakk> evs2 \<in> bankerb_gets; Key K \<notin> used evs2; K \<in> symKeys;
25.21 - Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
25.22 - \<Longrightarrow> Says Server A
25.23 - (Crypt (shrK A)
25.24 - \<lbrace>Number (CT evs2), Agent B, Key K,
25.25 - (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
25.26 - # evs2 \<in> bankerb_gets"
25.27 + Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
25.28 + \<Longrightarrow> Says Server A
25.29 + (Crypt (shrK A)
25.30 + \<lbrace>Number (CT evs2), Agent B, Key K,
25.31 + (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
25.32 + # evs2 \<in> bankerb_gets"
25.33
25.34
25.35 | BK3: "\<lbrakk> evs3 \<in> bankerb_gets;
25.36 - Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
25.37 - \<in> set evs3;
25.38 - Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
25.39 - \<not> expiredK Tk evs3 \<rbrakk>
25.40 - \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
25.41 - # evs3 \<in> bankerb_gets"
25.42 + Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
25.43 + \<in> set evs3;
25.44 + Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
25.45 + \<not> expiredK Tk evs3 \<rbrakk>
25.46 + \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
25.47 + # evs3 \<in> bankerb_gets"
25.48
25.49
25.50 | BK4: "\<lbrakk> evs4 \<in> bankerb_gets;
25.51 - Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
25.52 - (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
25.53 - \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
25.54 - \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
25.55 - \<in> bankerb_gets"
25.56 + Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
25.57 + (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
25.58 + \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
25.59 + \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
25.60 + \<in> bankerb_gets"
25.61
25.62 - (*Old session keys may become compromised*)
25.63 + (*Old session keys may become compromised*)
25.64 | Oops: "\<lbrakk> evso \<in> bankerb_gets;
25.65 Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
25.66 - \<in> set evso;
25.67 - expiredK Tk evso \<rbrakk>
25.68 - \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
25.69 + \<in> set evso;
25.70 + expiredK Tk evso \<rbrakk>
25.71 + \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
25.72
25.73
25.74 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
25.75 @@ -359,7 +359,7 @@
25.76 "\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
25.77 B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
25.78 \<Longrightarrow> \<exists> Tk. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
25.79 - Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
25.80 + Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
25.81 apply (erule rev_mp)
25.82 apply (erule bankerb_gets.induct)
25.83 apply auto
26.1 --- a/src/HOL/Auth/Message.thy Sat Oct 17 01:05:59 2009 +0200
26.2 +++ b/src/HOL/Auth/Message.thy Sat Oct 17 14:43:18 2009 +0200
26.3 @@ -40,13 +40,13 @@
26.4 agent = Server | Friend nat | Spy
26.5
26.6 datatype
26.7 - msg = Agent agent --{*Agent names*}
26.8 + msg = Agent agent --{*Agent names*}
26.9 | Number nat --{*Ordinary integers, timestamps, ...*}
26.10 | Nonce nat --{*Unguessable nonces*}
26.11 | Key key --{*Crypto keys*}
26.12 - | Hash msg --{*Hashing*}
26.13 - | MPair msg msg --{*Compound messages*}
26.14 - | Crypt key msg --{*Encryption, public- or shared-key*}
26.15 + | Hash msg --{*Hashing*}
26.16 + | MPair msg msg --{*Compound messages*}
26.17 + | Crypt key msg --{*Encryption, public- or shared-key*}
26.18
26.19
26.20 text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
26.21 @@ -873,8 +873,8 @@
26.22 (Fake_insert_simp_tac ss 1
26.23 THEN
26.24 IF_UNSOLVED (Blast.depth_tac
26.25 - (cs addIs [@{thm analz_insertI},
26.26 - impOfSubs @{thm analz_subset_parts}]) 4 1))
26.27 + (cs addIs [@{thm analz_insertI},
26.28 + impOfSubs @{thm analz_subset_parts}]) 4 1))
26.29
26.30 fun spy_analz_tac (cs,ss) i =
26.31 DETERM
27.1 --- a/src/HOL/Auth/NS_Public.thy Sat Oct 17 01:05:59 2009 +0200
27.2 +++ b/src/HOL/Auth/NS_Public.thy Sat Oct 17 14:43:18 2009 +0200
27.3 @@ -110,8 +110,8 @@
27.4 lemma A_trusts_NS2_lemma [rule_format]:
27.5 "\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
27.6 \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>
27.7 - Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
27.8 - Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
27.9 + Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
27.10 + Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
27.11 apply (erule ns_public.induct, simp_all)
27.12 (*Fake, NS1*)
27.13 apply (blast dest: Spy_not_see_NA)+
27.14 @@ -129,8 +129,8 @@
27.15 lemma B_trusts_NS1 [rule_format]:
27.16 "evs \<in> ns_public
27.17 \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
27.18 - Nonce NA \<notin> analz (spies evs) \<longrightarrow>
27.19 - Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
27.20 + Nonce NA \<notin> analz (spies evs) \<longrightarrow>
27.21 + Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
27.22 apply (erule ns_public.induct, simp_all)
27.23 (*Fake*)
27.24 apply (blast intro!: analz_insertI)
28.1 --- a/src/HOL/Auth/NS_Public_Bad.thy Sat Oct 17 01:05:59 2009 +0200
28.2 +++ b/src/HOL/Auth/NS_Public_Bad.thy Sat Oct 17 14:43:18 2009 +0200
28.3 @@ -116,8 +116,8 @@
28.4 lemma A_trusts_NS2_lemma [rule_format]:
28.5 "\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
28.6 \<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
28.7 - Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
28.8 - Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
28.9 + Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
28.10 + Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
28.11 apply (erule ns_public.induct)
28.12 apply (auto dest: Spy_not_see_NA unique_NA)
28.13 done
28.14 @@ -134,8 +134,8 @@
28.15 lemma B_trusts_NS1 [rule_format]:
28.16 "evs \<in> ns_public
28.17 \<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
28.18 - Nonce NA \<notin> analz (spies evs) \<longrightarrow>
28.19 - Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
28.20 + Nonce NA \<notin> analz (spies evs) \<longrightarrow>
28.21 + Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
28.22 apply (erule ns_public.induct, simp_all)
28.23 (*Fake*)
28.24 apply (blast intro!: analz_insertI)
29.1 --- a/src/HOL/Auth/NS_Shared.thy Sat Oct 17 01:05:59 2009 +0200
29.2 +++ b/src/HOL/Auth/NS_Shared.thy Sat Oct 17 14:43:18 2009 +0200
29.3 @@ -27,61 +27,61 @@
29.4
29.5 inductive_set ns_shared :: "event list set"
29.6 where
29.7 - (*Initial trace is empty*)
29.8 + (*Initial trace is empty*)
29.9 Nil: "[] \<in> ns_shared"
29.10 - (*The spy MAY say anything he CAN say. We do not expect him to
29.11 - invent new nonces here, but he can also use NS1. Common to
29.12 - all similar protocols.*)
29.13 + (*The spy MAY say anything he CAN say. We do not expect him to
29.14 + invent new nonces here, but he can also use NS1. Common to
29.15 + all similar protocols.*)
29.16 | Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk>
29.17 - \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
29.18 + \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
29.19
29.20 - (*Alice initiates a protocol run, requesting to talk to any B*)
29.21 + (*Alice initiates a protocol run, requesting to talk to any B*)
29.22 | NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk>
29.23 - \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared"
29.24 + \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared"
29.25
29.26 - (*Server's response to Alice's message.
29.27 - !! It may respond more than once to A's request !!
29.28 - Server doesn't know who the true sender is, hence the A' in
29.29 - the sender field.*)
29.30 + (*Server's response to Alice's message.
29.31 + !! It may respond more than once to A's request !!
29.32 + Server doesn't know who the true sender is, hence the A' in
29.33 + the sender field.*)
29.34 | NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; KAB \<in> symKeys;
29.35 - Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
29.36 - \<Longrightarrow> Says Server A
29.37 - (Crypt (shrK A)
29.38 - \<lbrace>Nonce NA, Agent B, Key KAB,
29.39 - (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
29.40 - # evs2 \<in> ns_shared"
29.41 + Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
29.42 + \<Longrightarrow> Says Server A
29.43 + (Crypt (shrK A)
29.44 + \<lbrace>Nonce NA, Agent B, Key KAB,
29.45 + (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
29.46 + # evs2 \<in> ns_shared"
29.47
29.48 - (*We can't assume S=Server. Agent A "remembers" her nonce.
29.49 - Need A \<noteq> Server because we allow messages to self.*)
29.50 + (*We can't assume S=Server. Agent A "remembers" her nonce.
29.51 + Need A \<noteq> Server because we allow messages to self.*)
29.52 | NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server;
29.53 - Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
29.54 - Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
29.55 - \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
29.56 + Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
29.57 + Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
29.58 + \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
29.59
29.60 - (*Bob's nonce exchange. He does not know who the message came
29.61 - from, but responds to A because she is mentioned inside.*)
29.62 + (*Bob's nonce exchange. He does not know who the message came
29.63 + from, but responds to A because she is mentioned inside.*)
29.64 | NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; K \<in> symKeys;
29.65 - Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
29.66 - \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
29.67 + Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
29.68 + \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
29.69
29.70 - (*Alice responds with Nonce NB if she has seen the key before.
29.71 - Maybe should somehow check Nonce NA again.
29.72 - We do NOT send NB-1 or similar as the Spy cannot spoof such things.
29.73 - Letting the Spy add or subtract 1 lets him send all nonces.
29.74 - Instead we distinguish the messages by sending the nonce twice.*)
29.75 + (*Alice responds with Nonce NB if she has seen the key before.
29.76 + Maybe should somehow check Nonce NA again.
29.77 + We do NOT send NB-1 or similar as the Spy cannot spoof such things.
29.78 + Letting the Spy add or subtract 1 lets him send all nonces.
29.79 + Instead we distinguish the messages by sending the nonce twice.*)
29.80 | NS5: "\<lbrakk>evs5 \<in> ns_shared; K \<in> symKeys;
29.81 - Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
29.82 - Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
29.83 - \<in> set evs5\<rbrakk>
29.84 - \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
29.85 + Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
29.86 + Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
29.87 + \<in> set evs5\<rbrakk>
29.88 + \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
29.89
29.90 - (*This message models possible leaks of session keys.
29.91 - The two Nonces identify the protocol run: the rule insists upon
29.92 - the true senders in order to make them accurate.*)
29.93 + (*This message models possible leaks of session keys.
29.94 + The two Nonces identify the protocol run: the rule insists upon
29.95 + the true senders in order to make them accurate.*)
29.96 | Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso;
29.97 - Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
29.98 - \<in> set evso\<rbrakk>
29.99 - \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
29.100 + Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
29.101 + \<in> set evso\<rbrakk>
29.102 + \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
29.103
29.104
29.105 declare Says_imp_knows_Spy [THEN parts.Inj, dest]
29.106 @@ -98,7 +98,7 @@
29.107 apply (intro exI bexI)
29.108 apply (rule_tac [2] ns_shared.Nil
29.109 [THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3,
29.110 - THEN ns_shared.NS4, THEN ns_shared.NS5])
29.111 + THEN ns_shared.NS4, THEN ns_shared.NS5])
29.112 apply (possibility, simp add: used_Cons)
29.113 done
29.114
29.115 @@ -275,7 +275,7 @@
29.116 apply blast
29.117 txt{*NS3*}
29.118 apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
29.119 - dest: Says_imp_knows_Spy analz.Inj unique_session_keys)
29.120 + dest: Says_imp_knows_Spy analz.Inj unique_session_keys)
29.121 txt{*Oops*}
29.122 apply (blast dest: unique_session_keys)
29.123 done
29.124 @@ -355,8 +355,8 @@
29.125 "\<lbrakk>B \<notin> bad; evs \<in> ns_shared\<rbrakk> \<Longrightarrow>
29.126 Key K \<notin> analz (spies evs) \<longrightarrow>
29.127 Says Server A
29.128 - (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
29.129 - Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
29.130 + (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
29.131 + Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
29.132 Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
29.133 Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
29.134 apply (erule ns_shared.induct, force)
29.135 @@ -366,7 +366,7 @@
29.136 apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
29.137 txt{*NS5*}
29.138 apply (blast dest!: A_trusts_NS2
29.139 - dest: Says_imp_knows_Spy [THEN analz.Inj]
29.140 + dest: Says_imp_knows_Spy [THEN analz.Inj]
29.141 unique_session_keys Crypt_Spy_analz_bad)
29.142 done
29.143
30.1 --- a/src/HOL/Auth/OtwayRees.thy Sat Oct 17 01:05:59 2009 +0200
30.2 +++ b/src/HOL/Auth/OtwayRees.thy Sat Oct 17 14:43:18 2009 +0200
30.3 @@ -61,7 +61,7 @@
30.4 # evs3 : otway"
30.5
30.6 (*Bob receives the Server's (?) message and compares the Nonces with
30.7 - those in the message he previously sent the Server.
30.8 + those in the message he previously sent the Server.
30.9 Need B \<noteq> Server because we allow messages to self.*)
30.10 | OR4: "[| evs4 \<in> otway; B \<noteq> Server;
30.11 Says B Server {|Nonce NA, Agent A, Agent B, X',
31.1 --- a/src/HOL/Auth/OtwayReesBella.thy Sat Oct 17 01:05:59 2009 +0200
31.2 +++ b/src/HOL/Auth/OtwayReesBella.thy Sat Oct 17 14:43:18 2009 +0200
31.3 @@ -22,50 +22,50 @@
31.4 Nil: "[]\<in> orb"
31.5
31.6 | Fake: "\<lbrakk>evsa\<in> orb; X\<in> synth (analz (knows Spy evsa))\<rbrakk>
31.7 - \<Longrightarrow> Says Spy B X # evsa \<in> orb"
31.8 + \<Longrightarrow> Says Spy B X # evsa \<in> orb"
31.9
31.10 | Reception: "\<lbrakk>evsr\<in> orb; Says A B X \<in> set evsr\<rbrakk>
31.11 - \<Longrightarrow> Gets B X # evsr \<in> orb"
31.12 + \<Longrightarrow> Gets B X # evsr \<in> orb"
31.13
31.14 | OR1: "\<lbrakk>evs1\<in> orb; Nonce NA \<notin> used evs1\<rbrakk>
31.15 - \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B,
31.16 - Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.17 - # evs1 \<in> orb"
31.18 + \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B,
31.19 + Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.20 + # evs1 \<in> orb"
31.21
31.22 | OR2: "\<lbrakk>evs2\<in> orb; Nonce NB \<notin> used evs2;
31.23 - Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
31.24 - \<Longrightarrow> Says B Server
31.25 - \<lbrace>Nonce M, Agent A, Agent B, X,
31.26 - Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.27 - # evs2 \<in> orb"
31.28 + Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
31.29 + \<Longrightarrow> Says B Server
31.30 + \<lbrace>Nonce M, Agent A, Agent B, X,
31.31 + Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.32 + # evs2 \<in> orb"
31.33
31.34 | OR3: "\<lbrakk>evs3\<in> orb; Key KAB \<notin> used evs3;
31.35 - Gets Server
31.36 - \<lbrace>Nonce M, Agent A, Agent B,
31.37 - Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>,
31.38 - Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.39 - \<in> set evs3\<rbrakk>
31.40 - \<Longrightarrow> Says Server B \<lbrace>Nonce M,
31.41 - Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
31.42 - Nonce NB, Key KAB\<rbrace>\<rbrace>
31.43 - # evs3 \<in> orb"
31.44 + Gets Server
31.45 + \<lbrace>Nonce M, Agent A, Agent B,
31.46 + Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>,
31.47 + Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.48 + \<in> set evs3\<rbrakk>
31.49 + \<Longrightarrow> Says Server B \<lbrace>Nonce M,
31.50 + Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
31.51 + Nonce NB, Key KAB\<rbrace>\<rbrace>
31.52 + # evs3 \<in> orb"
31.53
31.54 (*B can only check that the message he is bouncing is a ciphertext*)
31.55 (*Sending M back is omitted*)
31.56 | OR4: "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>;
31.57 - Says B Server \<lbrace>Nonce M, Agent A, Agent B, X',
31.58 - Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.59 - \<in> set evs4;
31.60 - Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
31.61 - \<in> set evs4\<rbrakk>
31.62 - \<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
31.63 + Says B Server \<lbrace>Nonce M, Agent A, Agent B, X',
31.64 + Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
31.65 + \<in> set evs4;
31.66 + Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
31.67 + \<in> set evs4\<rbrakk>
31.68 + \<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
31.69
31.70
31.71 | Oops: "\<lbrakk>evso\<in> orb;
31.72 - Says Server B \<lbrace>Nonce M,
31.73 - Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
31.74 - Nonce NB, Key KAB\<rbrace>\<rbrace>
31.75 - \<in> set evso\<rbrakk>
31.76 + Says Server B \<lbrace>Nonce M,
31.77 + Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
31.78 + Nonce NB, Key KAB\<rbrace>\<rbrace>
31.79 + \<in> set evso\<rbrakk>
31.80 \<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso
31.81 \<in> orb"
31.82
32.1 --- a/src/HOL/Auth/OtwayRees_AN.thy Sat Oct 17 01:05:59 2009 +0200
32.2 +++ b/src/HOL/Auth/OtwayRees_AN.thy Sat Oct 17 14:43:18 2009 +0200
32.3 @@ -35,7 +35,7 @@
32.4
32.5 | Reception: --{*A message that has been sent can be received by the
32.6 intended recipient.*}
32.7 - "[| evsr \<in> otway; Says A B X \<in>set evsr |]
32.8 + "[| evsr \<in> otway; Says A B X \<in>set evsr |]
32.9 ==> Gets B X # evsr \<in> otway"
32.10
32.11 | OR1: --{*Alice initiates a protocol run*}
32.12 @@ -43,14 +43,14 @@
32.13 ==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
32.14
32.15 | OR2: --{*Bob's response to Alice's message.*}
32.16 - "[| evs2 \<in> otway;
32.17 + "[| evs2 \<in> otway;
32.18 Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
32.19 ==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
32.20 # evs2 \<in> otway"
32.21
32.22 | OR3: --{*The Server receives Bob's message. Then he sends a new
32.23 session key to Bob with a packet for forwarding to Alice.*}
32.24 - "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
32.25 + "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
32.26 Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
32.27 \<in>set evs3 |]
32.28 ==> Says Server B
32.29 @@ -59,9 +59,9 @@
32.30 # evs3 \<in> otway"
32.31
32.32 | OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
32.33 - those in the message he previously sent the Server.
32.34 + those in the message he previously sent the Server.
32.35 Need @{term "B \<noteq> Server"} because we allow messages to self.*}
32.36 - "[| evs4 \<in> otway; B \<noteq> Server;
32.37 + "[| evs4 \<in> otway; B \<noteq> Server;
32.38 Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4;
32.39 Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
32.40 \<in>set evs4 |]
32.41 @@ -69,7 +69,7 @@
32.42
32.43 | Oops: --{*This message models possible leaks of session keys. The nonces
32.44 identify the protocol run.*}
32.45 - "[| evso \<in> otway;
32.46 + "[| evso \<in> otway;
32.47 Says Server B
32.48 {|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|},
32.49 Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
33.1 --- a/src/HOL/Auth/OtwayRees_Bad.thy Sat Oct 17 01:05:59 2009 +0200
33.2 +++ b/src/HOL/Auth/OtwayRees_Bad.thy Sat Oct 17 14:43:18 2009 +0200
33.3 @@ -32,18 +32,18 @@
33.4
33.5 | Reception: --{*A message that has been sent can be received by the
33.6 intended recipient.*}
33.7 - "[| evsr \<in> otway; Says A B X \<in>set evsr |]
33.8 + "[| evsr \<in> otway; Says A B X \<in>set evsr |]
33.9 ==> Gets B X # evsr \<in> otway"
33.10
33.11 | OR1: --{*Alice initiates a protocol run*}
33.12 - "[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
33.13 + "[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
33.14 ==> Says A B {|Nonce NA, Agent A, Agent B,
33.15 Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
33.16 # evs1 \<in> otway"
33.17
33.18 | OR2: --{*Bob's response to Alice's message.
33.19 This variant of the protocol does NOT encrypt NB.*}
33.20 - "[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
33.21 + "[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
33.22 Gets B {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
33.23 ==> Says B Server
33.24 {|Nonce NA, Agent A, Agent B, X, Nonce NB,
33.25 @@ -53,7 +53,7 @@
33.26 | OR3: --{*The Server receives Bob's message and checks that the three NAs
33.27 match. Then he sends a new session key to Bob with a packet for
33.28 forwarding to Alice.*}
33.29 - "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
33.30 + "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
33.31 Gets Server
33.32 {|Nonce NA, Agent A, Agent B,
33.33 Crypt (shrK A) {|Nonce NA, Agent A, Agent B|},
33.34 @@ -67,9 +67,9 @@
33.35 # evs3 \<in> otway"
33.36
33.37 | OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
33.38 - those in the message he previously sent the Server.
33.39 + those in the message he previously sent the Server.
33.40 Need @{term "B \<noteq> Server"} because we allow messages to self.*}
33.41 - "[| evs4 \<in> otway; B \<noteq> Server;
33.42 + "[| evs4 \<in> otway; B \<noteq> Server;
33.43 Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
33.44 Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
33.45 \<in> set evs4;
33.46 @@ -79,7 +79,7 @@
33.47
33.48 | Oops: --{*This message models possible leaks of session keys. The nonces
33.49 identify the protocol run.*}
33.50 - "[| evso \<in> otway;
33.51 + "[| evso \<in> otway;
33.52 Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
33.53 \<in> set evso |]
33.54 ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway"
34.1 --- a/src/HOL/Auth/Recur.thy Sat Oct 17 01:05:59 2009 +0200
34.2 +++ b/src/HOL/Auth/Recur.thy Sat Oct 17 14:43:18 2009 +0200
34.3 @@ -101,7 +101,7 @@
34.4 etc.
34.5
34.6 Oops: "[| evso \<in> recur; Says Server B RB \<in> set evso;
34.7 - RB \<in> responses evs'; Key K \<in> parts {RB} |]
34.8 + RB \<in> responses evs'; Key K \<in> parts {RB} |]
34.9 ==> Notes Spy {|Key K, RB|} # evso \<in> recur"
34.10 *)
34.11
34.12 @@ -140,10 +140,10 @@
34.13 apply (rule_tac [2]
34.14 recur.Nil
34.15 [THEN recur.RA1 [of _ NA],
34.16 - THEN recur.RA2 [of _ NB],
34.17 - THEN recur.RA3 [OF _ _ respond.One
34.18 + THEN recur.RA2 [of _ NB],
34.19 + THEN recur.RA3 [OF _ _ respond.One
34.20 [THEN respond.Cons [of _ _ K _ K']]],
34.21 - THEN recur.RA4], possibility)
34.22 + THEN recur.RA4], possibility)
34.23 apply (auto simp add: used_Cons)
34.24 done
34.25
34.26 @@ -241,7 +241,7 @@
34.27 (K \<in> KK | Key K \<in> analz (insert RB H))"
34.28 apply (erule responses.induct)
34.29 apply (simp_all del: image_insert
34.30 - add: analz_image_freshK_simps, auto)
34.31 + add: analz_image_freshK_simps, auto)
34.32 done
34.33
34.34
35.1 --- a/src/HOL/Auth/Smartcard/EventSC.thy Sat Oct 17 01:05:59 2009 +0200
35.2 +++ b/src/HOL/Auth/Smartcard/EventSC.thy Sat Oct 17 14:43:18 2009 +0200
35.3 @@ -38,13 +38,13 @@
35.4 specification (stolen)
35.5 (*The server's card is secure by assumption\<dots>*)
35.6 Card_Server_not_stolen [iff]: "Card Server \<notin> stolen"
35.7 - Card_Spy_not_stolen [iff]: "Card Spy \<notin> stolen"
35.8 + Card_Spy_not_stolen [iff]: "Card Spy \<notin> stolen"
35.9 apply blast done
35.10
35.11 specification (cloned)
35.12 (*\<dots> the spy's card is secure because she already can use it freely*)
35.13 Card_Server_not_cloned [iff]: "Card Server \<notin> cloned"
35.14 - Card_Spy_not_cloned [iff]: "Card Spy \<notin> cloned"
35.15 + Card_Spy_not_cloned [iff]: "Card Spy \<notin> cloned"
35.16 apply blast done
35.17
35.18
35.19 @@ -52,28 +52,28 @@
35.20 assumption of secure means*)
35.21 knows_Nil: "knows A [] = initState A"
35.22 knows_Cons: "knows A (ev # evs) =
35.23 - (case ev of
35.24 - Says A' B X =>
35.25 + (case ev of
35.26 + Says A' B X =>
35.27 if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
35.28 - | Notes A' X =>
35.29 - if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs)
35.30 + | Notes A' X =>
35.31 + if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs)
35.32 else knows A evs
35.33 | Gets A' X =>
35.34 - if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
35.35 + if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
35.36 else knows A evs
35.37 | Inputs A' C X =>
35.38 - if secureM then
35.39 + if secureM then
35.40 if A=A' then insert X (knows A evs) else knows A evs
35.41 - else
35.42 - if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
35.43 + else
35.44 + if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
35.45 | C_Gets C X => knows A evs
35.46 | Outpts C A' X =>
35.47 - if secureM then
35.48 + if secureM then
35.49 if A=A' then insert X (knows A evs) else knows A evs
35.50 else
35.51 - if A=Spy then insert X (knows A evs) else knows A evs
35.52 + if A=Spy then insert X (knows A evs) else knows A evs
35.53 | A_Gets A' X =>
35.54 - if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
35.55 + if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
35.56 else knows A evs)"
35.57
35.58
35.59 @@ -86,14 +86,14 @@
35.60 primrec
35.61 used_Nil: "used [] = (UN B. parts (initState B))"
35.62 used_Cons: "used (ev # evs) =
35.63 - (case ev of
35.64 - Says A B X => parts {X} \<union> (used evs)
35.65 - | Notes A X => parts {X} \<union> (used evs)
35.66 - | Gets A X => used evs
35.67 + (case ev of
35.68 + Says A B X => parts {X} \<union> (used evs)
35.69 + | Notes A X => parts {X} \<union> (used evs)
35.70 + | Gets A X => used evs
35.71 | Inputs A C X => parts{X} \<union> (used evs)
35.72 - | C_Gets C X => used evs
35.73 + | C_Gets C X => used evs
35.74 | Outpts C A X => parts{X} \<union> (used evs)
35.75 - | A_Gets A X => used evs)"
35.76 + | A_Gets A X => used evs)"
35.77 --{*@{term Gets} always follows @{term Says} in real protocols.
35.78 Likewise, @{term C_Gets} will always have to follow @{term Inputs}
35.79 and @{term A_Gets} will always have to follow @{term Outpts}*}
36.1 --- a/src/HOL/Auth/Smartcard/Smartcard.thy Sat Oct 17 01:05:59 2009 +0200
36.2 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Sat Oct 17 14:43:18 2009 +0200
36.3 @@ -386,8 +386,8 @@
36.4
36.5 val analz_image_freshK_ss =
36.6 @{simpset} delsimps [image_insert, image_Un]
36.7 - delsimps [@{thm imp_disjL}] (*reduces blow-up*)
36.8 - addsimps @{thms analz_image_freshK_simps}
36.9 + delsimps [@{thm imp_disjL}] (*reduces blow-up*)
36.10 + addsimps @{thms analz_image_freshK_simps}
36.11 end
36.12 *}
36.13
37.1 --- a/src/HOL/Auth/TLS.thy Sat Oct 17 01:05:59 2009 +0200
37.2 +++ b/src/HOL/Auth/TLS.thy Sat Oct 17 14:43:18 2009 +0200
37.3 @@ -36,7 +36,7 @@
37.4 Proofs would be simpler if ClientKeyExch included A's name within
37.5 Crypt KB (Nonce PMS). As things stand, there is much overlap between proofs
37.6 about that message (which B receives) and the stronger event
37.7 - Notes A {|Agent B, Nonce PMS|}.
37.8 +Notes A {|Agent B, Nonce PMS|}.
37.9 *)
37.10
37.11 header{*The TLS Protocol: Transport Layer Security*}
37.12 @@ -112,30 +112,30 @@
37.13 | SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
37.14 to available nonces*}
37.15 "[| evsSK \<in> tls;
37.16 - {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
37.17 + {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
37.18 ==> Notes Spy {| Nonce (PRF(M,NA,NB)),
37.19 - Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
37.20 + Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
37.21
37.22 | ClientHello:
37.23 - --{*(7.4.1.2)
37.24 - PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
37.25 - It is uninterpreted but will be confirmed in the FINISHED messages.
37.26 - NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
37.27 + --{*(7.4.1.2)
37.28 + PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
37.29 + It is uninterpreted but will be confirmed in the FINISHED messages.
37.30 + NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
37.31 UNIX TIME is omitted because the protocol doesn't use it.
37.32 May assume @{term "NA \<notin> range PRF"} because CLIENT RANDOM is
37.33 28 bytes while MASTER SECRET is 48 bytes*}
37.34 "[| evsCH \<in> tls; Nonce NA \<notin> used evsCH; NA \<notin> range PRF |]
37.35 ==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
37.36 - # evsCH \<in> tls"
37.37 + # evsCH \<in> tls"
37.38
37.39 | ServerHello:
37.40 --{*7.4.1.3 of the TLS Internet-Draft
37.41 - PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
37.42 + PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
37.43 SERVER CERTIFICATE (7.4.2) is always present.
37.44 @{text CERTIFICATE_REQUEST} (7.4.4) is implied.*}
37.45 "[| evsSH \<in> tls; Nonce NB \<notin> used evsSH; NB \<notin> range PRF;
37.46 Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
37.47 - \<in> set evsSH |]
37.48 + \<in> set evsSH |]
37.49 ==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH \<in> tls"
37.50
37.51 | Certificate:
37.52 @@ -148,28 +148,28 @@
37.53 She encrypts PMS using the supplied KB, which ought to be pubK B.
37.54 We assume @{term "PMS \<notin> range PRF"} because a clash betweem the PMS
37.55 and another MASTER SECRET is highly unlikely (even though
37.56 - both items have the same length, 48 bytes).
37.57 + both items have the same length, 48 bytes).
37.58 The Note event records in the trace that she knows PMS
37.59 (see REMARK at top). *}
37.60 "[| evsCX \<in> tls; Nonce PMS \<notin> used evsCX; PMS \<notin> range PRF;
37.61 Says B' A (certificate B KB) \<in> set evsCX |]
37.62 ==> Says A B (Crypt KB (Nonce PMS))
37.63 - # Notes A {|Agent B, Nonce PMS|}
37.64 - # evsCX \<in> tls"
37.65 + # Notes A {|Agent B, Nonce PMS|}
37.66 + # evsCX \<in> tls"
37.67
37.68 | CertVerify:
37.69 - --{*The optional Certificate Verify (7.4.8) message contains the
37.70 + --{*The optional Certificate Verify (7.4.8) message contains the
37.71 specific components listed in the security analysis, F.1.1.2.
37.72 It adds the pre-master-secret, which is also essential!
37.73 Checking the signature, which is the only use of A's certificate,
37.74 assures B of A's presence*}
37.75 "[| evsCV \<in> tls;
37.76 Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCV;
37.77 - Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
37.78 + Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
37.79 ==> Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))
37.80 # evsCV \<in> tls"
37.81
37.82 - --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
37.83 + --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
37.84 among other things. The master-secret is PRF(PMS,NA,NB).
37.85 Either party may send its message first.*}
37.86
37.87 @@ -181,60 +181,60 @@
37.88 could simply put @{term "A\<noteq>Spy"} into the rule, but one should not
37.89 expect the spy to be well-behaved.*}
37.90 "[| evsCF \<in> tls;
37.91 - Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
37.92 - \<in> set evsCF;
37.93 + Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
37.94 + \<in> set evsCF;
37.95 Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCF;
37.96 Notes A {|Agent B, Nonce PMS|} \<in> set evsCF;
37.97 - M = PRF(PMS,NA,NB) |]
37.98 + M = PRF(PMS,NA,NB) |]
37.99 ==> Says A B (Crypt (clientK(NA,NB,M))
37.100 - (Hash{|Number SID, Nonce M,
37.101 - Nonce NA, Number PA, Agent A,
37.102 - Nonce NB, Number PB, Agent B|}))
37.103 + (Hash{|Number SID, Nonce M,
37.104 + Nonce NA, Number PA, Agent A,
37.105 + Nonce NB, Number PB, Agent B|}))
37.106 # evsCF \<in> tls"
37.107
37.108 | ServerFinished:
37.109 - --{*Keeping A' and A'' distinct means B cannot even check that the
37.110 + --{*Keeping A' and A'' distinct means B cannot even check that the
37.111 two messages originate from the same source. *}
37.112 "[| evsSF \<in> tls;
37.113 - Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
37.114 - \<in> set evsSF;
37.115 - Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
37.116 - Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
37.117 - M = PRF(PMS,NA,NB) |]
37.118 + Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
37.119 + \<in> set evsSF;
37.120 + Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
37.121 + Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
37.122 + M = PRF(PMS,NA,NB) |]
37.123 ==> Says B A (Crypt (serverK(NA,NB,M))
37.124 - (Hash{|Number SID, Nonce M,
37.125 - Nonce NA, Number PA, Agent A,
37.126 - Nonce NB, Number PB, Agent B|}))
37.127 + (Hash{|Number SID, Nonce M,
37.128 + Nonce NA, Number PA, Agent A,
37.129 + Nonce NB, Number PB, Agent B|}))
37.130 # evsSF \<in> tls"
37.131
37.132 | ClientAccepts:
37.133 - --{*Having transmitted ClientFinished and received an identical
37.134 + --{*Having transmitted ClientFinished and received an identical
37.135 message encrypted with serverK, the client stores the parameters
37.136 needed to resume this session. The "Notes A ..." premise is
37.137 used to prove @{text Notes_master_imp_Crypt_PMS}.*}
37.138 "[| evsCA \<in> tls;
37.139 - Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
37.140 - M = PRF(PMS,NA,NB);
37.141 - X = Hash{|Number SID, Nonce M,
37.142 - Nonce NA, Number PA, Agent A,
37.143 - Nonce NB, Number PB, Agent B|};
37.144 + Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
37.145 + M = PRF(PMS,NA,NB);
37.146 + X = Hash{|Number SID, Nonce M,
37.147 + Nonce NA, Number PA, Agent A,
37.148 + Nonce NB, Number PB, Agent B|};
37.149 Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evsCA;
37.150 Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA |]
37.151 ==>
37.152 Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA \<in> tls"
37.153
37.154 | ServerAccepts:
37.155 - --{*Having transmitted ServerFinished and received an identical
37.156 + --{*Having transmitted ServerFinished and received an identical
37.157 message encrypted with clientK, the server stores the parameters
37.158 needed to resume this session. The "Says A'' B ..." premise is
37.159 used to prove @{text Notes_master_imp_Crypt_PMS}.*}
37.160 "[| evsSA \<in> tls;
37.161 - A \<noteq> B;
37.162 + A \<noteq> B;
37.163 Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
37.164 - M = PRF(PMS,NA,NB);
37.165 - X = Hash{|Number SID, Nonce M,
37.166 - Nonce NA, Number PA, Agent A,
37.167 - Nonce NB, Number PB, Agent B|};
37.168 + M = PRF(PMS,NA,NB);
37.169 + X = Hash{|Number SID, Nonce M,
37.170 + Nonce NA, Number PA, Agent A,
37.171 + Nonce NB, Number PB, Agent B|};
37.172 Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evsSA;
37.173 Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA |]
37.174 ==>
37.175 @@ -244,27 +244,27 @@
37.176 --{*If A recalls the @{text SESSION_ID}, then she sends a FINISHED
37.177 message using the new nonces and stored MASTER SECRET.*}
37.178 "[| evsCR \<in> tls;
37.179 - Says A B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
37.180 + Says A B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
37.181 Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCR;
37.182 Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsCR |]
37.183 ==> Says A B (Crypt (clientK(NA,NB,M))
37.184 - (Hash{|Number SID, Nonce M,
37.185 - Nonce NA, Number PA, Agent A,
37.186 - Nonce NB, Number PB, Agent B|}))
37.187 + (Hash{|Number SID, Nonce M,
37.188 + Nonce NA, Number PA, Agent A,
37.189 + Nonce NB, Number PB, Agent B|}))
37.190 # evsCR \<in> tls"
37.191
37.192 | ServerResume:
37.193 --{*Resumption (7.3): If B finds the @{text SESSION_ID} then he can
37.194 send a FINISHED message using the recovered MASTER SECRET*}
37.195 "[| evsSR \<in> tls;
37.196 - Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
37.197 - Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
37.198 + Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
37.199 + Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
37.200 Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsSR |]
37.201 ==> Says B A (Crypt (serverK(NA,NB,M))
37.202 - (Hash{|Number SID, Nonce M,
37.203 - Nonce NA, Number PA, Agent A,
37.204 - Nonce NB, Number PB, Agent B|})) # evsSR
37.205 - \<in> tls"
37.206 + (Hash{|Number SID, Nonce M,
37.207 + Nonce NA, Number PA, Agent A,
37.208 + Nonce NB, Number PB, Agent B|})) # evsSR
37.209 + \<in> tls"
37.210
37.211 | Oops:
37.212 --{*The most plausible compromise is of an old session key. Losing
37.213 @@ -273,7 +273,7 @@
37.214 otherwise the Spy could learn session keys merely by
37.215 replaying messages!*}
37.216 "[| evso \<in> tls; A \<noteq> Spy;
37.217 - Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
37.218 + Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
37.219 ==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso \<in> tls"
37.220
37.221 (*
37.222 @@ -328,7 +328,7 @@
37.223
37.224
37.225 (** These proofs assume that the Nonce_supply nonces
37.226 - (which have the form @ N. Nonce N \<notin> used evs)
37.227 + (which have the form @ N. Nonce N \<notin> used evs)
37.228 lie outside the range of PRF. It seems reasonable, but as it is needed
37.229 only for the possibility theorems, it is not taken as an axiom.
37.230 **)
37.231 @@ -381,11 +381,11 @@
37.232 \<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;
37.233 A \<noteq> B |]
37.234 ==> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
37.235 - X = Hash{|Number SID, Nonce M,
37.236 - Nonce NA, Number PA, Agent A,
37.237 - Nonce NB, Number PB, Agent B|} &
37.238 - Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs &
37.239 - Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
37.240 + X = Hash{|Number SID, Nonce M,
37.241 + Nonce NA, Number PA, Agent A,
37.242 + Nonce NB, Number PB, Agent B|} &
37.243 + Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs &
37.244 + Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
37.245 apply (intro exI bexI)
37.246 apply (rule_tac [2] tls.ClientHello
37.247 [THEN tls.ServerHello,
37.248 @@ -570,7 +570,7 @@
37.249 (priK B \<in> KK | B \<in> bad)"
37.250 apply (erule tls.induct)
37.251 apply (simp_all (no_asm_simp)
37.252 - del: image_insert
37.253 + del: image_insert
37.254 add: image_Un [THEN sym]
37.255 insert_Key_image Un_assoc [THEN sym])
37.256 txt{*Fake*}
37.257 @@ -598,16 +598,16 @@
37.258 lemma analz_image_keys [rule_format]:
37.259 "evs \<in> tls ==>
37.260 \<forall>KK. KK <= range sessionK -->
37.261 - (Nonce N \<in> analz (Key`KK Un (spies evs))) =
37.262 - (Nonce N \<in> analz (spies evs))"
37.263 + (Nonce N \<in> analz (Key`KK Un (spies evs))) =
37.264 + (Nonce N \<in> analz (spies evs))"
37.265 apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
37.266 apply (safe del: iffI)
37.267 apply (safe del: impI iffI intro!: analz_image_keys_lemma)
37.268 apply (simp_all (no_asm_simp) (*faster*)
37.269 del: image_insert imp_disjL (*reduces blow-up*)
37.270 - add: image_Un [THEN sym] Un_assoc [THEN sym]
37.271 - insert_Key_singleton
37.272 - range_sessionkeys_not_priK analz_image_priK)
37.273 + add: image_Un [THEN sym] Un_assoc [THEN sym]
37.274 + insert_Key_singleton
37.275 + range_sessionkeys_not_priK analz_image_priK)
37.276 apply (simp_all add: insert_absorb)
37.277 txt{*Fake*}
37.278 apply spy_analz
37.279 @@ -901,11 +901,11 @@
37.280 down to 433s on albatross*)
37.281
37.282 (*5/5/01: conversion to Isar script
37.283 - loads in 137s (perch)
37.284 + loads in 137s (perch)
37.285 the last ML version loaded in 122s on perch, a 600MHz machine:
37.286 - twice as fast as pike. No idea why it's so much slower!
37.287 - The Isar script is slower still, perhaps because simp_all simplifies
37.288 - the assumptions be default.
37.289 + twice as fast as pike. No idea why it's so much slower!
37.290 + The Isar script is slower still, perhaps because simp_all simplifies
37.291 + the assumptions be default.
37.292 *)
37.293
37.294 end
38.1 --- a/src/HOL/Auth/Yahalom.thy Sat Oct 17 01:05:59 2009 +0200
38.2 +++ b/src/HOL/Auth/Yahalom.thy Sat Oct 17 14:43:18 2009 +0200
38.3 @@ -58,7 +58,7 @@
38.4 uses the new session key to send Bob his Nonce. The premise
38.5 @{term "A \<noteq> Server"} is needed for @{text Says_Server_not_range}.
38.6 Alice can check that K is symmetric by its length.*}
38.7 - "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys;
38.8 + "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys;
38.9 Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
38.10 \<in> set evs4;
38.11 Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
38.12 @@ -481,9 +481,9 @@
38.13 text{*A vital theorem for B, that nonce NB remains secure from the Spy.*}
38.14 lemma Spy_not_see_NB :
38.15 "[| Says B Server
38.16 - {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
38.17 - \<in> set evs;
38.18 - (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
38.19 + {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
38.20 + \<in> set evs;
38.21 + (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
38.22 A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
38.23 ==> Nonce NB \<notin> analz (knows Spy evs)"
38.24 apply (erule rev_mp, erule rev_mp)
39.1 --- a/src/HOL/Auth/Yahalom2.thy Sat Oct 17 01:05:59 2009 +0200
39.2 +++ b/src/HOL/Auth/Yahalom2.thy Sat Oct 17 14:43:18 2009 +0200
39.3 @@ -51,7 +51,7 @@
39.4 Both agents are quoted in the 2nd certificate to prevent attacks!*)
39.5 | YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3;
39.6 Gets Server {|Agent B, Nonce NB,
39.7 - Crypt (shrK B) {|Agent A, Nonce NA|}|}
39.8 + Crypt (shrK B) {|Agent A, Nonce NA|}|}
39.9 \<in> set evs3 |]
39.10 ==> Says Server A
39.11 {|Nonce NB,
40.1 --- a/src/HOL/Auth/ZhouGollmann.thy Sat Oct 17 01:05:59 2009 +0200
40.2 +++ b/src/HOL/Auth/ZhouGollmann.thy Sat Oct 17 14:43:18 2009 +0200
40.3 @@ -35,7 +35,7 @@
40.4 Nil: "[] \<in> zg"
40.5
40.6 | Fake: "[| evsf \<in> zg; X \<in> synth (analz (spies evsf)) |]
40.7 - ==> Says Spy B X # evsf \<in> zg"
40.8 + ==> Says Spy B X # evsf \<in> zg"
40.9
40.10 | Reception: "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
40.11
40.12 @@ -44,26 +44,26 @@
40.13 We just assume that the protocol's objective is to deliver K fairly,
40.14 rather than to keep M secret.*)
40.15 | ZG1: "[| evs1 \<in> zg; Nonce L \<notin> used evs1; C = Crypt K (Number m);
40.16 - K \<in> symKeys;
40.17 - NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
40.18 + K \<in> symKeys;
40.19 + NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
40.20 ==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg"
40.21
40.22 (*B must check that NRO is A's signature to learn the sender's name*)
40.23 | ZG2: "[| evs2 \<in> zg;
40.24 - Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
40.25 - NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
40.26 - NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
40.27 + Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
40.28 + NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
40.29 + NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
40.30 ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} # evs2 \<in> zg"
40.31
40.32 (*A must check that NRR is B's signature to learn the sender's name;
40.33 without spy, the matching label would be enough*)
40.34 | ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
40.35 - Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
40.36 - Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
40.37 - NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
40.38 - sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
40.39 + Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
40.40 + Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
40.41 + NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
40.42 + sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
40.43 ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
40.44 - # evs3 \<in> zg"
40.45 + # evs3 \<in> zg"
40.46
40.47 (*TTP checks that sub_K is A's signature to learn who issued K, then
40.48 gives credentials to A and B. The Notes event models the availability of
40.49 @@ -72,15 +72,15 @@
40.50 also allowing lemma @{text Crypt_used_imp_spies} to omit the condition
40.51 @{term "K \<noteq> priK TTP"}. *)
40.52 | ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
40.53 - Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
40.54 - \<in> set evs4;
40.55 - sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
40.56 - con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
40.57 - Nonce L, Key K|}|]
40.58 + Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
40.59 + \<in> set evs4;
40.60 + sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
40.61 + con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
40.62 + Nonce L, Key K|}|]
40.63 ==> Says TTP Spy con_K
40.64 #
40.65 - Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
40.66 - # evs4 \<in> zg"
40.67 + Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
40.68 + # evs4 \<in> zg"
40.69
40.70
40.71 declare Says_imp_knows_Spy [THEN analz.Inj, dest]
41.1 --- a/src/HOL/Bali/AxCompl.thy Sat Oct 17 01:05:59 2009 +0200
41.2 +++ b/src/HOL/Bali/AxCompl.thy Sat Oct 17 14:43:18 2009 +0200
41.3 @@ -319,13 +319,13 @@
41.4 case 0
41.5 with is_cls
41.6 show ?thesis
41.7 - by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
41.8 + by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
41.9 next
41.10 case (Suc m)
41.11 with mgf_hyp have mgf_hyp': "\<And> t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}"
41.12 - by simp
41.13 + by simp
41.14 from is_cls obtain c where c: "the (class G C) = c"
41.15 - by auto
41.16 + by auto
41.17 let ?Q= "(\<lambda>Y s' (x,s) .
41.18 G\<turnstile> (x,init_class_obj G C s)
41.19 \<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s'
41.20 @@ -333,45 +333,45 @@
41.21 from c
41.22 show ?thesis
41.23 proof (rule ax_derivs.Init [where ?Q="?Q"])
41.24 - let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s
41.25 + let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s
41.26 \<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>m)"
41.27 - show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
41.28 + show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
41.29 .(if C = Object then Skip else Init (super c)).
41.30 {?Q}"
41.31 - proof (rule conseq1 [where ?P'="?P'"])
41.32 - show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
41.33 - proof (cases "C=Object")
41.34 - case True
41.35 - have "G,A\<turnstile>{?P'} .Skip. {?Q}"
41.36 - by (rule ax_derivs.Skip [THEN conseq1])
41.37 - (auto simp add: True intro: eval.Skip)
41.38 + proof (rule conseq1 [where ?P'="?P'"])
41.39 + show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
41.40 + proof (cases "C=Object")
41.41 + case True
41.42 + have "G,A\<turnstile>{?P'} .Skip. {?Q}"
41.43 + by (rule ax_derivs.Skip [THEN conseq1])
41.44 + (auto simp add: True intro: eval.Skip)
41.45 with True show ?thesis
41.46 - by simp
41.47 - next
41.48 - case False
41.49 - from mgf_hyp'
41.50 - have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
41.51 - by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
41.52 - with False show ?thesis
41.53 - by simp
41.54 - qed
41.55 - next
41.56 - from Suc is_cls
41.57 - show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
41.58 + by simp
41.59 + next
41.60 + case False
41.61 + from mgf_hyp'
41.62 + have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
41.63 + by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
41.64 + with False show ?thesis
41.65 + by simp
41.66 + qed
41.67 + next
41.68 + from Suc is_cls
41.69 + show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
41.70 \<Rightarrow> ?P'"
41.71 - by (fastsimp elim: nyinitcls_le_SucD)
41.72 - qed
41.73 + by (fastsimp elim: nyinitcls_le_SucD)
41.74 + qed
41.75 next
41.76 - from mgf_hyp'
41.77 - show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty}
41.78 + from mgf_hyp'
41.79 + show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty}
41.80 .init c.
41.81 {set_lvars l .; ?R}"
41.82 - apply (rule MGFnD' [THEN conseq12, THEN allI])
41.83 - apply (clarsimp simp add: split_paired_all)
41.84 - apply (rule eval.Init [OF c])
41.85 - apply (insert c)
41.86 - apply auto
41.87 - done
41.88 + apply (rule MGFnD' [THEN conseq12, THEN allI])
41.89 + apply (clarsimp simp add: split_paired_all)
41.90 + apply (rule eval.Init [OF c])
41.91 + apply (insert c)
41.92 + apply auto
41.93 + done
41.94 qed
41.95 qed
41.96 thus "G,A\<turnstile>{Normal ?P \<and>. Not \<circ> initd C} .Init C. {?R}"
41.97 @@ -399,7 +399,7 @@
41.98 mode: "mode = invmode statM e" and
41.99 T: "T =Inl (resTy statM)" and
41.100 eq_accC_accC': "accC=accC'"
41.101 - by cases fastsimp+
41.102 + by cases fastsimp+
41.103 let ?Q="(\<lambda>Y s1 (x,s) . x = None \<and>
41.104 (\<exists>a. G\<turnstile>Norm s \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
41.105 (normal s1 \<longrightarrow> G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT)
41.106 @@ -435,29 +435,29 @@
41.107 (\<exists>P. \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P))
41.108 \<and> s1\<Colon>\<preceq>(G, L)"
41.109 proof -
41.110 - from da obtain C where
41.111 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
41.112 + from da obtain C where
41.113 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
41.114 dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
41.115 - da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E"
41.116 - by cases (simp add: eq_accC_accC')
41.117 - from eval_e conf_s0 wt_e da_e wf
41.118 - obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
41.119 - and "s1\<Colon>\<preceq>(G, L)"
41.120 - by (rule eval_type_soundE) simp
41.121 - moreover
41.122 - {
41.123 - assume normal_s1: "normal s1"
41.124 - have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
41.125 - proof -
41.126 - from eval_e wt_e da_e wf normal_s1
41.127 - have "nrm C \<subseteq> dom (locals (store s1))"
41.128 - by (cases rule: da_good_approxE') iprover
41.129 - with da_ps show ?thesis
41.130 - by (rule da_weakenE) iprover
41.131 - qed
41.132 - }
41.133 - ultimately show ?thesis
41.134 - using eq_accC_accC' by simp
41.135 + da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E"
41.136 + by cases (simp add: eq_accC_accC')
41.137 + from eval_e conf_s0 wt_e da_e wf
41.138 + obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
41.139 + and "s1\<Colon>\<preceq>(G, L)"
41.140 + by (rule eval_type_soundE) simp
41.141 + moreover
41.142 + {
41.143 + assume normal_s1: "normal s1"
41.144 + have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
41.145 + proof -
41.146 + from eval_e wt_e da_e wf normal_s1
41.147 + have "nrm C \<subseteq> dom (locals (store s1))"
41.148 + by (cases rule: da_good_approxE') iprover
41.149 + with da_ps show ?thesis
41.150 + by (rule da_weakenE) iprover
41.151 + qed
41.152 + }
41.153 + ultimately show ?thesis
41.154 + using eq_accC_accC' by simp
41.155 qed
41.156 qed
41.157 next
41.158 @@ -467,36 +467,36 @@
41.159 show "?PS a"
41.160 proof (rule MGFnD' [OF mgf_ps, THEN conseq12],
41.161 clarsimp simp add: eq_accC_accC' [symmetric])
41.162 - fix s0 s1 s2 vs
41.163 - assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
41.164 - assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
41.165 - assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
41.166 - assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
41.167 - assume da_ps: "abrupt s1 = None \<longrightarrow>
41.168 + fix s0 s1 s2 vs
41.169 + assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
41.170 + assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
41.171 + assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
41.172 + assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
41.173 + assume da_ps: "abrupt s1 = None \<longrightarrow>
41.174 (\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
41.175 dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P)"
41.176 - show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
41.177 + show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
41.178 (abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT) \<and>
41.179 G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2) \<and>
41.180 s2\<Colon>\<preceq>(G, L)"
41.181 - proof (cases "normal s1")
41.182 - case True
41.183 - with da_ps obtain P where
41.184 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
41.185 - by auto
41.186 - from eval_ps conf_s1 wt_args this wf
41.187 - have "s2\<Colon>\<preceq>(G, L)"
41.188 - by (rule eval_type_soundE)
41.189 - with eval_e conf_a eval_ps
41.190 - show ?thesis
41.191 - by auto
41.192 - next
41.193 - case False
41.194 - with eval_ps have "s2=s1" by auto
41.195 - with eval_e conf_a eval_ps conf_s1
41.196 - show ?thesis
41.197 - by auto
41.198 - qed
41.199 + proof (cases "normal s1")
41.200 + case True
41.201 + with da_ps obtain P where
41.202 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
41.203 + by auto
41.204 + from eval_ps conf_s1 wt_args this wf
41.205 + have "s2\<Colon>\<preceq>(G, L)"
41.206 + by (rule eval_type_soundE)
41.207 + with eval_e conf_a eval_ps
41.208 + show ?thesis
41.209 + by auto
41.210 + next
41.211 + case False
41.212 + with eval_ps have "s2=s1" by auto
41.213 + with eval_e conf_a eval_ps conf_s1
41.214 + show ?thesis
41.215 + by auto
41.216 + qed
41.217 qed
41.218 qed
41.219 next
41.220 @@ -517,52 +517,52 @@
41.221 from mgf_methds [rule_format]
41.222 show "?METHD a vs invC declC l"
41.223 proof (rule MGFnD' [THEN conseq12],clarsimp)
41.224 - fix s4 s2 s1::state
41.225 - fix s0 v
41.226 - let ?D= "invocation_declclass G mode (store s2) a statT
41.227 + fix s4 s2 s1::state
41.228 + fix s0 v
41.229 + let ?D= "invocation_declclass G mode (store s2) a statT
41.230 \<lparr>name=mn,parTs=pTs'\<rparr>"
41.231 - let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
41.232 - assume inv_prop: "abrupt ?s3=None
41.233 + let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
41.234 + assume inv_prop: "abrupt ?s3=None
41.235 \<longrightarrow> G\<turnstile>mode\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
41.236 - assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
41.237 - assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
41.238 - assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
41.239 - assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
41.240 - assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
41.241 - show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
41.242 + assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
41.243 + assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
41.244 + assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
41.245 + assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
41.246 + assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
41.247 + show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
41.248 \<rightarrow> (set_lvars (locals (store s2))) s4"
41.249 - proof -
41.250 - obtain D where D: "D=?D" by simp
41.251 - obtain s3 where s3: "s3=?s3" by simp
41.252 - obtain s3' where
41.253 - s3': "s3' = check_method_access G accC statT mode
41.254 + proof -
41.255 + obtain D where D: "D=?D" by simp
41.256 + obtain s3 where s3: "s3=?s3" by simp
41.257 + obtain s3' where
41.258 + s3': "s3' = check_method_access G accC statT mode
41.259 \<lparr>name=mn,parTs=pTs'\<rparr> a s3"
41.260 - by simp
41.261 - have eq_s3'_s3: "s3'=s3"
41.262 - proof -
41.263 - from inv_prop s3 mode
41.264 - have "normal s3 \<Longrightarrow>
41.265 + by simp
41.266 + have eq_s3'_s3: "s3'=s3"
41.267 + proof -
41.268 + from inv_prop s3 mode
41.269 + have "normal s3 \<Longrightarrow>
41.270 G\<turnstile>invmode statM e\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
41.271 - by auto
41.272 - with eval_ps wt_e statM conf_s2 conf_a [rule_format]
41.273 - have "check_method_access G accC statT (invmode statM e)
41.274 + by auto
41.275 + with eval_ps wt_e statM conf_s2 conf_a [rule_format]
41.276 + have "check_method_access G accC statT (invmode statM e)
41.277 \<lparr>name=mn,parTs=pTs'\<rparr> a s3 = s3"
41.278 - by (rule error_free_call_access) (auto simp add: s3 mode wf)
41.279 - thus ?thesis
41.280 - by (simp add: s3' mode)
41.281 - qed
41.282 - with eval_mthd D s3
41.283 - have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
41.284 - by simp
41.285 - with eval_e eval_ps D _ s3'
41.286 - show ?thesis
41.287 - by (rule eval_Call) (auto simp add: s3 mode D)
41.288 - qed
41.289 + by (rule error_free_call_access) (auto simp add: s3 mode wf)
41.290 + thus ?thesis
41.291 + by (simp add: s3' mode)
41.292 + qed
41.293 + with eval_mthd D s3
41.294 + have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
41.295 + by simp
41.296 + with eval_e eval_ps D _ s3'
41.297 + show ?thesis
41.298 + by (rule eval_Call) (auto simp add: s3 mode D)
41.299 + qed
41.300 qed
41.301 qed
41.302 qed
41.303 qed
41.304 -
41.305 +
41.306 lemma eval_expression_no_jump':
41.307 assumes eval: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<rightarrow> s1"
41.308 and no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
41.309 @@ -610,36 +610,36 @@
41.310 case True
41.311 with normal_t eval_e normal_termination
41.312 show ?thesis
41.313 - by (auto intro: eval.Loop)
41.314 + by (auto intro: eval.Loop)
41.315 next
41.316 case False
41.317 note abrupt_s' = this
41.318 from eval_e _ wt wf
41.319 have no_cont: "abrupt s' \<noteq> Some (Jump (Cont l))"
41.320 - by (rule eval_expression_no_jump') (insert normal_t,simp)
41.321 + by (rule eval_expression_no_jump') (insert normal_t,simp)
41.322 have
41.323 - "if the_Bool v
41.324 + "if the_Bool v
41.325 then (G\<turnstile>s' \<midarrow>c\<rightarrow> s' \<and>
41.326 G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s')
41.327 - else s' = s'"
41.328 + else s' = s'"
41.329 proof (cases "the_Bool v")
41.330 - case False thus ?thesis by simp
41.331 + case False thus ?thesis by simp
41.332 next
41.333 - case True
41.334 - with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
41.335 - moreover from abrupt_s' no_cont
41.336 - have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
41.337 - by (cases s') (simp add: absorb_def split: split_if)
41.338 - moreover
41.339 - from no_absorb abrupt_s'
41.340 - have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
41.341 - by auto
41.342 - ultimately show ?thesis
41.343 - using True by simp
41.344 + case True
41.345 + with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
41.346 + moreover from abrupt_s' no_cont
41.347 + have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
41.348 + by (cases s') (simp add: absorb_def split: split_if)
41.349 + moreover
41.350 + from no_absorb abrupt_s'
41.351 + have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
41.352 + by auto
41.353 + ultimately show ?thesis
41.354 + using True by simp
41.355 qed
41.356 with eval_e
41.357 show ?thesis
41.358 - using normal_t by (auto intro: eval.Loop)
41.359 + using normal_t by (auto intro: eval.Loop)
41.360 qed
41.361 qed
41.362 next
41.363 @@ -690,49 +690,49 @@
41.364 Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s'))
41.365 \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>}"
41.366 proof (rule ax_derivs.Loop)
41.367 - from mfg_e
41.368 - show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n}
41.369 + from mfg_e
41.370 + show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n}
41.371 e-\<succ>
41.372 {(\<lambda>Y s' s. (\<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
41.373 Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s'))
41.374 \<and>. G\<turnstile>init\<le>n}"
41.375 - proof (rule MGFnD' [THEN conseq12],clarsimp)
41.376 - fix s Z s' v
41.377 - assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
41.378 - moreover
41.379 - assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
41.380 - ultimately
41.381 - show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
41.382 - by blast
41.383 - qed
41.384 + proof (rule MGFnD' [THEN conseq12],clarsimp)
41.385 + fix s Z s' v
41.386 + assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
41.387 + moreover
41.388 + assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
41.389 + ultimately
41.390 + show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
41.391 + by blast
41.392 + qed
41.393 next
41.394 - from mfg_c
41.395 - show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
41.396 + from mfg_c
41.397 + show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
41.398 Y = \<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')
41.399 \<and>. G\<turnstile>init\<le>n)\<leftarrow>=True)}
41.400 .c.
41.401 {abupd (absorb (Cont l)) .;
41.402 ((\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n)}"
41.403 - proof (rule MGFnD' [THEN conseq12],clarsimp)
41.404 - fix Z s' s v t
41.405 - assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
41.406 - assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s"
41.407 - assume true: "the_Bool v"
41.408 - assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
41.409 - show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
41.410 - proof -
41.411 - note unroll
41.412 - also
41.413 - from eval_e true eval_c
41.414 - have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c"
41.415 - by (unfold unroll_def) force
41.416 - ultimately show ?thesis ..
41.417 - qed
41.418 - qed
41.419 + proof (rule MGFnD' [THEN conseq12],clarsimp)
41.420 + fix Z s' s v t
41.421 + assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
41.422 + assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s"
41.423 + assume true: "the_Bool v"
41.424 + assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
41.425 + show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
41.426 + proof -
41.427 + note unroll
41.428 + also
41.429 + from eval_e true eval_c
41.430 + have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c"
41.431 + by (unfold unroll_def) force
41.432 + ultimately show ?thesis ..
41.433 + qed
41.434 + qed
41.435 qed
41.436 next
41.437 show
41.438 - "\<forall>Y s Z.
41.439 + "\<forall>Y s Z.
41.440 (Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)) Y s Z
41.441 \<longrightarrow> (\<forall>Y' s'.
41.442 (\<forall>Y Z'.
41.443 @@ -742,28 +742,28 @@
41.444 \<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>) Y' s' Z')
41.445 \<longrightarrow> G\<turnstile>Z \<midarrow>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ>\<rightarrow> (Y', s'))"
41.446 proof (clarsimp)
41.447 - fix Y' s' s
41.448 - assume asm:
41.449 - "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>*
41.450 + fix Y' s' s
41.451 + assume asm:
41.452 + "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>*
41.453 \<longrightarrow> card (nyinitcls G s') \<le> n \<and>
41.454 (\<exists>v. (\<exists>t. (Z', t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s') \<and>
41.455 (fst s' = None \<longrightarrow> \<not> the_Bool v)) \<and> Y' = \<diamondsuit>"
41.456 - show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
41.457 - proof -
41.458 - from asm obtain v t where
41.459 - -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}
41.460 - unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
41.461 + show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
41.462 + proof -
41.463 + from asm obtain v t where
41.464 + -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}
41.465 + unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
41.466 eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" and
41.467 normal_termination: "normal s' \<longrightarrow> \<not> the_Bool v" and
41.468 - Y': "Y' = \<diamondsuit>"
41.469 - by auto
41.470 - from unroll eval_e normal_termination wt_e wf
41.471 - have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
41.472 - by (rule unroll_while)
41.473 - with Y'
41.474 - show ?thesis
41.475 - by simp
41.476 - qed
41.477 + Y': "Y' = \<diamondsuit>"
41.478 + by auto
41.479 + from unroll eval_e normal_termination wt_e wf
41.480 + have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
41.481 + by (rule unroll_while)
41.482 + with Y'
41.483 + show ?thesis
41.484 + by simp
41.485 + qed
41.486 qed
41.487 qed
41.488 qed
41.489 @@ -810,39 +810,39 @@
41.490 show "(\<exists>E. \<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E) \<and>
41.491 s'\<Colon>\<preceq>(G, L)"
41.492 proof -
41.493 - from da
41.494 - obtain E where
41.495 - "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
41.496 - by cases simp
41.497 - moreover
41.498 - from eval_init
41.499 - have "dom (locals s) \<subseteq> dom (locals (store s'))"
41.500 - by (rule dom_locals_eval_mono [elim_format]) simp
41.501 - ultimately obtain E' where
41.502 - "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
41.503 - by (rule da_weakenE)
41.504 - moreover
41.505 - have "s'\<Colon>\<preceq>(G, L)"
41.506 - proof -
41.507 - have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
41.508 - proof -
41.509 - from wf wt_e
41.510 - have iscls_statC: "is_class G statC"
41.511 - by (auto dest: ty_expr_is_type type_is_class)
41.512 - with wf accfield
41.513 - have iscls_statDeclC: "is_class G statDeclC"
41.514 - by (auto dest!: accfield_fields dest: fields_declC)
41.515 - thus ?thesis by simp
41.516 - qed
41.517 - obtain I where
41.518 - da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
41.519 + from da
41.520 + obtain E where
41.521 + "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
41.522 + by cases simp
41.523 + moreover
41.524 + from eval_init
41.525 + have "dom (locals s) \<subseteq> dom (locals (store s'))"
41.526 + by (rule dom_locals_eval_mono [elim_format]) simp
41.527 + ultimately obtain E' where
41.528 + "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
41.529 + by (rule da_weakenE)
41.530 + moreover
41.531 + have "s'\<Colon>\<preceq>(G, L)"
41.532 + proof -
41.533 + have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
41.534 + proof -
41.535 + from wf wt_e
41.536 + have iscls_statC: "is_class G statC"
41.537 + by (auto dest: ty_expr_is_type type_is_class)
41.538 + with wf accfield
41.539 + have iscls_statDeclC: "is_class G statDeclC"
41.540 + by (auto dest!: accfield_fields dest: fields_declC)
41.541 + thus ?thesis by simp
41.542 + qed
41.543 + obtain I where
41.544 + da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
41.545 \<turnstile> dom (locals (store ((Norm s)::state))) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
41.546 - by (auto intro: da_Init [simplified] assigned.select_convs)
41.547 - from eval_init conf_s wt_init da_init wf
41.548 - show ?thesis
41.549 - by (rule eval_type_soundE)
41.550 - qed
41.551 - ultimately show ?thesis by iprover
41.552 + by (auto intro: da_Init [simplified] assigned.select_convs)
41.553 + from eval_init conf_s wt_init da_init wf
41.554 + show ?thesis
41.555 + by (rule eval_type_soundE)
41.556 + qed
41.557 + ultimately show ?thesis by iprover
41.558 qed
41.559 qed
41.560 next
41.561 @@ -857,30 +857,30 @@
41.562 assume da_e: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
41.563 show "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>fst ?fvar\<rightarrow> snd ?fvar"
41.564 proof -
41.565 - obtain v s2' where
41.566 - v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
41.567 - by simp
41.568 - obtain s3 where
41.569 - s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
41.570 - by simp
41.571 - have eq_s3_s2': "s3=s2'"
41.572 - proof -
41.573 - from eval_e conf_s1 wt_e da_e wf obtain
41.574 - conf_s2: "s2\<Colon>\<preceq>(G, L)" and
41.575 - conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
41.576 - by (rule eval_type_soundE) simp
41.577 - from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
41.578 - show ?thesis
41.579 - by (rule error_free_field_access
41.580 + obtain v s2' where
41.581 + v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
41.582 + by simp
41.583 + obtain s3 where
41.584 + s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
41.585 + by simp
41.586 + have eq_s3_s2': "s3=s2'"
41.587 + proof -
41.588 + from eval_e conf_s1 wt_e da_e wf obtain
41.589 + conf_s2: "s2\<Colon>\<preceq>(G, L)" and
41.590 + conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
41.591 + by (rule eval_type_soundE) simp
41.592 + from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
41.593 + show ?thesis
41.594 + by (rule error_free_field_access
41.595 [where ?v=v and ?s2'=s2',elim_format])
41.596 - (simp add: s3 v s2' stat)+
41.597 + (simp add: s3 v s2' stat)+
41.598 qed
41.599 - from eval_init eval_e
41.600 - show ?thesis
41.601 - apply (rule eval.FVar [where ?s2'=s2'])
41.602 - apply (simp add: s2')
41.603 - apply (simp add: s3 [symmetric] eq_s3_s2' eq_accC s2' [symmetric])
41.604 - done
41.605 + from eval_init eval_e
41.606 + show ?thesis
41.607 + apply (rule eval.FVar [where ?s2'=s2'])
41.608 + apply (simp add: s2')
41.609 + apply (simp add: s3 [symmetric] eq_s3_s2' eq_accC s2' [symmetric])
41.610 + done
41.611 qed
41.612 qed
41.613 qed
41.614 @@ -918,7 +918,7 @@
41.615 fix s0
41.616 assume "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> C"
41.617 thus "\<exists>C1. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1"
41.618 - by cases (auto simp add: inj_term_simps)
41.619 + by cases (auto simp add: inj_term_simps)
41.620 qed
41.621 next
41.622 from mgf_c2
41.623 @@ -933,25 +933,25 @@
41.624 show "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2
41.625 \<rightarrow> abupd (abrupt_if (\<exists>y. abrupt s1 = Some y) (abrupt s1)) s2"
41.626 proof -
41.627 - obtain abr1 str1 where s1: "s1=(abr1,str1)"
41.628 - by (cases s1)
41.629 - with eval_c1 eval_c2 obtain
41.630 - eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
41.631 - eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
41.632 - by simp
41.633 - obtain s3 where
41.634 - s3: "s3 = (if \<exists>err. abr1 = Some (Error err)
41.635 - then (abr1, str1)
41.636 + obtain abr1 str1 where s1: "s1=(abr1,str1)"
41.637 + by (cases s1)
41.638 + with eval_c1 eval_c2 obtain
41.639 + eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
41.640 + eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
41.641 + by simp
41.642 + obtain s3 where
41.643 + s3: "s3 = (if \<exists>err. abr1 = Some (Error err)
41.644 + then (abr1, str1)
41.645 else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
41.646 - by simp
41.647 - from eval_c1' conf_s0 wt_c1 _ wf
41.648 - have "error_free (abr1,str1)"
41.649 - by (rule eval_type_soundE) (insert da_c1,auto)
41.650 - with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
41.651 - by (simp add: error_free_def)
41.652 - from eval_c1' eval_c2' s3
41.653 - show ?thesis
41.654 - by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
41.655 + by simp
41.656 + from eval_c1' conf_s0 wt_c1 _ wf
41.657 + have "error_free (abr1,str1)"
41.658 + by (rule eval_type_soundE) (insert da_c1,auto)
41.659 + with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
41.660 + by (simp add: error_free_def)
41.661 + from eval_c1' eval_c2' s3
41.662 + show ?thesis
41.663 + by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
41.664 qed
41.665 qed
41.666 qed
41.667 @@ -1009,7 +1009,7 @@
41.668 fix s0
41.669 assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright> E"
41.670 thus "jumpNestingOkS {Ret} c"
41.671 - by cases simp
41.672 + by cases simp
41.673 qed
41.674 next
41.675 from mgf_c
41.676 @@ -1022,22 +1022,22 @@
41.677 show "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
41.678 \<rightarrow> abupd (absorb Ret) s2"
41.679 proof -
41.680 - from wt obtain d where
41.681 + from wt obtain d where
41.682 d: "class G D=Some d" and
41.683 wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
41.684 - by cases auto
41.685 - obtain s3 where
41.686 - s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
41.687 + by cases auto
41.688 + obtain s3 where
41.689 + s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
41.690 fst s2 = Some (Jump (Cont l))
41.691 then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2
41.692 else s2)"
41.693 - by simp
41.694 - from eval_init eval_c nestingOk wt_c d wf
41.695 - have eq_s3_s2: "s3=s2"
41.696 - by (rule Body_no_break [elim_format]) (simp add: s3)
41.697 - from eval_init eval_c s3
41.698 - show ?thesis
41.699 - by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
41.700 + by simp
41.701 + from eval_init eval_c nestingOk wt_c d wf
41.702 + have eq_s3_s2: "s3=s2"
41.703 + by (rule Body_no_break [elim_format]) (simp add: s3)
41.704 + from eval_init eval_c s3
41.705 + show ?thesis
41.706 + by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
41.707 qed
41.708 qed
41.709 qed
41.710 @@ -1062,312 +1062,312 @@
41.711 proof (induct rule: var_expr_stmt.inducts)
41.712 case (LVar v)
41.713 show "G,A\<turnstile>{=:n} \<langle>LVar v\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
41.714 - apply (rule MGFn_NormalI)
41.715 - apply (rule ax_derivs.LVar [THEN conseq1])
41.716 - apply (clarsimp)
41.717 - apply (rule eval.LVar)
41.718 - done
41.719 + apply (rule MGFn_NormalI)
41.720 + apply (rule ax_derivs.LVar [THEN conseq1])
41.721 + apply (clarsimp)
41.722 + apply (rule eval.LVar)
41.723 + done
41.724 next
41.725 case (FVar accC statDeclC stat e fn)
41.726 from MGFn_Init [OF hyp] and `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}` and wf
41.727 show ?case
41.728 - by (rule MGFn_FVar)
41.729 + by (rule MGFn_FVar)
41.730 next
41.731 case (AVar e1 e2)
41.732 note mgf_e1 = `G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
41.733 note mgf_e2 = `G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
41.734 show "G,A\<turnstile>{=:n} \<langle>e1.[e2]\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
41.735 - apply (rule MGFn_NormalI)
41.736 - apply (rule ax_derivs.AVar)
41.737 - apply (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
41.738 - apply (rule allI)
41.739 - apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
41.740 - apply (fastsimp intro: eval.AVar)
41.741 - done
41.742 + apply (rule MGFn_NormalI)
41.743 + apply (rule ax_derivs.AVar)
41.744 + apply (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
41.745 + apply (rule allI)
41.746 + apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
41.747 + apply (fastsimp intro: eval.AVar)
41.748 + done
41.749 next
41.750 case (InsInitV c v)
41.751 show ?case
41.752 - by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
41.753 + by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
41.754 next
41.755 case (NewC C)
41.756 show ?case
41.757 - apply (rule MGFn_NormalI)
41.758 - apply (rule ax_derivs.NewC)
41.759 - apply (rule MGFn_InitD [OF hyp, THEN conseq2])
41.760 - apply (fastsimp intro: eval.NewC)
41.761 - done
41.762 + apply (rule MGFn_NormalI)
41.763 + apply (rule ax_derivs.NewC)
41.764 + apply (rule MGFn_InitD [OF hyp, THEN conseq2])
41.765 + apply (fastsimp intro: eval.NewC)
41.766 + done
41.767 next
41.768 case (NewA T e)
41.769 thus ?case
41.770 - apply -
41.771 - apply (rule MGFn_NormalI)
41.772 - apply (rule ax_derivs.NewA
41.773 + apply -
41.774 + apply (rule MGFn_NormalI)
41.775 + apply (rule ax_derivs.NewA
41.776 [where ?Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty T)
41.777 \<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n"])
41.778 - apply (simp add: init_comp_ty_def split add: split_if)
41.779 - apply (rule conjI, clarsimp)
41.780 - apply (rule MGFn_InitD [OF hyp, THEN conseq2])
41.781 - apply (clarsimp intro: eval.Init)
41.782 - apply clarsimp
41.783 - apply (rule ax_derivs.Skip [THEN conseq1])
41.784 - apply (clarsimp intro: eval.Skip)
41.785 - apply (erule MGFnD' [THEN conseq12])
41.786 - apply (fastsimp intro: eval.NewA)
41.787 - done
41.788 + apply (simp add: init_comp_ty_def split add: split_if)
41.789 + apply (rule conjI, clarsimp)
41.790 + apply (rule MGFn_InitD [OF hyp, THEN conseq2])
41.791 + apply (clarsimp intro: eval.Init)
41.792 + apply clarsimp
41.793 + apply (rule ax_derivs.Skip [THEN conseq1])
41.794 + apply (clarsimp intro: eval.Skip)
41.795 + apply (erule MGFnD' [THEN conseq12])
41.796 + apply (fastsimp intro: eval.NewA)
41.797 + done
41.798 next
41.799 case (Cast C e)
41.800 thus ?case
41.801 - apply -
41.802 - apply (rule MGFn_NormalI)
41.803 - apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
41.804 - apply (fastsimp intro: eval.Cast)
41.805 - done
41.806 + apply -
41.807 + apply (rule MGFn_NormalI)
41.808 + apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
41.809 + apply (fastsimp intro: eval.Cast)
41.810 + done
41.811 next
41.812 case (Inst e C)
41.813 thus ?case
41.814 - apply -
41.815 - apply (rule MGFn_NormalI)
41.816 - apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
41.817 - apply (fastsimp intro: eval.Inst)
41.818 - done
41.819 + apply -
41.820 + apply (rule MGFn_NormalI)
41.821 + apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
41.822 + apply (fastsimp intro: eval.Inst)
41.823 + done
41.824 next
41.825 case (Lit v)
41.826 show ?case
41.827 - apply -
41.828 - apply (rule MGFn_NormalI)
41.829 - apply (rule ax_derivs.Lit [THEN conseq1])
41.830 - apply (fastsimp intro: eval.Lit)
41.831 - done
41.832 + apply -
41.833 + apply (rule MGFn_NormalI)
41.834 + apply (rule ax_derivs.Lit [THEN conseq1])
41.835 + apply (fastsimp intro: eval.Lit)
41.836 + done
41.837 next
41.838 case (UnOp unop e)
41.839 thus ?case
41.840 - apply -
41.841 - apply (rule MGFn_NormalI)
41.842 - apply (rule ax_derivs.UnOp)
41.843 - apply (erule MGFnD' [THEN conseq12])
41.844 - apply (fastsimp intro: eval.UnOp)
41.845 - done
41.846 + apply -
41.847 + apply (rule MGFn_NormalI)
41.848 + apply (rule ax_derivs.UnOp)
41.849 + apply (erule MGFnD' [THEN conseq12])
41.850 + apply (fastsimp intro: eval.UnOp)
41.851 + done
41.852 next
41.853 case (BinOp binop e1 e2)
41.854 thus ?case
41.855 - apply -
41.856 - apply (rule MGFn_NormalI)
41.857 - apply (rule ax_derivs.BinOp)
41.858 - apply (erule MGFnD [THEN ax_NormalD])
41.859 - apply (rule allI)
41.860 - apply (case_tac "need_second_arg binop v1")
41.861 - apply simp
41.862 - apply (erule MGFnD' [THEN conseq12])
41.863 - apply (fastsimp intro: eval.BinOp)
41.864 - apply simp
41.865 - apply (rule ax_Normal_cases)
41.866 - apply (rule ax_derivs.Skip [THEN conseq1])
41.867 - apply clarsimp
41.868 - apply (rule eval_BinOp_arg2_indepI)
41.869 - apply simp
41.870 - apply simp
41.871 - apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
41.872 - apply (fastsimp intro: eval.BinOp)
41.873 - done
41.874 + apply -
41.875 + apply (rule MGFn_NormalI)
41.876 + apply (rule ax_derivs.BinOp)
41.877 + apply (erule MGFnD [THEN ax_NormalD])
41.878 + apply (rule allI)
41.879 + apply (case_tac "need_second_arg binop v1")
41.880 + apply simp
41.881 + apply (erule MGFnD' [THEN conseq12])
41.882 + apply (fastsimp intro: eval.BinOp)
41.883 + apply simp
41.884 + apply (rule ax_Normal_cases)
41.885 + apply (rule ax_derivs.Skip [THEN conseq1])
41.886 + apply clarsimp
41.887 + apply (rule eval_BinOp_arg2_indepI)
41.888 + apply simp
41.889 + apply simp
41.890 + apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
41.891 + apply (fastsimp intro: eval.BinOp)
41.892 + done
41.893 next
41.894 case Super
41.895 show ?case
41.896 - apply -
41.897 - apply (rule MGFn_NormalI)
41.898 - apply (rule ax_derivs.Super [THEN conseq1])
41.899 - apply (fastsimp intro: eval.Super)
41.900 - done
41.901 + apply -
41.902 + apply (rule MGFn_NormalI)
41.903 + apply (rule ax_derivs.Super [THEN conseq1])
41.904 + apply (fastsimp intro: eval.Super)
41.905 + done
41.906 next
41.907 case (Acc v)
41.908 thus ?case
41.909 - apply -
41.910 - apply (rule MGFn_NormalI)
41.911 - apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
41.912 - apply (fastsimp intro: eval.Acc simp add: split_paired_all)
41.913 - done
41.914 + apply -
41.915 + apply (rule MGFn_NormalI)
41.916 + apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
41.917 + apply (fastsimp intro: eval.Acc simp add: split_paired_all)
41.918 + done
41.919 next
41.920 case (Ass v e)
41.921 thus "G,A\<turnstile>{=:n} \<langle>v:=e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
41.922 - apply -
41.923 - apply (rule MGFn_NormalI)
41.924 - apply (rule ax_derivs.Ass)
41.925 - apply (erule MGFnD [THEN ax_NormalD])
41.926 - apply (rule allI)
41.927 - apply (erule MGFnD'[THEN conseq12])
41.928 - apply (fastsimp intro: eval.Ass simp add: split_paired_all)
41.929 - done
41.930 + apply -
41.931 + apply (rule MGFn_NormalI)
41.932 + apply (rule ax_derivs.Ass)
41.933 + apply (erule MGFnD [THEN ax_NormalD])
41.934 + apply (rule allI)
41.935 + apply (erule MGFnD'[THEN conseq12])
41.936 + apply (fastsimp intro: eval.Ass simp add: split_paired_all)
41.937 + done
41.938 next
41.939 case (Cond e1 e2 e3)
41.940 thus "G,A\<turnstile>{=:n} \<langle>e1 ? e2 : e3\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
41.941 - apply -
41.942 - apply (rule MGFn_NormalI)
41.943 - apply (rule ax_derivs.Cond)
41.944 - apply (erule MGFnD [THEN ax_NormalD])
41.945 - apply (rule allI)
41.946 - apply (rule ax_Normal_cases)
41.947 - prefer 2
41.948 - apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
41.949 - apply (fastsimp intro: eval.Cond)
41.950 - apply (case_tac "b")
41.951 - apply simp
41.952 - apply (erule MGFnD'[THEN conseq12])
41.953 - apply (fastsimp intro: eval.Cond)
41.954 - apply simp
41.955 - apply (erule MGFnD'[THEN conseq12])
41.956 - apply (fastsimp intro: eval.Cond)
41.957 - done
41.958 + apply -
41.959 + apply (rule MGFn_NormalI)
41.960 + apply (rule ax_derivs.Cond)
41.961 + apply (erule MGFnD [THEN ax_NormalD])
41.962 + apply (rule allI)
41.963 + apply (rule ax_Normal_cases)
41.964 + prefer 2
41.965 + apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
41.966 + apply (fastsimp intro: eval.Cond)
41.967 + apply (case_tac "b")
41.968 + apply simp
41.969 + apply (erule MGFnD'[THEN conseq12])
41.970 + apply (fastsimp intro: eval.Cond)
41.971 + apply simp
41.972 + apply (erule MGFnD'[THEN conseq12])
41.973 + apply (fastsimp intro: eval.Cond)
41.974 + done
41.975 next
41.976 case (Call accC statT mode e mn pTs' ps)
41.977 note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
41.978 note mgf_ps = `G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}`
41.979 from mgf_methds mgf_e mgf_ps wf
41.980 show "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
41.981 - by (rule MGFn_Call)
41.982 + by (rule MGFn_Call)
41.983 next
41.984 case (Methd D mn)
41.985 from mgf_methds
41.986 show "G,A\<turnstile>{=:n} \<langle>Methd D mn\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
41.987 - by simp
41.988 + by simp
41.989 next
41.990 case (Body D c)
41.991 note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
41.992 from wf MGFn_Init [OF hyp] mgf_c
41.993 show "G,A\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
41.994 - by (rule MGFn_Body)
41.995 + by (rule MGFn_Body)
41.996 next
41.997 case (InsInitE c e)
41.998 show ?case
41.999 - by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
41.1000 + by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
41.1001 next
41.1002 case (Callee l e)
41.1003 show ?case
41.1004 - by (rule MGFn_NormalI) (rule ax_derivs.Callee)
41.1005 + by (rule MGFn_NormalI) (rule ax_derivs.Callee)
41.1006 next
41.1007 case Skip
41.1008 show ?case
41.1009 - apply -
41.1010 - apply (rule MGFn_NormalI)
41.1011 - apply (rule ax_derivs.Skip [THEN conseq1])
41.1012 - apply (fastsimp intro: eval.Skip)
41.1013 - done
41.1014 + apply -
41.1015 + apply (rule MGFn_NormalI)
41.1016 + apply (rule ax_derivs.Skip [THEN conseq1])
41.1017 + apply (fastsimp intro: eval.Skip)
41.1018 + done
41.1019 next
41.1020 case (Expr e)
41.1021 thus ?case
41.1022 - apply -
41.1023 - apply (rule MGFn_NormalI)
41.1024 - apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
41.1025 - apply (fastsimp intro: eval.Expr)
41.1026 - done
41.1027 + apply -
41.1028 + apply (rule MGFn_NormalI)
41.1029 + apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
41.1030 + apply (fastsimp intro: eval.Expr)
41.1031 + done
41.1032 next
41.1033 case (Lab l c)
41.1034 thus "G,A\<turnstile>{=:n} \<langle>l\<bullet> c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1035 - apply -
41.1036 - apply (rule MGFn_NormalI)
41.1037 - apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
41.1038 - apply (fastsimp intro: eval.Lab)
41.1039 - done
41.1040 + apply -
41.1041 + apply (rule MGFn_NormalI)
41.1042 + apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
41.1043 + apply (fastsimp intro: eval.Lab)
41.1044 + done
41.1045 next
41.1046 case (Comp c1 c2)
41.1047 thus "G,A\<turnstile>{=:n} \<langle>c1;; c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1048 - apply -
41.1049 - apply (rule MGFn_NormalI)
41.1050 - apply (rule ax_derivs.Comp)
41.1051 - apply (erule MGFnD [THEN ax_NormalD])
41.1052 - apply (erule MGFnD' [THEN conseq12])
41.1053 - apply (fastsimp intro: eval.Comp)
41.1054 - done
41.1055 + apply -
41.1056 + apply (rule MGFn_NormalI)
41.1057 + apply (rule ax_derivs.Comp)
41.1058 + apply (erule MGFnD [THEN ax_NormalD])
41.1059 + apply (erule MGFnD' [THEN conseq12])
41.1060 + apply (fastsimp intro: eval.Comp)
41.1061 + done
41.1062 next
41.1063 case (If' e c1 c2)
41.1064 thus "G,A\<turnstile>{=:n} \<langle>If(e) c1 Else c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1065 - apply -
41.1066 - apply (rule MGFn_NormalI)
41.1067 - apply (rule ax_derivs.If)
41.1068 - apply (erule MGFnD [THEN ax_NormalD])
41.1069 - apply (rule allI)
41.1070 - apply (rule ax_Normal_cases)
41.1071 - prefer 2
41.1072 - apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
41.1073 - apply (fastsimp intro: eval.If)
41.1074 - apply (case_tac "b")
41.1075 - apply simp
41.1076 - apply (erule MGFnD' [THEN conseq12])
41.1077 - apply (fastsimp intro: eval.If)
41.1078 - apply simp
41.1079 - apply (erule MGFnD' [THEN conseq12])
41.1080 - apply (fastsimp intro: eval.If)
41.1081 - done
41.1082 + apply -
41.1083 + apply (rule MGFn_NormalI)
41.1084 + apply (rule ax_derivs.If)
41.1085 + apply (erule MGFnD [THEN ax_NormalD])
41.1086 + apply (rule allI)
41.1087 + apply (rule ax_Normal_cases)
41.1088 + prefer 2
41.1089 + apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
41.1090 + apply (fastsimp intro: eval.If)
41.1091 + apply (case_tac "b")
41.1092 + apply simp
41.1093 + apply (erule MGFnD' [THEN conseq12])
41.1094 + apply (fastsimp intro: eval.If)
41.1095 + apply simp
41.1096 + apply (erule MGFnD' [THEN conseq12])
41.1097 + apply (fastsimp intro: eval.If)
41.1098 + done
41.1099 next
41.1100 case (Loop l e c)
41.1101 note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
41.1102 note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
41.1103 from mgf_e mgf_c wf
41.1104 show "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1105 - by (rule MGFn_Loop)
41.1106 + by (rule MGFn_Loop)
41.1107 next
41.1108 case (Jmp j)
41.1109 thus ?case
41.1110 - apply -
41.1111 - apply (rule MGFn_NormalI)
41.1112 - apply (rule ax_derivs.Jmp [THEN conseq1])
41.1113 - apply (auto intro: eval.Jmp simp add: abupd_def2)
41.1114 - done
41.1115 + apply -
41.1116 + apply (rule MGFn_NormalI)
41.1117 + apply (rule ax_derivs.Jmp [THEN conseq1])
41.1118 + apply (auto intro: eval.Jmp simp add: abupd_def2)
41.1119 + done
41.1120 next
41.1121 case (Throw e)
41.1122 thus ?case
41.1123 - apply -
41.1124 - apply (rule MGFn_NormalI)
41.1125 - apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
41.1126 - apply (fastsimp intro: eval.Throw)
41.1127 - done
41.1128 + apply -
41.1129 + apply (rule MGFn_NormalI)
41.1130 + apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
41.1131 + apply (fastsimp intro: eval.Throw)
41.1132 + done
41.1133 next
41.1134 case (TryC c1 C vn c2)
41.1135 thus "G,A\<turnstile>{=:n} \<langle>Try c1 Catch(C vn) c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1136 - apply -
41.1137 - apply (rule MGFn_NormalI)
41.1138 - apply (rule ax_derivs.Try [where
41.1139 + apply -
41.1140 + apply (rule MGFn_NormalI)
41.1141 + apply (rule ax_derivs.Try [where
41.1142 ?Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>\<langle>c1\<rangle>\<^sub>s\<succ>\<rightarrow> (Y',s'') \<and>
41.1143 G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n"])
41.1144 - apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
41.1145 - apply (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
41.1146 - apply (erule MGFnD'[THEN conseq12])
41.1147 - apply (fastsimp intro: eval.Try)
41.1148 - apply (fastsimp intro: eval.Try)
41.1149 - done
41.1150 + apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
41.1151 + apply (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
41.1152 + apply (erule MGFnD'[THEN conseq12])
41.1153 + apply (fastsimp intro: eval.Try)
41.1154 + apply (fastsimp intro: eval.Try)
41.1155 + done
41.1156 next
41.1157 case (Fin c1 c2)
41.1158 note mgf_c1 = `G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
41.1159 note mgf_c2 = `G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
41.1160 from wf mgf_c1 mgf_c2
41.1161 show "G,A\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1162 - by (rule MGFn_Fin)
41.1163 + by (rule MGFn_Fin)
41.1164 next
41.1165 case (FinA abr c)
41.1166 show ?case
41.1167 - by (rule MGFn_NormalI) (rule ax_derivs.FinA)
41.1168 + by (rule MGFn_NormalI) (rule ax_derivs.FinA)
41.1169 next
41.1170 case (Init C)
41.1171 from hyp
41.1172 show "G,A\<turnstile>{=:n} \<langle>Init C\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
41.1173 - by (rule MGFn_Init)
41.1174 + by (rule MGFn_Init)
41.1175 next
41.1176 case Nil_expr
41.1177 show "G,A\<turnstile>{=:n} \<langle>[]\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
41.1178 - apply -
41.1179 - apply (rule MGFn_NormalI)
41.1180 - apply (rule ax_derivs.Nil [THEN conseq1])
41.1181 - apply (fastsimp intro: eval.Nil)
41.1182 - done
41.1183 + apply -
41.1184 + apply (rule MGFn_NormalI)
41.1185 + apply (rule ax_derivs.Nil [THEN conseq1])
41.1186 + apply (fastsimp intro: eval.Nil)
41.1187 + done
41.1188 next
41.1189 case (Cons_expr e es)
41.1190 thus "G,A\<turnstile>{=:n} \<langle>e# es\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
41.1191 - apply -
41.1192 - apply (rule MGFn_NormalI)
41.1193 - apply (rule ax_derivs.Cons)
41.1194 - apply (erule MGFnD [THEN ax_NormalD])
41.1195 - apply (rule allI)
41.1196 - apply (erule MGFnD'[THEN conseq12])
41.1197 - apply (fastsimp intro: eval.Cons)
41.1198 - done
41.1199 + apply -
41.1200 + apply (rule MGFn_NormalI)
41.1201 + apply (rule ax_derivs.Cons)
41.1202 + apply (erule MGFnD [THEN ax_NormalD])
41.1203 + apply (rule allI)
41.1204 + apply (erule MGFnD'[THEN conseq12])
41.1205 + apply (fastsimp intro: eval.Cons)
41.1206 + done
41.1207 qed
41.1208 }
41.1209 thus ?thesis
41.1210 @@ -1519,15 +1519,15 @@
41.1211 proof -
41.1212 from eval_t type_ok wf
41.1213 obtain n where evaln: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s')"
41.1214 - by (rule eval_to_evaln [elim_format]) iprover
41.1215 + by (rule eval_to_evaln [elim_format]) iprover
41.1216 from valid have
41.1217 - valid_expanded:
41.1218 - "\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s
41.1219 + valid_expanded:
41.1220 + "\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s
41.1221 \<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s') \<longrightarrow> Q Y' s' Z)"
41.1222 - by (simp add: ax_valids_def triple_valid_def)
41.1223 + by (simp add: ax_valids_def triple_valid_def)
41.1224 from P type_ok evaln
41.1225 show "Q Y' s' Z"
41.1226 - by (rule valid_expanded [rule_format])
41.1227 + by (rule valid_expanded [rule_format])
41.1228 qed
41.1229 qed
41.1230 qed
42.1 --- a/src/HOL/Bali/AxSem.thy Sat Oct 17 01:05:59 2009 +0200
42.2 +++ b/src/HOL/Bali/AxSem.thy Sat Oct 17 14:43:18 2009 +0200
42.3 @@ -519,7 +519,7 @@
42.4 G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
42.5
42.6 | NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q}; G,A\<turnstile>{Q} e-\<succ>
42.7 - {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
42.8 + {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
42.9 G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}"
42.10
42.11 | Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
43.1 --- a/src/HOL/Bali/AxSound.thy Sat Oct 17 01:05:59 2009 +0200
43.2 +++ b/src/HOL/Bali/AxSound.thy Sat Oct 17 14:43:18 2009 +0200
43.3 @@ -118,14 +118,14 @@
43.4 case 0
43.5 show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>0\<Colon>t"
43.6 proof -
43.7 - {
43.8 - fix C sig
43.9 - assume "(C,sig) \<in> ms"
43.10 - have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
43.11 - by (rule Methd_triple_valid2_0)
43.12 - }
43.13 - thus ?thesis
43.14 - by (simp add: mtriples_def split_def)
43.15 + {
43.16 + fix C sig
43.17 + assume "(C,sig) \<in> ms"
43.18 + have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
43.19 + by (rule Methd_triple_valid2_0)
43.20 + }
43.21 + thus ?thesis
43.22 + by (simp add: mtriples_def split_def)
43.23 qed
43.24 next
43.25 case (Suc m)
43.26 @@ -133,28 +133,28 @@
43.27 note prem = `\<forall>t\<in>A. G\<Turnstile>Suc m\<Colon>t`
43.28 show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>Suc m\<Colon>t"
43.29 proof -
43.30 - {
43.31 - fix C sig
43.32 - assume m: "(C,sig) \<in> ms"
43.33 - have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
43.34 - proof -
43.35 - from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
43.36 - by (rule triples_valid2_Suc)
43.37 - hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
43.38 - by (rule hyp)
43.39 - with prem_m
43.40 - have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
43.41 - by (simp add: ball_Un)
43.42 - hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
43.43 - by (rule recursive)
43.44 - with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
43.45 - by (auto simp add: mtriples_def split_def)
43.46 - thus ?thesis
43.47 - by (rule Methd_triple_valid2_SucI)
43.48 - qed
43.49 - }
43.50 - thus ?thesis
43.51 - by (simp add: mtriples_def split_def)
43.52 + {
43.53 + fix C sig
43.54 + assume m: "(C,sig) \<in> ms"
43.55 + have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
43.56 + proof -
43.57 + from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
43.58 + by (rule triples_valid2_Suc)
43.59 + hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
43.60 + by (rule hyp)
43.61 + with prem_m
43.62 + have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
43.63 + by (simp add: ball_Un)
43.64 + hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
43.65 + by (rule recursive)
43.66 + with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
43.67 + by (auto simp add: mtriples_def split_def)
43.68 + thus ?thesis
43.69 + by (rule Methd_triple_valid2_SucI)
43.70 + qed
43.71 + }
43.72 + thus ?thesis
43.73 + by (simp add: mtriples_def split_def)
43.74 qed
43.75 qed
43.76 }
43.77 @@ -361,10 +361,10 @@
43.78 have "G\<Turnstile>n\<Colon>t" and "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
43.79 proof -
43.80 from valid_A valid_t show "G\<Turnstile>n\<Colon>t"
43.81 - by (simp add: ax_valids2_def)
43.82 + by (simp add: ax_valids2_def)
43.83 next
43.84 from valid_A valid_ts show "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
43.85 - by (unfold ax_valids2_def) blast
43.86 + by (unfold ax_valids2_def) blast
43.87 qed
43.88 hence "\<forall>t'\<in>insert t ts. G\<Turnstile>n\<Colon>t'"
43.89 by simp
43.90 @@ -402,27 +402,27 @@
43.91 proof -
43.92 from valid_A conf wt da eval P con
43.93 have "Q v s1 Z"
43.94 - apply (simp add: ax_valids2_def triple_valid2_def2)
43.95 - apply (tactic "smp_tac 3 1")
43.96 - apply clarify
43.97 - apply (tactic "smp_tac 1 1")
43.98 - apply (erule allE,erule allE, erule mp)
43.99 - apply (intro strip)
43.100 - apply (tactic "smp_tac 3 1")
43.101 - apply (tactic "smp_tac 2 1")
43.102 - apply (tactic "smp_tac 1 1")
43.103 - by blast
43.104 + apply (simp add: ax_valids2_def triple_valid2_def2)
43.105 + apply (tactic "smp_tac 3 1")
43.106 + apply clarify
43.107 + apply (tactic "smp_tac 1 1")
43.108 + apply (erule allE,erule allE, erule mp)
43.109 + apply (intro strip)
43.110 + apply (tactic "smp_tac 3 1")
43.111 + apply (tactic "smp_tac 2 1")
43.112 + apply (tactic "smp_tac 1 1")
43.113 + by blast
43.114 moreover have "s1\<Colon>\<preceq>(G, L)"
43.115 proof (cases "normal s0")
43.116 - case True
43.117 - from eval wt [OF True] da [OF True] conf wf
43.118 - show ?thesis
43.119 - by (rule evaln_type_sound [elim_format]) simp
43.120 + case True
43.121 + from eval wt [OF True] da [OF True] conf wf
43.122 + show ?thesis
43.123 + by (rule evaln_type_sound [elim_format]) simp
43.124 next
43.125 - case False
43.126 - with eval have "s1=s0"
43.127 - by auto
43.128 - with conf show ?thesis by simp
43.129 + case False
43.130 + with eval have "s1=s0"
43.131 + by auto
43.132 + with conf show ?thesis by simp
43.133 qed
43.134 ultimately show ?thesis ..
43.135 qed
43.136 @@ -461,13 +461,13 @@
43.137 show "P (In2 vf) s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.138 proof
43.139 from eval normal_s0 obtain "s1=s0" "vf=lvar vn (store s0)"
43.140 - by (fastsimp elim: evaln_elim_cases)
43.141 + by (fastsimp elim: evaln_elim_cases)
43.142 with P show "P (In2 vf) s1 Z"
43.143 - by simp
43.144 + by simp
43.145 next
43.146 from eval wt da conf_s0 wf
43.147 show "s1\<Colon>\<preceq>(G, L)"
43.148 - by (rule evaln_type_sound [elim_format]) simp
43.149 + by (rule evaln_type_sound [elim_format]) simp
43.150 qed
43.151 qed
43.152 next
43.153 @@ -490,78 +490,78 @@
43.154 from wt obtain statC f where
43.155 wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
43.156 accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
43.157 - eq_accC: "accC=accC'" and
43.158 + eq_accC: "accC=accC'" and
43.159 stat: "stat=is_static f" and
43.160 - T: "T=(type f)"
43.161 - by (cases) (auto simp add: member_is_static_simp)
43.162 + T: "T=(type f)"
43.163 + by (cases) (auto simp add: member_is_static_simp)
43.164 from da eq_accC
43.165 have da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V"
43.166 - by cases simp
43.167 + by cases simp
43.168 from eval obtain a s1 s2 s2' where
43.169 - eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and
43.170 + eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and
43.171 eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s2" and
43.172 - fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
43.173 - s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
43.174 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.175 + fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
43.176 + s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
43.177 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.178 have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
43.179 proof -
43.180 - from wf wt_e
43.181 - have iscls_statC: "is_class G statC"
43.182 - by (auto dest: ty_expr_is_type type_is_class)
43.183 - with wf accfield
43.184 - have iscls_statDeclC: "is_class G statDeclC"
43.185 - by (auto dest!: accfield_fields dest: fields_declC)
43.186 - thus ?thesis by simp
43.187 + from wf wt_e
43.188 + have iscls_statC: "is_class G statC"
43.189 + by (auto dest: ty_expr_is_type type_is_class)
43.190 + with wf accfield
43.191 + have iscls_statDeclC: "is_class G statDeclC"
43.192 + by (auto dest!: accfield_fields dest: fields_declC)
43.193 + thus ?thesis by simp
43.194 qed
43.195 obtain I where
43.196 - da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.197 + da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.198 \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
43.199 - by (auto intro: da_Init [simplified] assigned.select_convs)
43.200 + by (auto intro: da_Init [simplified] assigned.select_convs)
43.201 from valid_init P valid_A conf_s0 eval_init wt_init da_init
43.202 obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
43.203 - by (rule validE)
43.204 + by (rule validE)
43.205 obtain
43.206 - R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and
43.207 + R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and
43.208 conf_s2: "s2\<Colon>\<preceq>(G, L)" and
43.209 - conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
43.210 + conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
43.211 proof (cases "normal s1")
43.212 - case True
43.213 - obtain V' where
43.214 - da_e':
43.215 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
43.216 - proof -
43.217 - from eval_init
43.218 - have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
43.219 - by (rule dom_locals_evaln_mono_elim)
43.220 - with da_e show thesis
43.221 - by (rule da_weakenE) (rule that)
43.222 - qed
43.223 - with valid_e Q valid_A conf_s1 eval_e wt_e
43.224 - obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
43.225 - by (rule validE) (simp add: fvar [symmetric])
43.226 - moreover
43.227 - from eval_e wt_e da_e' conf_s1 wf
43.228 - have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
43.229 - by (rule evaln_type_sound [elim_format]) simp
43.230 - ultimately show ?thesis ..
43.231 + case True
43.232 + obtain V' where
43.233 + da_e':
43.234 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
43.235 + proof -
43.236 + from eval_init
43.237 + have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
43.238 + by (rule dom_locals_evaln_mono_elim)
43.239 + with da_e show thesis
43.240 + by (rule da_weakenE) (rule that)
43.241 + qed
43.242 + with valid_e Q valid_A conf_s1 eval_e wt_e
43.243 + obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
43.244 + by (rule validE) (simp add: fvar [symmetric])
43.245 + moreover
43.246 + from eval_e wt_e da_e' conf_s1 wf
43.247 + have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
43.248 + by (rule evaln_type_sound [elim_format]) simp
43.249 + ultimately show ?thesis ..
43.250 next
43.251 - case False
43.252 - with valid_e Q valid_A conf_s1 eval_e
43.253 - obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
43.254 - by (cases rule: validE) (simp add: fvar [symmetric])+
43.255 - moreover from False eval_e have "\<not> normal s2"
43.256 - by auto
43.257 - hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
43.258 - by auto
43.259 - ultimately show ?thesis ..
43.260 + case False
43.261 + with valid_e Q valid_A conf_s1 eval_e
43.262 + obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
43.263 + by (cases rule: validE) (simp add: fvar [symmetric])+
43.264 + moreover from False eval_e have "\<not> normal s2"
43.265 + by auto
43.266 + hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
43.267 + by auto
43.268 + ultimately show ?thesis ..
43.269 qed
43.270 from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat s3 wf
43.271 have eq_s3_s2': "s3=s2'"
43.272 - using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
43.273 + using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
43.274 moreover
43.275 from eval wt da conf_s0 wf
43.276 have "s3\<Colon>\<preceq>(G, L)"
43.277 - by (rule evaln_type_sound [elim_format]) simp
43.278 + by (rule evaln_type_sound [elim_format]) simp
43.279 ultimately show ?thesis using Q R by simp
43.280 qed
43.281 qed
43.282 @@ -584,48 +584,48 @@
43.283 show "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z \<and> s2'\<Colon>\<preceq>(G, L)"
43.284 proof -
43.285 from wt obtain
43.286 - wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
43.287 + wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
43.288 wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-PrimT Integer"
43.289 - by (rule wt_elim_cases) simp
43.290 + by (rule wt_elim_cases) simp
43.291 from da obtain E1 where
43.292 - da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
43.293 - da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
43.294 - by (rule da_elim_cases) simp
43.295 + da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
43.296 + da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
43.297 + by (rule da_elim_cases) simp
43.298 from eval obtain s1 a i s2 where
43.299 - eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
43.300 - eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
43.301 - avar: "avar G i a s2 =(vf, s2')"
43.302 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.303 + eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
43.304 + eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
43.305 + avar: "avar G i a s2 =(vf, s2')"
43.306 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.307 from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
43.308 obtain Q: "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
43.309 - by (rule validE)
43.310 + by (rule validE)
43.311 from Q have Q': "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
43.312 - by simp
43.313 + by simp
43.314 have "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z"
43.315 proof (cases "normal s1")
43.316 - case True
43.317 - obtain V' where
43.318 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
43.319 - proof -
43.320 - from eval_e1 wt_e1 da_e1 wf True
43.321 - have "nrm E1 \<subseteq> dom (locals (store s1))"
43.322 - by (cases rule: da_good_approx_evalnE) iprover
43.323 - with da_e2 show thesis
43.324 - by (rule da_weakenE) (rule that)
43.325 - qed
43.326 - with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2
43.327 - show ?thesis
43.328 - by (rule validE) (simp add: avar)
43.329 + case True
43.330 + obtain V' where
43.331 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
43.332 + proof -
43.333 + from eval_e1 wt_e1 da_e1 wf True
43.334 + have "nrm E1 \<subseteq> dom (locals (store s1))"
43.335 + by (cases rule: da_good_approx_evalnE) iprover
43.336 + with da_e2 show thesis
43.337 + by (rule da_weakenE) (rule that)
43.338 + qed
43.339 + with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2
43.340 + show ?thesis
43.341 + by (rule validE) (simp add: avar)
43.342 next
43.343 - case False
43.344 - with valid_e2 Q' valid_A conf_s1 eval_e2
43.345 - show ?thesis
43.346 - by (cases rule: validE) (simp add: avar)+
43.347 + case False
43.348 + with valid_e2 Q' valid_A conf_s1 eval_e2
43.349 + show ?thesis
43.350 + by (cases rule: validE) (simp add: avar)+
43.351 qed
43.352 moreover
43.353 from eval wt da conf_s0 wf
43.354 have "s2'\<Colon>\<preceq>(G, L)"
43.355 - by (rule evaln_type_sound [elim_format]) simp
43.356 + by (rule evaln_type_sound [elim_format]) simp
43.357 ultimately show ?thesis ..
43.358 qed
43.359 qed
43.360 @@ -646,26 +646,26 @@
43.361 show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
43.362 proof -
43.363 from wt obtain is_cls_C: "is_class G C"
43.364 - by (rule wt_elim_cases) (auto dest: is_acc_classD)
43.365 + by (rule wt_elim_cases) (auto dest: is_acc_classD)
43.366 hence wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>"
43.367 - by auto
43.368 + by auto
43.369 obtain I where
43.370 - da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
43.371 - by (auto intro: da_Init [simplified] assigned.select_convs)
43.372 + da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
43.373 + by (auto intro: da_Init [simplified] assigned.select_convs)
43.374 from eval obtain s1 a where
43.375 - eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and
43.376 + eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and
43.377 alloc: "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" and
43.378 - v: "v=Addr a"
43.379 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.380 + v: "v=Addr a"
43.381 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.382 from valid_init P valid_A conf_s0 eval_init wt_init da_init
43.383 obtain "(Alloc G (CInst C) Q) \<diamondsuit> s1 Z"
43.384 - by (rule validE)
43.385 + by (rule validE)
43.386 with alloc v have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.387 - by simp
43.388 + by simp
43.389 moreover
43.390 from eval wt da conf_s0 wf
43.391 have "s2\<Colon>\<preceq>(G, L)"
43.392 - by (rule evaln_type_sound [elim_format]) simp
43.393 + by (rule evaln_type_sound [elim_format]) simp
43.394 ultimately show ?thesis ..
43.395 qed
43.396 qed
43.397 @@ -687,58 +687,58 @@
43.398 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
43.399 proof -
43.400 from wt obtain
43.401 - wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
43.402 - wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer"
43.403 - by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
43.404 + wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
43.405 + wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer"
43.406 + by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
43.407 from da obtain
43.408 - da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.409 - by cases simp
43.410 + da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.411 + by cases simp
43.412 from eval obtain s1 i s2 a where
43.413 - eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and
43.414 + eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and
43.415 eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>i\<midarrow>n\<rightarrow> s2" and
43.416 alloc: "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" and
43.417 v: "v=Addr a"
43.418 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.419 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.420 obtain I where
43.421 - da_init:
43.422 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
43.423 + da_init:
43.424 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
43.425 proof (cases "\<exists>C. T = Class C")
43.426 - case True
43.427 - thus ?thesis
43.428 - by - (rule that, (auto intro: da_Init [simplified]
43.429 + case True
43.430 + thus ?thesis
43.431 + by - (rule that, (auto intro: da_Init [simplified]
43.432 assigned.select_convs
43.433 simp add: init_comp_ty_def))
43.434 - (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
43.435 + (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
43.436 next
43.437 - case False
43.438 - thus ?thesis
43.439 - by - (rule that, (auto intro: da_Skip [simplified]
43.440 + case False
43.441 + thus ?thesis
43.442 + by - (rule that, (auto intro: da_Skip [simplified]
43.443 assigned.select_convs
43.444 simp add: init_comp_ty_def))
43.445 (* simplified: to rewrite \<langle>Skip\<rangle> to In1r (Skip) *)
43.446 qed
43.447 with valid_init P valid_A conf_s0 eval_init wt_init
43.448 obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
43.449 - by (rule validE)
43.450 + by (rule validE)
43.451 obtain E' where
43.452 "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
43.453 proof -
43.454 - from eval_init
43.455 - have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
43.456 - by (rule dom_locals_evaln_mono_elim)
43.457 - with da_e show thesis
43.458 - by (rule da_weakenE) (rule that)
43.459 + from eval_init
43.460 + have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
43.461 + by (rule dom_locals_evaln_mono_elim)
43.462 + with da_e show thesis
43.463 + by (rule da_weakenE) (rule that)
43.464 qed
43.465 with valid_e Q valid_A conf_s1 eval_e wt_e
43.466 have "(\<lambda>Val:i:. abupd (check_neg i) .;
43.467 Alloc G (Arr T (the_Intg i)) R) \<lfloor>i\<rfloor>\<^sub>e s2 Z"
43.468 - by (rule validE)
43.469 + by (rule validE)
43.470 with alloc v have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
43.471 - by simp
43.472 + by simp
43.473 moreover
43.474 from eval wt da conf_s0 wf
43.475 have "s3\<Colon>\<preceq>(G, L)"
43.476 - by (rule evaln_type_sound [elim_format]) simp
43.477 + by (rule evaln_type_sound [elim_format]) simp
43.478 ultimately show ?thesis ..
43.479 qed
43.480 qed
43.481 @@ -760,25 +760,25 @@
43.482 show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
43.483 proof -
43.484 from wt obtain eT where
43.485 - wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.486 - by cases simp
43.487 + wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.488 + by cases simp
43.489 from da obtain
43.490 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.491 - by cases simp
43.492 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.493 + by cases simp
43.494 from eval obtain s1 where
43.495 - eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
43.496 + eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
43.497 s2: "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits T) ClassCast) s1"
43.498 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.499 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.500 from valid_e P valid_A conf_s0 eval_e wt_e da_e
43.501 have "(\<lambda>Val:v:. \<lambda>s.. abupd (raise_if (\<not> G,s\<turnstile>v fits T) ClassCast) .;
43.502 Q\<leftarrow>In1 v) \<lfloor>v\<rfloor>\<^sub>e s1 Z"
43.503 - by (rule validE)
43.504 + by (rule validE)
43.505 with s2 have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.506 - by simp
43.507 + by simp
43.508 moreover
43.509 from eval wt da conf_s0 wf
43.510 have "s2\<Colon>\<preceq>(G, L)"
43.511 - by (rule evaln_type_sound [elim_format]) simp
43.512 + by (rule evaln_type_sound [elim_format]) simp
43.513 ultimately show ?thesis ..
43.514 qed
43.515 qed
43.516 @@ -799,25 +799,25 @@
43.517 show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.518 proof -
43.519 from wt obtain eT where
43.520 - wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.521 - by cases simp
43.522 + wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.523 + by cases simp
43.524 from da obtain
43.525 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.526 - by cases simp
43.527 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.528 + by cases simp
43.529 from eval obtain a where
43.530 - eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
43.531 + eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
43.532 v: "v = Bool (a \<noteq> Null \<and> G,store s1\<turnstile>a fits RefT T)"
43.533 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.534 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.535 from valid_e P valid_A conf_s0 eval_e wt_e da_e
43.536 have "(\<lambda>Val:v:. \<lambda>s.. Q\<leftarrow>In1 (Bool (v \<noteq> Null \<and> G,s\<turnstile>v fits RefT T)))
43.537 \<lfloor>a\<rfloor>\<^sub>e s1 Z"
43.538 - by (rule validE)
43.539 + by (rule validE)
43.540 with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
43.541 - by simp
43.542 + by simp
43.543 moreover
43.544 from eval wt da conf_s0 wf
43.545 have "s1\<Colon>\<preceq>(G, L)"
43.546 - by (rule evaln_type_sound [elim_format]) simp
43.547 + by (rule evaln_type_sound [elim_format]) simp
43.548 ultimately show ?thesis ..
43.549 qed
43.550 qed
43.551 @@ -833,7 +833,7 @@
43.552 show "P \<lfloor>v'\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.553 proof -
43.554 from eval have "s1=s0" and "v'=v"
43.555 - using normal_s0 by (auto elim: evaln_elim_cases)
43.556 + using normal_s0 by (auto elim: evaln_elim_cases)
43.557 with P conf_s0 show ?thesis by simp
43.558 qed
43.559 qed
43.560 @@ -853,24 +853,24 @@
43.561 show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.562 proof -
43.563 from wt obtain eT where
43.564 - wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.565 - by cases simp
43.566 + wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.567 + by cases simp
43.568 from da obtain
43.569 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.570 - by cases simp
43.571 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.572 + by cases simp
43.573 from eval obtain ve where
43.574 - eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
43.575 + eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
43.576 v: "v = eval_unop unop ve"
43.577 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.578 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.579 from valid_e P valid_A conf_s0 eval_e wt_e da_e
43.580 have "(\<lambda>Val:v:. Q\<leftarrow>In1 (eval_unop unop v)) \<lfloor>ve\<rfloor>\<^sub>e s1 Z"
43.581 - by (rule validE)
43.582 + by (rule validE)
43.583 with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
43.584 - by simp
43.585 + by simp
43.586 moreover
43.587 from eval wt da conf_s0 wf
43.588 have "s1\<Colon>\<preceq>(G, L)"
43.589 - by (rule evaln_type_sound [elim_format]) simp
43.590 + by (rule evaln_type_sound [elim_format]) simp
43.591 ultimately show ?thesis ..
43.592 qed
43.593 qed
43.594 @@ -897,63 +897,63 @@
43.595 from wt obtain e1T e2T where
43.596 wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-e1T" and
43.597 wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-e2T" and
43.598 - wt_binop: "wt_binop G binop e1T e2T"
43.599 - by cases simp
43.600 + wt_binop: "wt_binop G binop e1T e2T"
43.601 + by cases simp
43.602 have wt_Skip: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>Skip\<Colon>\<surd>"
43.603 - by simp
43.604 + by simp
43.605 (*
43.606 obtain S where
43.607 - daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.608 + daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.609 \<turnstile> dom (locals (store s1)) \<guillemotright>In1r Skip\<guillemotright> S"
43.610 - by (auto intro: da_Skip [simplified] assigned.select_convs) *)
43.611 + by (auto intro: da_Skip [simplified] assigned.select_convs) *)
43.612 from da obtain E1 where
43.613 - da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
43.614 - by cases simp+
43.615 + da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
43.616 + by cases simp+
43.617 from eval obtain v1 s1 v2 where
43.618 - eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
43.619 - eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
43.620 + eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
43.621 + eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
43.622 \<succ>\<midarrow>n\<rightarrow> (\<lfloor>v2\<rfloor>\<^sub>e, s2)" and
43.623 v: "v=eval_binop binop v1 v2"
43.624 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.625 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.626 from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
43.627 obtain Q: "Q \<lfloor>v1\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.628 - by (rule validE)
43.629 + by (rule validE)
43.630 from Q have Q': "\<And> v. (Q\<leftarrow>In1 v1) v s1 Z"
43.631 - by simp
43.632 + by simp
43.633 have "(\<lambda>Val:v2:. R\<leftarrow>In1 (eval_binop binop v1 v2)) \<lfloor>v2\<rfloor>\<^sub>e s2 Z"
43.634 proof (cases "normal s1")
43.635 - case True
43.636 - from eval_e1 wt_e1 da_e1 conf_s0 wf
43.637 - have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T"
43.638 - by (rule evaln_type_sound [elim_format]) (insert True,simp)
43.639 - from eval_e1
43.640 - have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
43.641 - by (rule evaln_eval)
43.642 - from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
43.643 - obtain E2 where
43.644 - da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1))
43.645 + case True
43.646 + from eval_e1 wt_e1 da_e1 conf_s0 wf
43.647 + have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T"
43.648 + by (rule evaln_type_sound [elim_format]) (insert True,simp)
43.649 + from eval_e1
43.650 + have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
43.651 + by (rule evaln_eval)
43.652 + from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
43.653 + obtain E2 where
43.654 + da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1))
43.655 \<guillemotright>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<guillemotright> E2"
43.656 - by (rule da_e2_BinOp [elim_format]) iprover
43.657 - from wt_e2 wt_Skip obtain T2
43.658 - where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.659 + by (rule da_e2_BinOp [elim_format]) iprover
43.660 + from wt_e2 wt_Skip obtain T2
43.661 + where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.662 \<turnstile>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<Colon>T2"
43.663 - by (cases "need_second_arg binop v1") auto
43.664 - note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2 this da_e2]
43.665 - (* chaining Q', without extra OF causes unification error *)
43.666 - thus ?thesis
43.667 - by (rule ve)
43.668 + by (cases "need_second_arg binop v1") auto
43.669 + note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2 this da_e2]
43.670 + (* chaining Q', without extra OF causes unification error *)
43.671 + thus ?thesis
43.672 + by (rule ve)
43.673 next
43.674 - case False
43.675 - note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
43.676 - with False show ?thesis
43.677 - by iprover
43.678 + case False
43.679 + note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
43.680 + with False show ?thesis
43.681 + by iprover
43.682 qed
43.683 with v have "R \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.684 - by simp
43.685 + by simp
43.686 moreover
43.687 from eval wt da conf_s0 wf
43.688 have "s2\<Colon>\<preceq>(G, L)"
43.689 - by (rule evaln_type_sound [elim_format]) simp
43.690 + by (rule evaln_type_sound [elim_format]) simp
43.691 ultimately show ?thesis ..
43.692 qed
43.693 qed
43.694 @@ -969,7 +969,7 @@
43.695 show "P \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.696 proof -
43.697 from eval have "s1=s0" and "v=val_this (store s0)"
43.698 - using normal_s0 by (auto elim: evaln_elim_cases)
43.699 + using normal_s0 by (auto elim: evaln_elim_cases)
43.700 with P conf_s0 show ?thesis by simp
43.701 qed
43.702 qed
43.703 @@ -989,23 +989,23 @@
43.704 show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.705 proof -
43.706 from wt obtain
43.707 - wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T"
43.708 - by cases simp
43.709 + wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T"
43.710 + by cases simp
43.711 from da obtain V where
43.712 - da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
43.713 - by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
43.714 + da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
43.715 + by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
43.716 from eval obtain w upd where
43.717 - eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
43.718 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.719 + eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
43.720 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.721 from valid_var P valid_A conf_s0 eval_var wt_var da_var
43.722 have "(\<lambda>Var:(v, f):. Q\<leftarrow>In1 v) \<lfloor>(v, upd)\<rfloor>\<^sub>v s1 Z"
43.723 - by (rule validE)
43.724 + by (rule validE)
43.725 then have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
43.726 - by simp
43.727 + by simp
43.728 moreover
43.729 from eval wt da conf_s0 wf
43.730 have "s1\<Colon>\<preceq>(G, L)"
43.731 - by (rule evaln_type_sound [elim_format]) simp
43.732 + by (rule evaln_type_sound [elim_format]) simp
43.733 ultimately show ?thesis ..
43.734 qed
43.735 qed
43.736 @@ -1028,96 +1028,96 @@
43.737 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
43.738 proof -
43.739 from wt obtain varT where
43.740 - wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=varT" and
43.741 - wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
43.742 - by cases simp
43.743 + wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=varT" and
43.744 + wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
43.745 + by cases simp
43.746 from eval obtain w upd s1 s2 where
43.747 - eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
43.748 + eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
43.749 eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s2" and
43.750 - s3: "s3=assign upd v s2"
43.751 - using normal_s0 by (auto elim: evaln_elim_cases)
43.752 + s3: "s3=assign upd v s2"
43.753 + using normal_s0 by (auto elim: evaln_elim_cases)
43.754 have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
43.755 proof (cases "\<exists> vn. var = LVar vn")
43.756 - case False
43.757 - with da obtain V where
43.758 - da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.759 + case False
43.760 + with da obtain V where
43.761 + da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.762 \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V" and
43.763 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.764 - by cases simp+
43.765 - from valid_var P valid_A conf_s0 eval_var wt_var da_var
43.766 - obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.767 - by (rule validE)
43.768 - hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
43.769 - by simp
43.770 - have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.771 - proof (cases "normal s1")
43.772 - case True
43.773 - obtain E' where
43.774 - da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
43.775 - proof -
43.776 - from eval_var wt_var da_var wf True
43.777 - have "nrm V \<subseteq> dom (locals (store s1))"
43.778 - by (cases rule: da_good_approx_evalnE) iprover
43.779 - with da_e show thesis
43.780 - by (rule da_weakenE) (rule that)
43.781 - qed
43.782 - note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
43.783 - show ?thesis
43.784 - by (rule ve)
43.785 - next
43.786 - case False
43.787 - note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
43.788 - with False show ?thesis
43.789 - by iprover
43.790 - qed
43.791 - with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
43.792 - by simp
43.793 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.794 + by cases simp+
43.795 + from valid_var P valid_A conf_s0 eval_var wt_var da_var
43.796 + obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.797 + by (rule validE)
43.798 + hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
43.799 + by simp
43.800 + have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.801 + proof (cases "normal s1")
43.802 + case True
43.803 + obtain E' where
43.804 + da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
43.805 + proof -
43.806 + from eval_var wt_var da_var wf True
43.807 + have "nrm V \<subseteq> dom (locals (store s1))"
43.808 + by (cases rule: da_good_approx_evalnE) iprover
43.809 + with da_e show thesis
43.810 + by (rule da_weakenE) (rule that)
43.811 + qed
43.812 + note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
43.813 + show ?thesis
43.814 + by (rule ve)
43.815 + next
43.816 + case False
43.817 + note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
43.818 + with False show ?thesis
43.819 + by iprover
43.820 + qed
43.821 + with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
43.822 + by simp
43.823 next
43.824 - case True
43.825 - then obtain vn where
43.826 - vn: "var = LVar vn"
43.827 - by auto
43.828 - with da obtain E where
43.829 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.830 - by cases simp+
43.831 - from da.LVar vn obtain V where
43.832 - da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.833 + case True
43.834 + then obtain vn where
43.835 + vn: "var = LVar vn"
43.836 + by auto
43.837 + with da obtain E where
43.838 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
43.839 + by cases simp+
43.840 + from da.LVar vn obtain V where
43.841 + da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.842 \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
43.843 - by auto
43.844 - from valid_var P valid_A conf_s0 eval_var wt_var da_var
43.845 - obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.846 - by (rule validE)
43.847 - hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
43.848 - by simp
43.849 - have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.850 - proof (cases "normal s1")
43.851 - case True
43.852 - obtain E' where
43.853 - da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.854 + by auto
43.855 + from valid_var P valid_A conf_s0 eval_var wt_var da_var
43.856 + obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.857 + by (rule validE)
43.858 + hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
43.859 + by simp
43.860 + have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.861 + proof (cases "normal s1")
43.862 + case True
43.863 + obtain E' where
43.864 + da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.865 \<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
43.866 - proof -
43.867 - from eval_var
43.868 - have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
43.869 - by (rule dom_locals_evaln_mono_elim)
43.870 - with da_e show thesis
43.871 - by (rule da_weakenE) (rule that)
43.872 - qed
43.873 - note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
43.874 - show ?thesis
43.875 - by (rule ve)
43.876 - next
43.877 - case False
43.878 - note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
43.879 - with False show ?thesis
43.880 - by iprover
43.881 - qed
43.882 - with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
43.883 - by simp
43.884 + proof -
43.885 + from eval_var
43.886 + have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
43.887 + by (rule dom_locals_evaln_mono_elim)
43.888 + with da_e show thesis
43.889 + by (rule da_weakenE) (rule that)
43.890 + qed
43.891 + note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
43.892 + show ?thesis
43.893 + by (rule ve)
43.894 + next
43.895 + case False
43.896 + note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
43.897 + with False show ?thesis
43.898 + by iprover
43.899 + qed
43.900 + with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
43.901 + by simp
43.902 qed
43.903 moreover
43.904 from eval wt da conf_s0 wf
43.905 have "s3\<Colon>\<preceq>(G, L)"
43.906 - by (rule evaln_type_sound [elim_format]) simp
43.907 + by (rule evaln_type_sound [elim_format]) simp
43.908 ultimately show ?thesis ..
43.909 qed
43.910 qed
43.911 @@ -1139,75 +1139,75 @@
43.912 show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
43.913 proof -
43.914 from wt obtain T1 T2 where
43.915 - wt_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
43.916 - wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T1" and
43.917 - wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-T2"
43.918 - by cases simp
43.919 + wt_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
43.920 + wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T1" and
43.921 + wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-T2"
43.922 + by cases simp
43.923 from da obtain E0 E1 E2 where
43.924 da_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e0\<rangle>\<^sub>e\<guillemotright> E0" and
43.925 da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.926 \<turnstile>(dom (locals (store s0)) \<union> assigns_if True e0)\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
43.927 da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.928 \<turnstile>(dom (locals (store s0)) \<union> assigns_if False e0)\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2"
43.929 - by cases simp+
43.930 + by cases simp+
43.931 from eval obtain b s1 where
43.932 - eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
43.933 + eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
43.934 eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2"
43.935 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.936 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.937 from valid_e0 P valid_A conf_s0 eval_e0 wt_e0 da_e0
43.938 obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.939 - by (rule validE)
43.940 + by (rule validE)
43.941 hence P': "\<And> v. (P'\<leftarrow>=(the_Bool b)) v s1 Z"
43.942 - by (cases "normal s1") auto
43.943 + by (cases "normal s1") auto
43.944 have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
43.945 proof (cases "normal s1")
43.946 - case True
43.947 - note normal_s1=this
43.948 - from wt_e1 wt_e2 obtain T' where
43.949 - wt_then_else:
43.950 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
43.951 - by (cases "the_Bool b") simp+
43.952 - have s0_s1: "dom (locals (store s0))
43.953 + case True
43.954 + note normal_s1=this
43.955 + from wt_e1 wt_e2 obtain T' where
43.956 + wt_then_else:
43.957 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
43.958 + by (cases "the_Bool b") simp+
43.959 + have s0_s1: "dom (locals (store s0))
43.960 \<union> assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
43.961 - proof -
43.962 - from eval_e0
43.963 - have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
43.964 - by (rule evaln_eval)
43.965 - hence
43.966 - "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
43.967 - by (rule dom_locals_eval_mono_elim)
43.968 + proof -
43.969 + from eval_e0
43.970 + have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
43.971 + by (rule evaln_eval)
43.972 + hence
43.973 + "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
43.974 + by (rule dom_locals_eval_mono_elim)
43.975 moreover
43.976 - from eval_e0' True wt_e0
43.977 - have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
43.978 - by (rule assigns_if_good_approx')
43.979 - ultimately show ?thesis by (rule Un_least)
43.980 - qed
43.981 - obtain E' where
43.982 - da_then_else:
43.983 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.984 + from eval_e0' True wt_e0
43.985 + have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
43.986 + by (rule assigns_if_good_approx')
43.987 + ultimately show ?thesis by (rule Un_least)
43.988 + qed
43.989 + obtain E' where
43.990 + da_then_else:
43.991 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.992 \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then e1 else e2\<rangle>\<^sub>e\<guillemotright> E'"
43.993 - proof (cases "the_Bool b")
43.994 - case True
43.995 - with that da_e1 s0_s1 show ?thesis
43.996 - by simp (erule da_weakenE,auto)
43.997 - next
43.998 - case False
43.999 - with that da_e2 s0_s1 show ?thesis
43.1000 - by simp (erule da_weakenE,auto)
43.1001 - qed
43.1002 - with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
43.1003 - show ?thesis
43.1004 - by (rule validE)
43.1005 + proof (cases "the_Bool b")
43.1006 + case True
43.1007 + with that da_e1 s0_s1 show ?thesis
43.1008 + by simp (erule da_weakenE,auto)
43.1009 + next
43.1010 + case False
43.1011 + with that da_e2 s0_s1 show ?thesis
43.1012 + by simp (erule da_weakenE,auto)
43.1013 + qed
43.1014 + with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
43.1015 + show ?thesis
43.1016 + by (rule validE)
43.1017 next
43.1018 - case False
43.1019 - with valid_then_else P' valid_A conf_s1 eval_then_else
43.1020 - show ?thesis
43.1021 - by (cases rule: validE) iprover+
43.1022 + case False
43.1023 + with valid_then_else P' valid_A conf_s1 eval_then_else
43.1024 + show ?thesis
43.1025 + by (cases rule: validE) iprover+
43.1026 qed
43.1027 moreover
43.1028 from eval wt da conf_s0 wf
43.1029 have "s2\<Colon>\<preceq>(G, L)"
43.1030 - by (rule evaln_type_sound [elim_format]) simp
43.1031 + by (rule evaln_type_sound [elim_format]) simp
43.1032 ultimately show ?thesis ..
43.1033 qed
43.1034 qed
43.1035 @@ -1248,93 +1248,93 @@
43.1036 mode: "mode = invmode statM e" and
43.1037 T: "T =(resTy statM)" and
43.1038 eq_accC_accC': "accC=accC'"
43.1039 - by cases fastsimp+
43.1040 + by cases fastsimp+
43.1041 from da obtain C where
43.1042 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
43.1043 - da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E"
43.1044 - by cases simp
43.1045 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
43.1046 + da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E"
43.1047 + by cases simp
43.1048 from eval eq_accC_accC' obtain a s1 vs s2 s3 s3' s4 invDeclC where
43.1049 - evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
43.1050 + evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
43.1051 evaln_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
43.1052 - invDeclC: "invDeclC = invocation_declclass
43.1053 + invDeclC: "invDeclC = invocation_declclass
43.1054 G mode (store s2) a statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
43.1055 s3: "s3 = init_lvars G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> mode a vs s2" and
43.1056 check: "s3' = check_method_access G
43.1057 accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a s3" and
43.1058 - evaln_methd:
43.1059 + evaln_methd:
43.1060 "G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s4" and
43.1061 s5: "s5=(set_lvars (locals (store s2))) s4"
43.1062 - using normal_s0 by (auto elim: evaln_elim_cases)
43.1063 + using normal_s0 by (auto elim: evaln_elim_cases)
43.1064
43.1065 from evaln_e
43.1066 have eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
43.1067 - by (rule evaln_eval)
43.1068 + by (rule evaln_eval)
43.1069
43.1070 from eval_e _ wt_e wf
43.1071 have s1_no_return: "abrupt s1 \<noteq> Some (Jump Ret)"
43.1072 - by (rule eval_expression_no_jump
43.1073 + by (rule eval_expression_no_jump
43.1074 [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
43.1075 - (insert normal_s0,auto)
43.1076 + (insert normal_s0,auto)
43.1077
43.1078 from valid_e P valid_A conf_s0 evaln_e wt_e da_e
43.1079 obtain "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.1080 - by (rule validE)
43.1081 + by (rule validE)
43.1082 hence Q: "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
43.1083 - by simp
43.1084 + by simp
43.1085 obtain
43.1086 - R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and
43.1087 - conf_s2: "s2\<Colon>\<preceq>(G,L)" and
43.1088 - s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
43.1089 + R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and
43.1090 + conf_s2: "s2\<Colon>\<preceq>(G,L)" and
43.1091 + s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
43.1092 proof (cases "normal s1")
43.1093 - case True
43.1094 - obtain E' where
43.1095 - da_args':
43.1096 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
43.1097 - proof -
43.1098 - from evaln_e wt_e da_e wf True
43.1099 - have "nrm C \<subseteq> dom (locals (store s1))"
43.1100 - by (cases rule: da_good_approx_evalnE) iprover
43.1101 - with da_args show thesis
43.1102 - by (rule da_weakenE) (rule that)
43.1103 - qed
43.1104 - with valid_args Q valid_A conf_s1 evaln_args wt_args
43.1105 - obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
43.1106 - by (rule validE)
43.1107 - moreover
43.1108 - from evaln_args
43.1109 - have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
43.1110 - by (rule evaln_eval)
43.1111 - from this s1_no_return wt_args wf
43.1112 - have "abrupt s2 \<noteq> Some (Jump Ret)"
43.1113 - by (rule eval_expression_list_no_jump
43.1114 + case True
43.1115 + obtain E' where
43.1116 + da_args':
43.1117 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
43.1118 + proof -
43.1119 + from evaln_e wt_e da_e wf True
43.1120 + have "nrm C \<subseteq> dom (locals (store s1))"
43.1121 + by (cases rule: da_good_approx_evalnE) iprover
43.1122 + with da_args show thesis
43.1123 + by (rule da_weakenE) (rule that)
43.1124 + qed
43.1125 + with valid_args Q valid_A conf_s1 evaln_args wt_args
43.1126 + obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
43.1127 + by (rule validE)
43.1128 + moreover
43.1129 + from evaln_args
43.1130 + have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
43.1131 + by (rule evaln_eval)
43.1132 + from this s1_no_return wt_args wf
43.1133 + have "abrupt s2 \<noteq> Some (Jump Ret)"
43.1134 + by (rule eval_expression_list_no_jump
43.1135 [where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
43.1136 - ultimately show ?thesis ..
43.1137 + ultimately show ?thesis ..
43.1138 next
43.1139 - case False
43.1140 - with valid_args Q valid_A conf_s1 evaln_args
43.1141 - obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
43.1142 - by (cases rule: validE) iprover+
43.1143 - moreover
43.1144 - from False evaln_args have "s2=s1"
43.1145 - by auto
43.1146 - with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
43.1147 - by simp
43.1148 - ultimately show ?thesis ..
43.1149 + case False
43.1150 + with valid_args Q valid_A conf_s1 evaln_args
43.1151 + obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
43.1152 + by (cases rule: validE) iprover+
43.1153 + moreover
43.1154 + from False evaln_args have "s2=s1"
43.1155 + by auto
43.1156 + with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
43.1157 + by simp
43.1158 + ultimately show ?thesis ..
43.1159 qed
43.1160
43.1161 obtain invC where
43.1162 - invC: "invC = invocation_class mode (store s2) a statT"
43.1163 - by simp
43.1164 + invC: "invC = invocation_class mode (store s2) a statT"
43.1165 + by simp
43.1166 with s3
43.1167 have invC': "invC = (invocation_class mode (store s3) a statT)"
43.1168 - by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
43.1169 + by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
43.1170 obtain l where
43.1171 - l: "l = locals (store s2)"
43.1172 - by simp
43.1173 + l: "l = locals (store s2)"
43.1174 + by simp
43.1175
43.1176 from eval wt da conf_s0 wf
43.1177 have conf_s5: "s5\<Colon>\<preceq>(G, L)"
43.1178 - by (rule evaln_type_sound [elim_format]) simp
43.1179 + by (rule evaln_type_sound [elim_format]) simp
43.1180 let "PROP ?R" = "\<And> v.
43.1181 (R a\<leftarrow>In3 vs \<and>.
43.1182 (\<lambda>s. invDeclC = invocation_declclass G mode (store s) a statT
43.1183 @@ -1345,128 +1345,128 @@
43.1184 (\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)
43.1185 ) v s3' Z"
43.1186 {
43.1187 - assume abrupt_s3: "\<not> normal s3"
43.1188 - have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
43.1189 - proof -
43.1190 - from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
43.1191 - by (auto simp add: check_method_access_def Let_def)
43.1192 - with R s3 invDeclC invC l abrupt_s3
43.1193 - have R': "PROP ?R"
43.1194 - by auto
43.1195 - have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
43.1196 - (* we need an arbirary environment (here empty) that s2' conforms to
43.1197 + assume abrupt_s3: "\<not> normal s3"
43.1198 + have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
43.1199 + proof -
43.1200 + from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
43.1201 + by (auto simp add: check_method_access_def Let_def)
43.1202 + with R s3 invDeclC invC l abrupt_s3
43.1203 + have R': "PROP ?R"
43.1204 + by auto
43.1205 + have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
43.1206 + (* we need an arbirary environment (here empty) that s2' conforms to
43.1207 to apply validE *)
43.1208 - proof -
43.1209 - from s2_no_return s3
43.1210 - have "abrupt s3 \<noteq> Some (Jump Ret)"
43.1211 - by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
43.1212 - moreover
43.1213 - obtain abr2 str2 where s2: "s2=(abr2,str2)"
43.1214 - by (cases s2)
43.1215 - from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
43.1216 - by (auto simp add: init_lvars_def2 split: split_if_asm)
43.1217 - ultimately show ?thesis
43.1218 - using s3 s2 eq_s3'_s3
43.1219 - apply (simp add: init_lvars_def2)
43.1220 - apply (rule conforms_set_locals [OF _ wlconf_empty])
43.1221 - by auto
43.1222 - qed
43.1223 - from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
43.1224 - have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
43.1225 - by (cases rule: validE) simp+
43.1226 - with s5 l show ?thesis
43.1227 - by simp
43.1228 - qed
43.1229 + proof -
43.1230 + from s2_no_return s3
43.1231 + have "abrupt s3 \<noteq> Some (Jump Ret)"
43.1232 + by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
43.1233 + moreover
43.1234 + obtain abr2 str2 where s2: "s2=(abr2,str2)"
43.1235 + by (cases s2)
43.1236 + from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
43.1237 + by (auto simp add: init_lvars_def2 split: split_if_asm)
43.1238 + ultimately show ?thesis
43.1239 + using s3 s2 eq_s3'_s3
43.1240 + apply (simp add: init_lvars_def2)
43.1241 + apply (rule conforms_set_locals [OF _ wlconf_empty])
43.1242 + by auto
43.1243 + qed
43.1244 + from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
43.1245 + have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
43.1246 + by (cases rule: validE) simp+
43.1247 + with s5 l show ?thesis
43.1248 + by simp
43.1249 + qed
43.1250 } note abrupt_s3_lemma = this
43.1251
43.1252 have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
43.1253 proof (cases "normal s2")
43.1254 - case False
43.1255 - with s3 have abrupt_s3: "\<not> normal s3"
43.1256 - by (cases s2) (simp add: init_lvars_def2)
43.1257 - thus ?thesis
43.1258 - by (rule abrupt_s3_lemma)
43.1259 + case False
43.1260 + with s3 have abrupt_s3: "\<not> normal s3"
43.1261 + by (cases s2) (simp add: init_lvars_def2)
43.1262 + thus ?thesis
43.1263 + by (rule abrupt_s3_lemma)
43.1264 next
43.1265 - case True
43.1266 - note normal_s2 = this
43.1267 - with evaln_args
43.1268 - have normal_s1: "normal s1"
43.1269 - by (rule evaln_no_abrupt)
43.1270 - obtain E' where
43.1271 - da_args':
43.1272 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
43.1273 - proof -
43.1274 - from evaln_e wt_e da_e wf normal_s1
43.1275 - have "nrm C \<subseteq> dom (locals (store s1))"
43.1276 - by (cases rule: da_good_approx_evalnE) iprover
43.1277 - with da_args show thesis
43.1278 - by (rule da_weakenE) (rule that)
43.1279 - qed
43.1280 - from evaln_args
43.1281 - have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
43.1282 - by (rule evaln_eval)
43.1283 - from evaln_e wt_e da_e conf_s0 wf
43.1284 - have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
43.1285 - by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
43.1286 - with normal_s1 normal_s2 eval_args
43.1287 - have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
43.1288 - by (auto dest: eval_gext intro: conf_gext)
43.1289 - from evaln_args wt_args da_args' conf_s1 wf
43.1290 - have conf_args: "list_all2 (conf G (store s2)) vs pTs"
43.1291 - by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
43.1292 - from statM
43.1293 - obtain
43.1294 - statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>"
43.1295 - and
43.1296 - pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
43.1297 - by (blast dest: max_spec2mheads)
43.1298 - show ?thesis
43.1299 - proof (cases "normal s3")
43.1300 - case False
43.1301 - thus ?thesis
43.1302 - by (rule abrupt_s3_lemma)
43.1303 - next
43.1304 - case True
43.1305 - note normal_s3 = this
43.1306 - with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
43.1307 - by (cases s2) (auto simp add: init_lvars_def2)
43.1308 - from conf_s2 conf_a_s2 wf notNull invC
43.1309 - have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
43.1310 - by (cases s2) (auto intro: DynT_propI)
43.1311 + case True
43.1312 + note normal_s2 = this
43.1313 + with evaln_args
43.1314 + have normal_s1: "normal s1"
43.1315 + by (rule evaln_no_abrupt)
43.1316 + obtain E' where
43.1317 + da_args':
43.1318 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
43.1319 + proof -
43.1320 + from evaln_e wt_e da_e wf normal_s1
43.1321 + have "nrm C \<subseteq> dom (locals (store s1))"
43.1322 + by (cases rule: da_good_approx_evalnE) iprover
43.1323 + with da_args show thesis
43.1324 + by (rule da_weakenE) (rule that)
43.1325 + qed
43.1326 + from evaln_args
43.1327 + have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
43.1328 + by (rule evaln_eval)
43.1329 + from evaln_e wt_e da_e conf_s0 wf
43.1330 + have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
43.1331 + by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
43.1332 + with normal_s1 normal_s2 eval_args
43.1333 + have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
43.1334 + by (auto dest: eval_gext intro: conf_gext)
43.1335 + from evaln_args wt_args da_args' conf_s1 wf
43.1336 + have conf_args: "list_all2 (conf G (store s2)) vs pTs"
43.1337 + by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
43.1338 + from statM
43.1339 + obtain
43.1340 + statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>"
43.1341 + and
43.1342 + pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
43.1343 + by (blast dest: max_spec2mheads)
43.1344 + show ?thesis
43.1345 + proof (cases "normal s3")
43.1346 + case False
43.1347 + thus ?thesis
43.1348 + by (rule abrupt_s3_lemma)
43.1349 + next
43.1350 + case True
43.1351 + note normal_s3 = this
43.1352 + with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
43.1353 + by (cases s2) (auto simp add: init_lvars_def2)
43.1354 + from conf_s2 conf_a_s2 wf notNull invC
43.1355 + have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
43.1356 + by (cases s2) (auto intro: DynT_propI)
43.1357
43.1358 - with wt_e statM' invC mode wf
43.1359 - obtain dynM where
43.1360 + with wt_e statM' invC mode wf
43.1361 + obtain dynM where
43.1362 dynM: "dynlookup G statT invC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
43.1363 acc_dynM: "G \<turnstile>Methd \<lparr>name=mn,parTs=pTs'\<rparr> dynM
43.1364 in invC dyn_accessible_from accC"
43.1365 - by (force dest!: call_access_ok)
43.1366 - with invC' check eq_accC_accC'
43.1367 - have eq_s3'_s3: "s3'=s3"
43.1368 - by (auto simp add: check_method_access_def Let_def)
43.1369 -
43.1370 - with dynT_prop R s3 invDeclC invC l
43.1371 - have R': "PROP ?R"
43.1372 - by auto
43.1373 + by (force dest!: call_access_ok)
43.1374 + with invC' check eq_accC_accC'
43.1375 + have eq_s3'_s3: "s3'=s3"
43.1376 + by (auto simp add: check_method_access_def Let_def)
43.1377 +
43.1378 + with dynT_prop R s3 invDeclC invC l
43.1379 + have R': "PROP ?R"
43.1380 + by auto
43.1381
43.1382 - from dynT_prop wf wt_e statM' mode invC invDeclC dynM
43.1383 - obtain
43.1384 + from dynT_prop wf wt_e statM' mode invC invDeclC dynM
43.1385 + obtain
43.1386 dynM: "dynlookup G statT invC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
43.1387 - wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
43.1388 - dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
43.1389 + wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
43.1390 + dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
43.1391 iscls_invDeclC: "is_class G invDeclC" and
43.1392 - invDeclC': "invDeclC = declclass dynM" and
43.1393 - invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
43.1394 - resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
43.1395 - is_static_eq: "is_static dynM = is_static statM" and
43.1396 - involved_classes_prop:
43.1397 + invDeclC': "invDeclC = declclass dynM" and
43.1398 + invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
43.1399 + resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
43.1400 + is_static_eq: "is_static dynM = is_static statM" and
43.1401 + involved_classes_prop:
43.1402 "(if invmode statM e = IntVir
43.1403 then \<forall>statC. statT = ClassT statC \<longrightarrow> G\<turnstile>invC\<preceq>\<^sub>C statC
43.1404 else ((\<exists>statC. statT = ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C invDeclC) \<or>
43.1405 (\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
43.1406 statDeclT = ClassT invDeclC)"
43.1407 - by (cases rule: DynT_mheadsE) simp
43.1408 - obtain L' where
43.1409 - L':"L'=(\<lambda> k.
43.1410 + by (cases rule: DynT_mheadsE) simp
43.1411 + obtain L' where
43.1412 + L':"L'=(\<lambda> k.
43.1413 (case k of
43.1414 EName e
43.1415 \<Rightarrow> (case e of
43.1416 @@ -1476,121 +1476,121 @@
43.1417 | Res \<Rightarrow> Some (resTy dynM))
43.1418 | This \<Rightarrow> if is_static statM
43.1419 then None else Some (Class invDeclC)))"
43.1420 - by simp
43.1421 - from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
43.1422 + by simp
43.1423 + from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
43.1424 wf eval_args conf_a mode notNull wf_dynM involved_classes_prop
43.1425 - have conf_s3: "s3\<Colon>\<preceq>(G,L')"
43.1426 - apply -
43.1427 + have conf_s3: "s3\<Colon>\<preceq>(G,L')"
43.1428 + apply -
43.1429 (* FIXME confomrs_init_lvars should be
43.1430 adjusted to be more directy applicable *)
43.1431 - apply (drule conforms_init_lvars [of G invDeclC
43.1432 + apply (drule conforms_init_lvars [of G invDeclC
43.1433 "\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2"
43.1434 L statT invC a "(statDeclT,statM)" e])
43.1435 - apply (rule wf)
43.1436 - apply (rule conf_args)
43.1437 - apply (simp add: pTs_widen)
43.1438 - apply (cases s2,simp)
43.1439 - apply (rule dynM')
43.1440 - apply (force dest: ty_expr_is_type)
43.1441 - apply (rule invC_widen)
43.1442 - apply (force intro: conf_gext dest: eval_gext)
43.1443 - apply simp
43.1444 - apply simp
43.1445 - apply (simp add: invC)
43.1446 - apply (simp add: invDeclC)
43.1447 - apply (simp add: normal_s2)
43.1448 - apply (cases s2, simp add: L' init_lvars_def2 s3
43.1449 - cong add: lname.case_cong ename.case_cong)
43.1450 - done
43.1451 - with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
43.1452 - from is_static_eq wf_dynM L'
43.1453 - obtain mthdT where
43.1454 - "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
43.1455 + apply (rule wf)
43.1456 + apply (rule conf_args)
43.1457 + apply (simp add: pTs_widen)
43.1458 + apply (cases s2,simp)
43.1459 + apply (rule dynM')
43.1460 + apply (force dest: ty_expr_is_type)
43.1461 + apply (rule invC_widen)
43.1462 + apply (force intro: conf_gext dest: eval_gext)
43.1463 + apply simp
43.1464 + apply simp
43.1465 + apply (simp add: invC)
43.1466 + apply (simp add: invDeclC)
43.1467 + apply (simp add: normal_s2)
43.1468 + apply (cases s2, simp add: L' init_lvars_def2 s3
43.1469 + cong add: lname.case_cong ename.case_cong)
43.1470 + done
43.1471 + with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
43.1472 + from is_static_eq wf_dynM L'
43.1473 + obtain mthdT where
43.1474 + "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
43.1475 \<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" and
43.1476 - mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
43.1477 - by - (drule wf_mdecl_bodyD,
43.1478 + mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
43.1479 + by - (drule wf_mdecl_bodyD,
43.1480 auto simp add: callee_lcl_def
43.1481 cong add: lname.case_cong ename.case_cong)
43.1482 - with dynM' iscls_invDeclC invDeclC'
43.1483 - have
43.1484 - wt_methd:
43.1485 - "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
43.1486 + with dynM' iscls_invDeclC invDeclC'
43.1487 + have
43.1488 + wt_methd:
43.1489 + "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
43.1490 \<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
43.1491 - by (auto intro: wt.Methd)
43.1492 - obtain M where
43.1493 - da_methd:
43.1494 - "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
43.1495 - \<turnstile> dom (locals (store s3'))
43.1496 + by (auto intro: wt.Methd)
43.1497 + obtain M where
43.1498 + da_methd:
43.1499 + "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
43.1500 + \<turnstile> dom (locals (store s3'))
43.1501 \<guillemotright>\<langle>Methd invDeclC \<lparr>name=mn,parTs=pTs'\<rparr>\<rangle>\<^sub>e\<guillemotright> M"
43.1502 - proof -
43.1503 - from wf_dynM
43.1504 - obtain M' where
43.1505 - da_body:
43.1506 - "\<lparr>prg=G, cls=invDeclC
43.1507 + proof -
43.1508 + from wf_dynM
43.1509 + obtain M' where
43.1510 + da_body:
43.1511 + "\<lparr>prg=G, cls=invDeclC
43.1512 ,lcl=callee_lcl invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> (mthd dynM)
43.1513 \<rparr> \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'" and
43.1514 res: "Result \<in> nrm M'"
43.1515 - by (rule wf_mdeclE) iprover
43.1516 - from da_body is_static_eq L' have
43.1517 - "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
43.1518 + by (rule wf_mdeclE) iprover
43.1519 + from da_body is_static_eq L' have
43.1520 + "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
43.1521 \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'"
43.1522 - by (simp add: callee_lcl_def
43.1523 + by (simp add: callee_lcl_def
43.1524 cong add: lname.case_cong ename.case_cong)
43.1525 - moreover have "parameters (mthd dynM) \<subseteq> dom (locals (store s3'))"
43.1526 - proof -
43.1527 - from is_static_eq
43.1528 - have "(invmode (mthd dynM) e) = (invmode statM e)"
43.1529 - by (simp add: invmode_def)
43.1530 - moreover
43.1531 - have "length (pars (mthd dynM)) = length vs"
43.1532 - proof -
43.1533 - from normal_s2 conf_args
43.1534 - have "length vs = length pTs"
43.1535 - by (simp add: list_all2_def)
43.1536 - also from pTs_widen
43.1537 - have "\<dots> = length pTs'"
43.1538 - by (simp add: widens_def list_all2_def)
43.1539 - also from wf_dynM
43.1540 - have "\<dots> = length (pars (mthd dynM))"
43.1541 - by (simp add: wf_mdecl_def wf_mhead_def)
43.1542 - finally show ?thesis ..
43.1543 - qed
43.1544 - moreover note s3 dynM' is_static_eq normal_s2 mode
43.1545 - ultimately
43.1546 - have "parameters (mthd dynM) = dom (locals (store s3))"
43.1547 - using dom_locals_init_lvars
43.1548 + moreover have "parameters (mthd dynM) \<subseteq> dom (locals (store s3'))"
43.1549 + proof -
43.1550 + from is_static_eq
43.1551 + have "(invmode (mthd dynM) e) = (invmode statM e)"
43.1552 + by (simp add: invmode_def)
43.1553 + moreover
43.1554 + have "length (pars (mthd dynM)) = length vs"
43.1555 + proof -
43.1556 + from normal_s2 conf_args
43.1557 + have "length vs = length pTs"
43.1558 + by (simp add: list_all2_def)
43.1559 + also from pTs_widen
43.1560 + have "\<dots> = length pTs'"
43.1561 + by (simp add: widens_def list_all2_def)
43.1562 + also from wf_dynM
43.1563 + have "\<dots> = length (pars (mthd dynM))"
43.1564 + by (simp add: wf_mdecl_def wf_mhead_def)
43.1565 + finally show ?thesis ..
43.1566 + qed
43.1567 + moreover note s3 dynM' is_static_eq normal_s2 mode
43.1568 + ultimately
43.1569 + have "parameters (mthd dynM) = dom (locals (store s3))"
43.1570 + using dom_locals_init_lvars
43.1571 [of "mthd dynM" G invDeclC "\<lparr>name=mn,parTs=pTs'\<rparr>" vs e a s2]
43.1572 - by simp
43.1573 - thus ?thesis using eq_s3'_s3 by simp
43.1574 - qed
43.1575 - ultimately obtain M2 where
43.1576 - da:
43.1577 - "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
43.1578 + by simp
43.1579 + thus ?thesis using eq_s3'_s3 by simp
43.1580 + qed
43.1581 + ultimately obtain M2 where
43.1582 + da:
43.1583 + "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
43.1584 \<turnstile> dom (locals (store s3')) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M2" and
43.1585 M2: "nrm M' \<subseteq> nrm M2"
43.1586 - by (rule da_weakenE)
43.1587 - from res M2 have "Result \<in> nrm M2"
43.1588 - by blast
43.1589 - moreover from wf_dynM
43.1590 - have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
43.1591 - by (rule wf_mdeclE)
43.1592 - ultimately
43.1593 - obtain M3 where
43.1594 - "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3'))
43.1595 + by (rule da_weakenE)
43.1596 + from res M2 have "Result \<in> nrm M2"
43.1597 + by blast
43.1598 + moreover from wf_dynM
43.1599 + have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
43.1600 + by (rule wf_mdeclE)
43.1601 + ultimately
43.1602 + obtain M3 where
43.1603 + "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3'))
43.1604 \<guillemotright>\<langle>Body (declclass dynM) (stmt (mbody (mthd dynM)))\<rangle>\<guillemotright> M3"
43.1605 - using da
43.1606 - by (iprover intro: da.Body assigned.select_convs)
43.1607 - from _ this [simplified]
43.1608 - show thesis
43.1609 - by (rule da.Methd [simplified,elim_format])
43.1610 - (auto intro: dynM' that)
43.1611 - qed
43.1612 - from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
43.1613 - have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
43.1614 - by (cases rule: validE) iprover+
43.1615 - with s5 l show ?thesis
43.1616 - by simp
43.1617 - qed
43.1618 + using da
43.1619 + by (iprover intro: da.Body assigned.select_convs)
43.1620 + from _ this [simplified]
43.1621 + show thesis
43.1622 + by (rule da.Methd [simplified,elim_format])
43.1623 + (auto intro: dynM' that)
43.1624 + qed
43.1625 + from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
43.1626 + have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
43.1627 + by (cases rule: validE) iprover+
43.1628 + with s5 l show ?thesis
43.1629 + by simp
43.1630 + qed
43.1631 qed
43.1632 with conf_s5 show ?thesis by iprover
43.1633 qed
43.1634 @@ -1618,61 +1618,61 @@
43.1635 show "R \<lfloor>v\<rfloor>\<^sub>e s4 Z \<and> s4\<Colon>\<preceq>(G, L)"
43.1636 proof -
43.1637 from wt obtain
43.1638 - iscls_D: "is_class G D" and
43.1639 + iscls_D: "is_class G D" and
43.1640 wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Init D\<Colon>\<surd>" and
43.1641 wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>"
43.1642 - by cases auto
43.1643 + by cases auto
43.1644 obtain I where
43.1645 - da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
43.1646 - by (auto intro: da_Init [simplified] assigned.select_convs)
43.1647 + da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
43.1648 + by (auto intro: da_Init [simplified] assigned.select_convs)
43.1649 from da obtain C where
43.1650 - da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
43.1651 - jmpOk: "jumpNestingOkS {Ret} c"
43.1652 - by cases simp
43.1653 + da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
43.1654 + jmpOk: "jumpNestingOkS {Ret} c"
43.1655 + by cases simp
43.1656 from eval obtain s1 s2 s3 where
43.1657 - eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
43.1658 + eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
43.1659 eval_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2" and
43.1660 - v: "v = the (locals (store s2) Result)" and
43.1661 + v: "v = the (locals (store s2) Result)" and
43.1662 s3: "s3 =(if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or>
43.1663 abrupt s2 = Some (Jump (Cont l))
43.1664 then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)"and
43.1665 s4: "s4 = abupd (absorb Ret) s3"
43.1666 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1667 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1668 obtain C' where
43.1669 - da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
43.1670 + da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
43.1671 proof -
43.1672 - from eval_init
43.1673 - have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
43.1674 - by (rule dom_locals_evaln_mono_elim)
43.1675 - with da_c show thesis by (rule da_weakenE) (rule that)
43.1676 + from eval_init
43.1677 + have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
43.1678 + by (rule dom_locals_evaln_mono_elim)
43.1679 + with da_c show thesis by (rule da_weakenE) (rule that)
43.1680 qed
43.1681 from valid_init P valid_A conf_s0 eval_init wt_init da_init
43.1682 obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.1683 - by (rule validE)
43.1684 + by (rule validE)
43.1685 from valid_c Q valid_A conf_s1 eval_c wt_c da_c'
43.1686 have R: "(\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result)))
43.1687 \<diamondsuit> s2 Z"
43.1688 - by (rule validE)
43.1689 + by (rule validE)
43.1690 have "s3=s2"
43.1691 proof -
43.1692 - from eval_init [THEN evaln_eval] wf
43.1693 - have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
43.1694 - by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
43.1695 + from eval_init [THEN evaln_eval] wf
43.1696 + have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
43.1697 + by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
43.1698 insert normal_s0,auto)
43.1699 - from eval_c [THEN evaln_eval] _ wt_c wf
43.1700 - have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
43.1701 - by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
43.1702 - moreover note s3
43.1703 - ultimately show ?thesis
43.1704 - by (force split: split_if)
43.1705 + from eval_c [THEN evaln_eval] _ wt_c wf
43.1706 + have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
43.1707 + by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
43.1708 + moreover note s3
43.1709 + ultimately show ?thesis
43.1710 + by (force split: split_if)
43.1711 qed
43.1712 with R v s4
43.1713 have "R \<lfloor>v\<rfloor>\<^sub>e s4 Z"
43.1714 - by simp
43.1715 + by simp
43.1716 moreover
43.1717 from eval wt da conf_s0 wf
43.1718 have "s4\<Colon>\<preceq>(G, L)"
43.1719 - by (rule evaln_type_sound [elim_format]) simp
43.1720 + by (rule evaln_type_sound [elim_format]) simp
43.1721 ultimately show ?thesis ..
43.1722 qed
43.1723 qed
43.1724 @@ -1688,9 +1688,9 @@
43.1725 show "P \<lfloor>vs\<rfloor>\<^sub>l s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.1726 proof -
43.1727 from eval obtain "vs=[]" "s1=s0"
43.1728 - using normal_s0 by (auto elim: evaln_elim_cases)
43.1729 + using normal_s0 by (auto elim: evaln_elim_cases)
43.1730 with P conf_s0 show ?thesis
43.1731 - by simp
43.1732 + by simp
43.1733 qed
43.1734 qed
43.1735 next
43.1736 @@ -1711,50 +1711,50 @@
43.1737 show "R \<lfloor>v\<rfloor>\<^sub>l s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
43.1738 proof -
43.1739 from wt obtain eT esT where
43.1740 - wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT" and
43.1741 - wt_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>esT"
43.1742 - by cases simp
43.1743 + wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT" and
43.1744 + wt_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>esT"
43.1745 + by cases simp
43.1746 from da obtain E1 where
43.1747 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
43.1748 - da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E"
43.1749 - by cases simp
43.1750 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
43.1751 + da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E"
43.1752 + by cases simp
43.1753 from eval obtain s1 ve vs where
43.1754 - eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
43.1755 - eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
43.1756 - v: "v=ve#vs"
43.1757 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1758 + eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
43.1759 + eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
43.1760 + v: "v=ve#vs"
43.1761 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1762 from valid_e P valid_A conf_s0 eval_e wt_e da_e
43.1763 obtain Q: "Q \<lfloor>ve\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.1764 - by (rule validE)
43.1765 + by (rule validE)
43.1766 from Q have Q': "\<And> v. (Q\<leftarrow>\<lfloor>ve\<rfloor>\<^sub>e) v s1 Z"
43.1767 - by simp
43.1768 + by simp
43.1769 have "(\<lambda>Vals:vs:. R\<leftarrow>\<lfloor>(ve # vs)\<rfloor>\<^sub>l) \<lfloor>vs\<rfloor>\<^sub>l s2 Z"
43.1770 proof (cases "normal s1")
43.1771 - case True
43.1772 - obtain E' where
43.1773 - da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
43.1774 - proof -
43.1775 - from eval_e wt_e da_e wf True
43.1776 - have "nrm E1 \<subseteq> dom (locals (store s1))"
43.1777 - by (cases rule: da_good_approx_evalnE) iprover
43.1778 - with da_es show thesis
43.1779 - by (rule da_weakenE) (rule that)
43.1780 - qed
43.1781 - from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
43.1782 - show ?thesis
43.1783 - by (rule validE)
43.1784 + case True
43.1785 + obtain E' where
43.1786 + da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
43.1787 + proof -
43.1788 + from eval_e wt_e da_e wf True
43.1789 + have "nrm E1 \<subseteq> dom (locals (store s1))"
43.1790 + by (cases rule: da_good_approx_evalnE) iprover
43.1791 + with da_es show thesis
43.1792 + by (rule da_weakenE) (rule that)
43.1793 + qed
43.1794 + from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
43.1795 + show ?thesis
43.1796 + by (rule validE)
43.1797 next
43.1798 - case False
43.1799 - with valid_es Q' valid_A conf_s1 eval_es
43.1800 - show ?thesis
43.1801 - by (cases rule: validE) iprover+
43.1802 + case False
43.1803 + with valid_es Q' valid_A conf_s1 eval_es
43.1804 + show ?thesis
43.1805 + by (cases rule: validE) iprover+
43.1806 qed
43.1807 with v have "R \<lfloor>v\<rfloor>\<^sub>l s2 Z"
43.1808 - by simp
43.1809 + by simp
43.1810 moreover
43.1811 from eval wt da conf_s0 wf
43.1812 have "s2\<Colon>\<preceq>(G, L)"
43.1813 - by (rule evaln_type_sound [elim_format]) simp
43.1814 + by (rule evaln_type_sound [elim_format]) simp
43.1815 ultimately show ?thesis ..
43.1816 qed
43.1817 qed
43.1818 @@ -1770,9 +1770,9 @@
43.1819 show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.1820 proof -
43.1821 from eval obtain "s1=s0"
43.1822 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1823 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1824 with P conf_s0 show ?thesis
43.1825 - by simp
43.1826 + by simp
43.1827 qed
43.1828 qed
43.1829 next
43.1830 @@ -1791,17 +1791,17 @@
43.1831 show "Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
43.1832 proof -
43.1833 from wt obtain eT where
43.1834 - wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.1835 - by cases simp
43.1836 + wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
43.1837 + by cases simp
43.1838 from da obtain E where
43.1839 - da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
43.1840 - by cases simp
43.1841 + da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
43.1842 + by cases simp
43.1843 from eval obtain v where
43.1844 - eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
43.1845 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1846 + eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
43.1847 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1848 from valid_e P valid_A conf_s0 eval_e wt_e da_e
43.1849 obtain Q: "(Q\<leftarrow>\<diamondsuit>) \<lfloor>v\<rfloor>\<^sub>e s1 Z" and "s1\<Colon>\<preceq>(G,L)"
43.1850 - by (rule validE)
43.1851 + by (rule validE)
43.1852 thus ?thesis by simp
43.1853 qed
43.1854 qed
43.1855 @@ -1821,24 +1821,24 @@
43.1856 show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
43.1857 proof -
43.1858 from wt obtain
43.1859 - wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
43.1860 - by cases simp
43.1861 + wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
43.1862 + by cases simp
43.1863 from da obtain E where
43.1864 - da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
43.1865 - by cases simp
43.1866 + da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
43.1867 + by cases simp
43.1868 from eval obtain s1 where
43.1869 - eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
43.1870 - s2: "s2 = abupd (absorb l) s1"
43.1871 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1872 + eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
43.1873 + s2: "s2 = abupd (absorb l) s1"
43.1874 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1875 from valid_c P valid_A conf_s0 eval_c wt_c da_c
43.1876 obtain Q: "(abupd (absorb l) .; Q) \<diamondsuit> s1 Z"
43.1877 - by (rule validE)
43.1878 + by (rule validE)
43.1879 with s2 have "Q \<diamondsuit> s2 Z"
43.1880 - by simp
43.1881 + by simp
43.1882 moreover
43.1883 from eval wt da conf_s0 wf
43.1884 have "s2\<Colon>\<preceq>(G, L)"
43.1885 - by (rule evaln_type_sound [elim_format]) simp
43.1886 + by (rule evaln_type_sound [elim_format]) simp
43.1887 ultimately show ?thesis ..
43.1888 qed
43.1889 qed
43.1890 @@ -1859,45 +1859,45 @@
43.1891 show "R \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
43.1892 proof -
43.1893 from eval obtain s1 where
43.1894 - eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
43.1895 - eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
43.1896 - using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1897 + eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
43.1898 + eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
43.1899 + using normal_s0 by (fastsimp elim: evaln_elim_cases)
43.1900 from wt obtain
43.1901 - wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
43.1902 + wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
43.1903 wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
43.1904 - by cases simp
43.1905 + by cases simp
43.1906 from da obtain C1 C2 where
43.1907 - da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
43.1908 - da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
43.1909 - by cases simp
43.1910 + da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
43.1911 + da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
43.1912 + by cases simp
43.1913 from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
43.1914 obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.1915 - by (rule validE)
43.1916 + by (rule validE)
43.1917 have "R \<diamondsuit> s2 Z"
43.1918 proof (cases "normal s1")
43.1919 - case True
43.1920 - obtain C2' where
43.1921 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
43.1922 - proof -
43.1923 - from eval_c1 wt_c1 da_c1 wf True
43.1924 - have "nrm C1 \<subseteq> dom (locals (store s1))"
43.1925 - by (cases rule: da_good_approx_evalnE) iprover
43.1926 - with da_c2 show thesis
43.1927 - by (rule da_weakenE) (rule that)
43.1928 - qed
43.1929 - with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2
43.1930 - show ?thesis
43.1931 - by (rule validE)
43.1932 + case True
43.1933 + obtain C2' where
43.1934 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
43.1935 + proof -
43.1936 + from eval_c1 wt_c1 da_c1 wf True
43.1937 + have "nrm C1 \<subseteq> dom (locals (store s1))"
43.1938 + by (cases rule: da_good_approx_evalnE) iprover
43.1939 + with da_c2 show thesis
43.1940 + by (rule da_weakenE) (rule that)
43.1941 + qed
43.1942 + with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2
43.1943 + show ?thesis
43.1944 + by (rule validE)
43.1945 next
43.1946 - case False
43.1947 - from valid_c2 Q valid_A conf_s1 eval_c2 False
43.1948 - show ?thesis
43.1949 - by (cases rule: validE) iprover+
43.1950 + case False
43.1951 + from valid_c2 Q valid_A conf_s1 eval_c2 False
43.1952 + show ?thesis
43.1953 + by (cases rule: validE) iprover+
43.1954 qed
43.1955 moreover
43.1956 from eval wt da conf_s0 wf
43.1957 have "s2\<Colon>\<preceq>(G, L)"
43.1958 - by (rule evaln_type_sound [elim_format]) simp
43.1959 + by (rule evaln_type_sound [elim_format]) simp
43.1960 ultimately show ?thesis ..
43.1961 qed
43.1962 qed
43.1963 @@ -1920,61 +1920,61 @@
43.1964 show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
43.1965 proof -
43.1966 from eval obtain b s1 where
43.1967 - eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
43.1968 - eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
43.1969 - using normal_s0 by (auto elim: evaln_elim_cases)
43.1970 + eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
43.1971 + eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
43.1972 + using normal_s0 by (auto elim: evaln_elim_cases)
43.1973 from wt obtain
43.1974 - wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
43.1975 - wt_then_else: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
43.1976 - by cases (simp split: split_if)
43.1977 + wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
43.1978 + wt_then_else: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
43.1979 + by cases (simp split: split_if)
43.1980 from da obtain E S where
43.1981 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
43.1982 - da_then_else:
43.1983 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
43.1984 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
43.1985 + da_then_else:
43.1986 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
43.1987 (dom (locals (store s0)) \<union> assigns_if (the_Bool b) e)
43.1988 \<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S"
43.1989 - by cases (cases "the_Bool b",auto)
43.1990 + by cases (cases "the_Bool b",auto)
43.1991 from valid_e P valid_A conf_s0 eval_e wt_e da_e
43.1992 obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
43.1993 - by (rule validE)
43.1994 + by (rule validE)
43.1995 hence P': "\<And>v. (P'\<leftarrow>=the_Bool b) v s1 Z"
43.1996 - by (cases "normal s1") auto
43.1997 + by (cases "normal s1") auto
43.1998 have "Q \<diamondsuit> s2 Z"
43.1999 proof (cases "normal s1")
43.2000 - case True
43.2001 - have s0_s1: "dom (locals (store s0))
43.2002 + case True
43.2003 + have s0_s1: "dom (locals (store s0))
43.2004 \<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
43.2005 - proof -
43.2006 - from eval_e
43.2007 - have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
43.2008 - by (rule evaln_eval)
43.2009 - hence
43.2010 - "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
43.2011 - by (rule dom_locals_eval_mono_elim)
43.2012 + proof -
43.2013 + from eval_e
43.2014 + have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
43.2015 + by (rule evaln_eval)
43.2016 + hence
43.2017 + "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
43.2018 + by (rule dom_locals_eval_mono_elim)
43.2019 moreover
43.2020 - from eval_e' True wt_e
43.2021 - have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
43.2022 - by (rule assigns_if_good_approx')
43.2023 - ultimately show ?thesis by (rule Un_least)
43.2024 - qed
43.2025 - with da_then_else
43.2026 - obtain S' where
43.2027 - "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.2028 + from eval_e' True wt_e
43.2029 + have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
43.2030 + by (rule assigns_if_good_approx')
43.2031 + ultimately show ?thesis by (rule Un_least)
43.2032 + qed
43.2033 + with da_then_else
43.2034 + obtain S' where
43.2035 + "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.2036 \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S'"
43.2037 - by (rule da_weakenE)
43.2038 - with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
43.2039 - show ?thesis
43.2040 - by (rule validE)
43.2041 + by (rule da_weakenE)
43.2042 + with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
43.2043 + show ?thesis
43.2044 + by (rule validE)
43.2045 next
43.2046 - case False
43.2047 - with valid_then_else P' valid_A conf_s1 eval_then_else
43.2048 - show ?thesis
43.2049 - by (cases rule: validE) iprover+
43.2050 + case False
43.2051 + with valid_then_else P' valid_A conf_s1 eval_then_else
43.2052 + show ?thesis
43.2053 + by (cases rule: validE) iprover+
43.2054 qed
43.2055 moreover
43.2056 from eval wt da conf_s0 wf
43.2057 have "s2\<Colon>\<preceq>(G, L)"
43.2058 - by (rule evaln_type_sound [elim_format]) simp
43.2059 + by (rule evaln_type_sound [elim_format]) simp
43.2060 ultimately show ?thesis ..
43.2061 qed
43.2062 qed
43.2063 @@ -1998,7 +1998,7 @@
43.2064 proof -
43.2065 --{* From the given hypothesises @{text valid_e} and @{text valid_c}
43.2066 we can only reach the state after unfolding the loop once, i.e.
43.2067 - @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
43.2068 + @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
43.2069 @{term c}. To gain validity of the further execution of while, to
43.2070 finally get @{term "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"} we have to get
43.2071 a hypothesis about the subsequent unfoldings (the whole loop again),
43.2072 @@ -2008,202 +2008,202 @@
43.2073 @{text valid_c} in the goal.
43.2074 *}
43.2075 {
43.2076 - fix t s s' v
43.2077 - assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
43.2078 - hence "\<And> Y' T E.
43.2079 + fix t s s' v
43.2080 + assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
43.2081 + hence "\<And> Y' T E.
43.2082 \<lbrakk>t = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s; \<forall>t\<in>A. G\<Turnstile>n\<Colon>t; P Y' s Z; s\<Colon>\<preceq>(G, L);
43.2083 normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
43.2084 normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>E
43.2085 \<rbrakk>\<Longrightarrow> (P'\<leftarrow>=False\<down>=\<diamondsuit>) v s' Z"
43.2086 - (is "PROP ?Hyp n t s v s'")
43.2087 - proof (induct)
43.2088 - case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
43.2089 - note while = `(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
43.2090 + (is "PROP ?Hyp n t s v s'")
43.2091 + proof (induct)
43.2092 + case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
43.2093 + note while = `(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
43.2094 hence eqs: "l'=l" "e'=e" "c'=c" by simp_all
43.2095 - note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
43.2096 - note P = `P Y' (Norm s0') Z`
43.2097 - note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
43.2098 + note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
43.2099 + note P = `P Y' (Norm s0') Z`
43.2100 + note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
43.2101 have wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<Colon>T"
43.2102 - using Loop.prems eqs by simp
43.2103 - have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
43.2104 + using Loop.prems eqs by simp
43.2105 + have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
43.2106 dom (locals (store ((Norm s0')::state)))\<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright>E"
43.2107 - using Loop.prems eqs by simp
43.2108 - have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'"
43.2109 - using Loop.hyps eqs by simp
43.2110 - show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
43.2111 - proof -
43.2112 - from wt obtain
43.2113 - wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
43.2114 + using Loop.prems eqs by simp
43.2115 + have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'"
43.2116 + using Loop.hyps eqs by simp
43.2117 + show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
43.2118 + proof -
43.2119 + from wt obtain
43.2120 + wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
43.2121 wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>"
43.2122 - by cases (simp add: eqs)
43.2123 - from da obtain E S where
43.2124 - da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.2125 + by cases (simp add: eqs)
43.2126 + from da obtain E S where
43.2127 + da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.2128 \<turnstile> dom (locals (store ((Norm s0')::state))) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
43.2129 - da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.2130 + da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
43.2131 \<turnstile> (dom (locals (store ((Norm s0')::state)))
43.2132 \<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S"
43.2133 - by cases (simp add: eqs)
43.2134 - from evaln_e
43.2135 - have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
43.2136 - by (rule evaln_eval)
43.2137 - from valid_e P valid_A conf_s0' evaln_e wt_e da_e
43.2138 - obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
43.2139 - by (rule validE)
43.2140 - show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
43.2141 - proof (cases "normal s1'")
43.2142 - case True
43.2143 - note normal_s1'=this
43.2144 - show ?thesis
43.2145 - proof (cases "the_Bool b")
43.2146 - case True
43.2147 - with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
43.2148 - by auto
43.2149 - from True Loop.hyps obtain
43.2150 - eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and
43.2151 - eval_while:
43.2152 - "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
43.2153 - by (simp add: eqs)
43.2154 - from True Loop.hyps have
43.2155 - hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s
43.2156 + by cases (simp add: eqs)
43.2157 + from evaln_e
43.2158 + have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
43.2159 + by (rule evaln_eval)
43.2160 + from valid_e P valid_A conf_s0' evaln_e wt_e da_e
43.2161 + obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
43.2162 + by (rule validE)
43.2163 + show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
43.2164 + proof (cases "normal s1'")
43.2165 + case True
43.2166 + note normal_s1'=this
43.2167 + show ?thesis
43.2168 + proof (cases "the_Bool b")
43.2169 + case True
43.2170 + with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
43.2171 + by auto
43.2172 + from True Loop.hyps obtain
43.2173 + eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and
43.2174 + eval_while:
43.2175 + "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
43.2176 + by (simp add: eqs)
43.2177 + from True Loop.hyps have
43.2178 + hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s
43.2179 (abupd (absorb (Cont l')) s2') \<diamondsuit> s3'"
43.2180 - apply (simp only: True if_True eqs)
43.2181 - apply (elim conjE)