eliminated hard tabulators, guessing at each author's individual tab-width;
tuned headers;
--- a/doc-src/Intro/bool.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/Intro/bool.thy Sat Oct 17 14:43:18 2009 +0200
@@ -1,5 +1,5 @@
Bool = FOL +
-types bool 0
-arities bool :: term
-consts tt,ff :: "bool"
+types bool 0
+arities bool :: term
+consts tt,ff :: "bool"
end
--- a/doc-src/Intro/list.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/Intro/list.thy Sat Oct 17 14:43:18 2009 +0200
@@ -1,6 +1,6 @@
List = FOL +
-types list 1
-arities list :: (term)term
-consts Nil :: "'a list"
- Cons :: "['a, 'a list] => 'a list"
+types list 1
+arities list :: (term)term
+consts Nil :: "'a list"
+ Cons :: "['a, 'a list] => 'a list"
end
--- a/doc-src/Locales/Locales/Examples.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/Locales/Locales/Examples.thy Sat Oct 17 14:43:18 2009 +0200
@@ -171,13 +171,13 @@
proof (rule anti_sym)
from inf' show "i \<sqsubseteq> i'"
proof (rule is_inf_greatest)
- from inf show "i \<sqsubseteq> x" ..
- from inf show "i \<sqsubseteq> y" ..
+ from inf show "i \<sqsubseteq> x" ..
+ from inf show "i \<sqsubseteq> y" ..
qed
from inf show "i' \<sqsubseteq> i"
proof (rule is_inf_greatest)
- from inf' show "i' \<sqsubseteq> x" ..
- from inf' show "i' \<sqsubseteq> y" ..
+ from inf' show "i' \<sqsubseteq> x" ..
+ from inf' show "i' \<sqsubseteq> y" ..
qed
qed
qed
@@ -213,13 +213,13 @@
proof (rule anti_sym)
from sup show "s \<sqsubseteq> s'"
proof (rule is_sup_least)
- from sup' show "x \<sqsubseteq> s'" ..
- from sup' show "y \<sqsubseteq> s'" ..
+ from sup' show "x \<sqsubseteq> s'" ..
+ from sup' show "y \<sqsubseteq> s'" ..
qed
from sup' show "s' \<sqsubseteq> s"
proof (rule is_sup_least)
- from sup show "x \<sqsubseteq> s" ..
- from sup show "y \<sqsubseteq> s" ..
+ from sup show "x \<sqsubseteq> s" ..
+ from sup show "y \<sqsubseteq> s" ..
qed
qed
qed
@@ -346,9 +346,9 @@
show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> x" ..
show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y"
proof -
- have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
- also have "\<dots> \<sqsubseteq> y" ..
- finally show ?thesis .
+ have "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> y \<sqinter> z" ..
+ also have "\<dots> \<sqsubseteq> y" ..
+ finally show ?thesis .
qed
qed
show "x \<sqinter> (y \<sqinter> z) \<sqsubseteq> z"
@@ -362,19 +362,19 @@
proof
show "w \<sqsubseteq> x"
proof -
- have "w \<sqsubseteq> x \<sqinter> y" by fact
- also have "\<dots> \<sqsubseteq> x" ..
- finally show ?thesis .
+ have "w \<sqsubseteq> x \<sqinter> y" by fact
+ also have "\<dots> \<sqsubseteq> x" ..
+ finally show ?thesis .
qed
show "w \<sqsubseteq> y \<sqinter> z"
proof
- show "w \<sqsubseteq> y"
- proof -
- have "w \<sqsubseteq> x \<sqinter> y" by fact
- also have "\<dots> \<sqsubseteq> y" ..
- finally show ?thesis .
- qed
- show "w \<sqsubseteq> z" by fact
+ show "w \<sqsubseteq> y"
+ proof -
+ have "w \<sqsubseteq> x \<sqinter> y" by fact
+ also have "\<dots> \<sqsubseteq> y" ..
+ finally show ?thesis .
+ qed
+ show "w \<sqsubseteq> z" by fact
qed
qed
qed
@@ -402,9 +402,9 @@
show "x \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
show "y \<sqsubseteq> x \<squnion> (y \<squnion> z)"
proof -
- have "y \<sqsubseteq> y \<squnion> z" ..
- also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
- finally show ?thesis .
+ have "y \<sqsubseteq> y \<squnion> z" ..
+ also have "... \<sqsubseteq> x \<squnion> (y \<squnion> z)" ..
+ finally show ?thesis .
qed
qed
show "z \<sqsubseteq> x \<squnion> (y \<squnion> z)"
@@ -418,19 +418,19 @@
proof
show "x \<sqsubseteq> w"
proof -
- have "x \<sqsubseteq> x \<squnion> y" ..
- also have "\<dots> \<sqsubseteq> w" by fact
- finally show ?thesis .
+ have "x \<sqsubseteq> x \<squnion> y" ..
+ also have "\<dots> \<sqsubseteq> w" by fact
+ finally show ?thesis .
qed
show "y \<squnion> z \<sqsubseteq> w"
proof
- show "y \<sqsubseteq> w"
- proof -
- have "y \<sqsubseteq> x \<squnion> y" ..
- also have "... \<sqsubseteq> w" by fact
- finally show ?thesis .
- qed
- show "z \<sqsubseteq> w" by fact
+ show "y \<sqsubseteq> w"
+ proof -
+ have "y \<sqsubseteq> x \<squnion> y" ..
+ also have "... \<sqsubseteq> w" by fact
+ finally show ?thesis .
+ qed
+ show "z \<sqsubseteq> w" by fact
qed
qed
qed
@@ -650,17 +650,17 @@
txt {* Jacobson I, p.\ 462 *}
proof -
{ assume c: "y \<sqsubseteq> x" "z \<sqsubseteq> x"
- from c have "?l = y \<squnion> z"
- by (metis c join_connection2 join_related2 meet_related2 total)
- also from c have "... = ?r" by (metis meet_related2)
- finally have "?l = ?r" . }
+ from c have "?l = y \<squnion> z"
+ by (metis c join_connection2 join_related2 meet_related2 total)
+ also from c have "... = ?r" by (metis meet_related2)
+ finally have "?l = ?r" . }
moreover
{ assume c: "x \<sqsubseteq> y \<or> x \<sqsubseteq> z"
- from c have "?l = x"
- by (metis join_connection2 join_related2 meet_connection total trans)
- also from c have "... = ?r"
- by (metis join_commute join_related2 meet_connection meet_related2 total)
- finally have "?l = ?r" . }
+ from c have "?l = x"
+ by (metis join_connection2 join_related2 meet_connection total trans)
+ also from c have "... = ?r"
+ by (metis join_commute join_related2 meet_connection meet_related2 total)
+ finally have "?l = ?r" . }
moreover note total
ultimately show ?thesis by blast
qed
--- a/doc-src/TutorialI/Overview/LNCS/Sets.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/TutorialI/Overview/LNCS/Sets.thy Sat Oct 17 14:43:18 2009 +0200
@@ -149,7 +149,7 @@
proof (rule converse_rtrancl_induct)
show "t \<in> lfp ?F"
proof (subst lfp_unfold[OF mono_ef])
- show "t \<in> ?F(lfp ?F)" using tA by blast
+ show "t \<in> ?F(lfp ?F)" using tA by blast
qed
next
fix s s'
@@ -157,7 +157,7 @@
and IH: "s' \<in> lfp ?F"
show "s \<in> lfp ?F"
proof (subst lfp_unfold[OF mono_ef])
- show "s \<in> ?F(lfp ?F)" using prems by blast
+ show "s \<in> ?F(lfp ?F)" using prems by blast
qed
qed
qed
--- a/doc-src/TutorialI/Protocol/Event.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Event.thy Sat Oct 17 14:43:18 2009 +0200
@@ -22,7 +22,7 @@
| Notes agent msg
consts
- bad :: "agent set" (*compromised agents*)
+ bad :: "agent set" -- {* compromised agents *}
knows :: "agent => event list => msg set"
@@ -43,19 +43,19 @@
knows_Cons:
"knows A (ev # evs) =
(if A = Spy then
- (case ev of
- Says A' B X => insert X (knows Spy evs)
- | Gets A' X => knows Spy evs
- | Notes A' X =>
- if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
- else
- (case ev of
- Says A' B X =>
- if A'=A then insert X (knows A evs) else knows A evs
- | Gets A' X =>
- if A'=A then insert X (knows A evs) else knows A evs
- | Notes A' X =>
- if A'=A then insert X (knows A evs) else knows A evs))"
+ (case ev of
+ Says A' B X => insert X (knows Spy evs)
+ | Gets A' X => knows Spy evs
+ | Notes A' X =>
+ if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+ else
+ (case ev of
+ Says A' B X =>
+ if A'=A then insert X (knows A evs) else knows A evs
+ | Gets A' X =>
+ if A'=A then insert X (knows A evs) else knows A evs
+ | Notes A' X =>
+ if A'=A then insert X (knows A evs) else knows A evs))"
(*
Case A=Spy on the Gets event
@@ -71,10 +71,10 @@
primrec
used_Nil: "used [] = (UN B. parts (initState B))"
used_Cons: "used (ev # evs) =
- (case ev of
- Says A B X => parts {X} \<union> used evs
- | Gets A X => used evs
- | Notes A X => parts {X} \<union> used evs)"
+ (case ev of
+ Says A B X => parts {X} \<union> used evs
+ | Gets A X => used evs
+ | Notes A X => parts {X} \<union> used evs)"
--{*The case for @{term Gets} seems anomalous, but @{term Gets} always
follows @{term Says} in real protocols. Seems difficult to change.
See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
--- a/doc-src/TutorialI/Protocol/Message.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Message.thy Sat Oct 17 14:43:18 2009 +0200
@@ -61,8 +61,8 @@
msg = Agent agent
| Nonce nat
| Key key
- | MPair msg msg
- | Crypt key msg
+ | MPair msg msg
+ | Crypt key msg
text {*
\noindent
@@ -855,8 +855,8 @@
(Fake_insert_simp_tac ss 1
THEN
IF_UNSOLVED (Blast.depth_tac
- (cs addIs [analz_insertI,
- impOfSubs analz_subset_parts]) 4 1))
+ (cs addIs [analz_insertI,
+ impOfSubs analz_subset_parts]) 4 1))
fun spy_analz_tac (cs,ss) i =
DETERM
--- a/doc-src/TutorialI/Protocol/NS_Public.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/TutorialI/Protocol/NS_Public.thy Sat Oct 17 14:43:18 2009 +0200
@@ -221,8 +221,8 @@
lemma A_trusts_NS2_lemma [rule_format]:
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
\<Longrightarrow> Crypt (pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
- Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
- Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
+ Says A B (Crypt(pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+ Says B A (Crypt(pubK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
apply (erule ns_public.induct, simp_all)
(*Fake, NS1*)
apply (blast dest: Spy_not_see_NA)+
@@ -240,8 +240,8 @@
lemma B_trusts_NS1 [rule_format]:
"evs \<in> ns_public
\<Longrightarrow> Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (knows Spy evs) \<longrightarrow>
- Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
- Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+ Nonce NA \<notin> analz (knows Spy evs) \<longrightarrow>
+ Says A B (Crypt (pubK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
apply (erule ns_public.induct, simp_all)
(*Fake*)
apply (blast intro!: analz_insertI)
--- a/doc-src/TutorialI/Protocol/Public.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/TutorialI/Protocol/Public.thy Sat Oct 17 14:43:18 2009 +0200
@@ -25,11 +25,11 @@
primrec
(*Agents know their private key and all public keys*)
initState_Server: "initState Server =
- insert (Key (priK Server)) (Key ` range pubK)"
+ insert (Key (priK Server)) (Key ` range pubK)"
initState_Friend: "initState (Friend i) =
- insert (Key (priK (Friend i))) (Key ` range pubK)"
+ insert (Key (priK (Friend i))) (Key ` range pubK)"
initState_Spy: "initState Spy =
- (Key`invKey`pubK`bad) Un (Key ` range pubK)"
+ (Key`invKey`pubK`bad) Un (Key ` range pubK)"
(*>*)
text {*
--- a/doc-src/TutorialI/Rules/Basic.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/doc-src/TutorialI/Rules/Basic.thy Sat Oct 17 14:43:18 2009 +0200
@@ -1,4 +1,3 @@
-(* ID: $Id$ *)
theory Basic imports Main begin
lemma conj_rule: "\<lbrakk> P; Q \<rbrakk> \<Longrightarrow> P \<and> (Q \<and> P)"
@@ -149,9 +148,9 @@
lemma "\<lbrakk>\<not>(P\<longrightarrow>Q); \<not>(R\<longrightarrow>Q)\<rbrakk> \<Longrightarrow> R"
apply (erule_tac Q="R\<longrightarrow>Q" in contrapos_np)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (intro impI)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
by (erule notE)
text {*
@@ -161,11 +160,11 @@
lemma "(P \<or> Q) \<and> R \<Longrightarrow> P \<or> Q \<and> R"
apply (intro disjCI conjI)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (elim conjE disjE)
apply assumption
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
by (erule contrapos_np, rule conjI)
text{*
@@ -241,18 +240,18 @@
text{*rename_tac*}
lemma "x < y \<Longrightarrow> \<forall>x y. P x (f y)"
apply (intro allI)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rename_tac v w)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
oops
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
apply (frule spec)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (drule mp, assumption)
apply (drule spec)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
by (drule mp)
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (f x); P a\<rbrakk> \<Longrightarrow> P(f (f a))"
@@ -276,11 +275,11 @@
lemma "\<lbrakk>\<forall>x. P x \<longrightarrow> P (h x); P a\<rbrakk> \<Longrightarrow> P(h (h a))"
apply (frule spec)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (drule mp, assumption)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (drule_tac x = "h a" in spec)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
by (drule mp)
text {*
@@ -290,15 +289,15 @@
lemma mult_dvd_mono: "\<lbrakk>i dvd m; j dvd n\<rbrakk> \<Longrightarrow> i*j dvd (m*n :: nat)"
apply (simp add: dvd_def)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (erule exE)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (erule exE)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rename_tac l)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply (rule_tac x="k*l" in exI)
- --{* @{subgoals[display,indent=0,margin=65]} *}
+ --{* @{subgoals[display,indent=0,margin=65]} *}
apply simp
done
--- a/src/FOL/FOL.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/FOL/FOL.thy Sat Oct 17 14:43:18 2009 +0200
@@ -174,13 +174,13 @@
structure Blast = Blast
(
val thy = @{theory}
- type claset = Cla.claset
+ type claset = Cla.claset
val equality_name = @{const_name "op ="}
val not_name = @{const_name Not}
- val notE = @{thm notE}
- val ccontr = @{thm ccontr}
+ val notE = @{thm notE}
+ val ccontr = @{thm ccontr}
val contr_tac = Cla.contr_tac
- val dup_intr = Cla.dup_intr
+ val dup_intr = Cla.dup_intr
val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac
val rep_cs = Cla.rep_cs
val cla_modifiers = Cla.cla_modifiers
--- a/src/FOL/ex/Classical.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/FOL/ex/Classical.thy Sat Oct 17 14:43:18 2009 +0200
@@ -418,7 +418,7 @@
by fast
text{*Halting problem: Formulation of Li Dafa (AAR Newsletter 27, Oct 1994.)
- author U. Egly*}
+ author U. Egly*}
lemma "((\<exists>x. A(x) & (\<forall>y. C(y) --> (\<forall>z. D(x,y,z)))) -->
(\<exists>w. C(w) & (\<forall>y. C(y) --> (\<forall>z. D(w,y,z)))))
&
--- a/src/HOL/Algebra/Coset.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/Coset.thy Sat Oct 17 14:43:18 2009 +0200
@@ -609,7 +609,7 @@
proof (simp add: r_congruent_def sym_def, clarify)
fix x y
assume [simp]: "x \<in> carrier G" "y \<in> carrier G"
- and "inv x \<otimes> y \<in> H"
+ and "inv x \<otimes> y \<in> H"
hence "inv (inv x \<otimes> y) \<in> H" by (simp add: m_inv_closed)
thus "inv y \<otimes> x \<in> H" by (simp add: inv_mult_group)
qed
@@ -618,10 +618,10 @@
proof (simp add: r_congruent_def trans_def, clarify)
fix x y z
assume [simp]: "x \<in> carrier G" "y \<in> carrier G" "z \<in> carrier G"
- and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
+ and "inv x \<otimes> y \<in> H" and "inv y \<otimes> z \<in> H"
hence "(inv x \<otimes> y) \<otimes> (inv y \<otimes> z) \<in> H" by simp
hence "inv x \<otimes> (y \<otimes> inv y) \<otimes> z \<in> H"
- by (simp add: m_assoc del: r_inv Units_r_inv)
+ by (simp add: m_assoc del: r_inv Units_r_inv)
thus "inv x \<otimes> z \<in> H" by simp
qed
qed
--- a/src/HOL/Algebra/Divisibility.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/Divisibility.thy Sat Oct 17 14:43:18 2009 +0200
@@ -2579,33 +2579,33 @@
by force
from this obtain p'
- where "p' \<in> set (as@bs)"
- and pp': "p \<sim> p'" by auto
+ where "p' \<in> set (as@bs)"
+ and pp': "p \<sim> p'" by auto
hence "p' \<in> set as \<or> p' \<in> set bs" by simp
moreover
{
- assume p'elem: "p' \<in> set as"
- with ascarr have [simp]: "p' \<in> carrier G" by fast
-
- note pp'
- also from afac p'elem
- have "p' divides a" by (rule factors_dividesI) fact+
- finally
- have "p divides a" by simp
+ assume p'elem: "p' \<in> set as"
+ with ascarr have [simp]: "p' \<in> carrier G" by fast
+
+ note pp'
+ also from afac p'elem
+ have "p' divides a" by (rule factors_dividesI) fact+
+ finally
+ have "p divides a" by simp
}
moreover
{
- assume p'elem: "p' \<in> set bs"
- with bscarr have [simp]: "p' \<in> carrier G" by fast
-
- note pp'
- also from bfac
- have "p' divides b" by (rule factors_dividesI) fact+
- finally have "p divides b" by simp
+ assume p'elem: "p' \<in> set bs"
+ with bscarr have [simp]: "p' \<in> carrier G" by fast
+
+ note pp'
+ also from bfac
+ have "p' divides b" by (rule factors_dividesI) fact+
+ finally have "p divides b" by simp
}
ultimately
- show "p divides a \<or> p divides b" by fast
+ show "p divides a \<or> p divides b" by fast
qed
qed
qed
@@ -3176,7 +3176,7 @@
have "c = c \<otimes> \<one>" by simp
also from abrelprime[symmetric]
have "\<dots> \<sim> c \<otimes> somegcd G a b"
- by (rule assoc_subst) (simp add: mult_cong_r)+
+ by (rule assoc_subst) (simp add: mult_cong_r)+
also have "\<dots> \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by (rule gcd_mult) fact+
finally
have c: "c \<sim> somegcd G (c \<otimes> a) (c \<otimes> b)" by simp
@@ -3188,13 +3188,13 @@
have "somegcd G a (b \<otimes> c) \<sim> somegcd G a (c \<otimes> b)" by (simp add: m_comm)
also from a
have "\<dots> \<sim> somegcd G (somegcd G a (c \<otimes> a)) (c \<otimes> b)"
- by (rule assoc_subst) (simp add: gcd_cong_l)+
+ by (rule assoc_subst) (simp add: gcd_cong_l)+
also from gcd_assoc
have "\<dots> \<sim> somegcd G a (somegcd G (c \<otimes> a) (c \<otimes> b))"
by (rule assoc_subst) simp+
also from c[symmetric]
have "\<dots> \<sim> somegcd G a c"
- by (rule assoc_subst) (simp add: gcd_cong_r)+
+ by (rule assoc_subst) (simp add: gcd_cong_r)+
also note acrelprime
finally
show "somegcd G a (b \<otimes> c) \<sim> \<one>" by simp
--- a/src/HOL/Algebra/FiniteProduct.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/FiniteProduct.thy Sat Oct 17 14:43:18 2009 +0200
@@ -423,17 +423,17 @@
then have "finprod G f A = finprod G f (insert x B)" by simp
also from insert have "... = f x \<otimes> finprod G f B"
proof (intro finprod_insert)
- show "finite B" by fact
+ show "finite B" by fact
next
- show "x ~: B" by fact
+ show "x ~: B" by fact
next
- assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
- "g \<in> insert x B \<rightarrow> carrier G"
- thus "f \<in> B -> carrier G" by fastsimp
+ assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
+ "g \<in> insert x B \<rightarrow> carrier G"
+ thus "f \<in> B -> carrier G" by fastsimp
next
- assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
- "g \<in> insert x B \<rightarrow> carrier G"
- thus "f x \<in> carrier G" by fastsimp
+ assume "x ~: B" "!!i. i \<in> insert x B \<Longrightarrow> f i = g i"
+ "g \<in> insert x B \<rightarrow> carrier G"
+ thus "f x \<in> carrier G" by fastsimp
qed
also from insert have "... = g x \<otimes> finprod G g B" by fastsimp
also from insert have "... = finprod G g (insert x B)"
--- a/src/HOL/Algebra/Group.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/Group.thy Sat Oct 17 14:43:18 2009 +0200
@@ -785,16 +785,16 @@
assume H: "H \<in> A"
with L have subgroupH: "subgroup H G" by auto
from subgroupH have groupH: "group (G (| carrier := H |))" (is "group ?H")
- by (rule subgroup_imp_group)
+ by (rule subgroup_imp_group)
from groupH have monoidH: "monoid ?H"
- by (rule group.is_monoid)
+ by (rule group.is_monoid)
from H have Int_subset: "?Int \<subseteq> H" by fastsimp
then show "le ?L ?Int H" by simp
next
fix H
assume H: "H \<in> Lower ?L A"
with L Int_subgroup show "le ?L H ?Int"
- by (fastsimp simp: Lower_def intro: Inter_greatest)
+ by (fastsimp simp: Lower_def intro: Inter_greatest)
next
show "A \<subseteq> carrier ?L" by (rule L)
next
--- a/src/HOL/Algebra/Lattice.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/Lattice.thy Sat Oct 17 14:43:18 2009 +0200
@@ -533,22 +533,22 @@
assume y: "y \<in> Upper L (insert x A)"
show "s \<sqsubseteq> y"
proof (rule least_le [OF least_s], rule Upper_memI)
- fix z
- assume z: "z \<in> {a, x}"
- then show "z \<sqsubseteq> y"
- proof
+ fix z
+ assume z: "z \<in> {a, x}"
+ then show "z \<sqsubseteq> y"
+ proof
have y': "y \<in> Upper L A"
apply (rule subsetD [where A = "Upper L (insert x A)"])
apply (rule Upper_antimono)
- apply blast
- apply (rule y)
+ apply blast
+ apply (rule y)
done
assume "z = a"
with y' least_a show ?thesis by (fast dest: least_le)
- next
- assume "z \<in> {x}" (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
+ next
+ assume "z \<in> {x}" (* FIXME "z = x"; declare specific elim rule for "insert x {}" (!?) *)
with y L show ?thesis by blast
- qed
+ qed
qed (rule Upper_closed [THEN subsetD, OF y])
next
from L show "insert x A \<subseteq> carrier L" by simp
@@ -569,9 +569,9 @@
case True
with insert show ?thesis
by simp (simp add: least_cong [OF weak_sup_of_singleton]
- sup_of_singleton_closed sup_of_singletonI)
- (* The above step is hairy; least_cong can make simp loop.
- Would want special version of simp to apply least_cong. *)
+ sup_of_singleton_closed sup_of_singletonI)
+ (* The above step is hairy; least_cong can make simp loop.
+ Would want special version of simp to apply least_cong. *)
next
case False
with insert have "least L (\<Squnion>A) (Upper L A)" by simp
@@ -774,22 +774,22 @@
assume y: "y \<in> Lower L (insert x A)"
show "y \<sqsubseteq> i"
proof (rule greatest_le [OF greatest_i], rule Lower_memI)
- fix z
- assume z: "z \<in> {a, x}"
- then show "y \<sqsubseteq> z"
- proof
+ fix z
+ assume z: "z \<in> {a, x}"
+ then show "y \<sqsubseteq> z"
+ proof
have y': "y \<in> Lower L A"
apply (rule subsetD [where A = "Lower L (insert x A)"])
apply (rule Lower_antimono)
- apply blast
- apply (rule y)
+ apply blast
+ apply (rule y)
done
assume "z = a"
with y' greatest_a show ?thesis by (fast dest: greatest_le)
- next
+ next
assume "z \<in> {x}"
with y L show ?thesis by blast
- qed
+ qed
qed (rule Lower_closed [THEN subsetD, OF y])
next
from L show "insert x A \<subseteq> carrier L" by simp
@@ -809,7 +809,7 @@
case True
with insert show ?thesis
by simp (simp add: greatest_cong [OF weak_inf_of_singleton]
- inf_of_singleton_closed inf_of_singletonI)
+ inf_of_singleton_closed inf_of_singletonI)
next
case False
from insert show ?thesis
@@ -1291,7 +1291,7 @@
proof
from B show "greatest ?L (\<Inter> B \<inter> A) (Lower ?L B)"
txt {* @{term "\<Inter> B"} is not the infimum of @{term B}:
- @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
+ @{term "\<Inter> {} = UNIV"} which is in general bigger than @{term "A"}! *}
by (fastsimp intro!: greatest_LowerI simp: Lower_def)
qed
qed
--- a/src/HOL/Algebra/UnivPoly.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/UnivPoly.thy Sat Oct 17 14:43:18 2009 +0200
@@ -349,19 +349,19 @@
fix nn assume Succ: "n = Suc nn"
have "coeff P (p \<otimes>\<^bsub>P\<^esub> \<one>\<^bsub>P\<^esub>) (Suc nn) = coeff P p (Suc nn)"
proof -
- 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
- 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>))"
- 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
- also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
- proof -
- have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
- 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
- unfolding Pi_def by simp
- also have "\<dots> = \<zero>" by simp
- finally show ?thesis using r_zero R by simp
- qed
- also have "\<dots> = coeff P p (Suc nn)" using R by simp
- finally show ?thesis by simp
+ 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
+ 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>))"
+ 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
+ also have "\<dots> = coeff P p (Suc nn) \<otimes> (if Suc nn \<le> Suc nn then \<one> else \<zero>)"
+ proof -
+ have "(\<Oplus>i\<in>{..nn}. coeff P p i \<otimes> (if Suc nn \<le> i then \<one> else \<zero>)) = (\<Oplus>i\<in>{..nn}. \<zero>)"
+ 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
+ unfolding Pi_def by simp
+ also have "\<dots> = \<zero>" by simp
+ finally show ?thesis using r_zero R by simp
+ qed
+ also have "\<dots> = coeff P p (Suc nn)" using R by simp
+ finally show ?thesis by simp
qed
then show ?thesis using Succ by simp
}
@@ -627,11 +627,11 @@
then show "monom P \<one> (Suc (Suc k)) = monom P \<one> 1 \<otimes>\<^bsub>P\<^esub> monom P \<one> (Suc k)"
proof -
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"
- unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
+ unfolding monom_one_Suc [of "Suc k"] unfolding hypo ..
note cl = monom_closed [OF R.one_closed, of 1]
note clk = monom_closed [OF R.one_closed, of k]
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"
- unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc [OF cl clk cl]] ..
+ unfolding monom_one_Suc [of k] unfolding sym [OF m_assoc [OF cl clk cl]] ..
from lhs rhs show ?thesis by simp
qed
}
@@ -670,25 +670,25 @@
case True
{
show ?thesis
- unfolding True [symmetric]
- coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"]
- coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
- 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>))"
- "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
- a_in_R b_in_R
- unfolding simp_implies_def
- using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
- unfolding Pi_def by auto
+ unfolding True [symmetric]
+ coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of "n + m"]
+ coeff_monom [OF a_in_R, of n] coeff_monom [OF b_in_R, of m]
+ 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>))"
+ "(\<lambda>i. if n = i then a \<otimes> b else \<zero>)"]
+ a_in_R b_in_R
+ unfolding simp_implies_def
+ using R.finsum_singleton [of n "{.. n + m}" "(\<lambda>i. a \<otimes> b)"]
+ unfolding Pi_def by auto
}
next
case False
{
show ?thesis
- unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
- unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
- unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
- 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>)"]
- unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
+ unfolding coeff_monom [OF R.m_closed [OF a_in_R b_in_R], of "n + m" k] apply (simp add: False)
+ unfolding coeff_mult [OF monom_closed [OF a_in_R, of n] monom_closed [OF b_in_R, of m], of k]
+ unfolding coeff_monom [OF a_in_R, of n] unfolding coeff_monom [OF b_in_R, of m] using False
+ 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>)"]
+ unfolding Pi_def simp_implies_def using a_in_R b_in_R by force
}
qed
qed (simp_all add: a_in_R b_in_R)
@@ -1517,7 +1517,7 @@
then have max_sl: "max (deg R p) (deg R q) < m" by simp
then have "deg R (p \<oplus>\<^bsub>P\<^esub> q) < m" using deg_add [OF p_in_P q_in_P] by arith
with deg_R_p deg_R_q show ?thesis using coeff_add [OF p_in_P q_in_P, of m]
- using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp
+ using deg_aboveD [of "p \<oplus>\<^bsub>P\<^esub> q" m] using p_in_P q_in_P by simp
qed
qed (simp add: p_in_P q_in_P)
moreover have deg_ne: "deg R (p \<oplus>\<^bsub>P\<^esub> q) \<noteq> deg R r"
@@ -1582,114 +1582,114 @@
(*JE: we now apply the induction hypothesis with some additional facts required*)
from f_in_P deg_g_le_deg_f show ?thesis
proof (induct n \<equiv> "deg R f" arbitrary: "f" rule: nat_less_induct)
- fix n f
- assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
+ fix n f
+ assume hypo: "\<forall>m<n. \<forall>x. x \<in> carrier P \<longrightarrow>
deg R g \<le> deg R x \<longrightarrow>
- m = deg R x \<longrightarrow>
- (\<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))"
- and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
- and deg_g_le_deg_f: "deg R g \<le> deg R f"
- 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)"
- and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
- 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)"
- proof -
- (*JE: we first extablish the existence of a triple satisfying the previous equation.
- Then we will have to prove the second part of the predicate.*)
- 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"
- using minus_add
- 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"]]
- using r_neg by auto
- show ?thesis
- proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
- (*JE: if the degree of the remainder satisfies the statement property we are done*)
- case True
- {
- show ?thesis
- proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
- 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
- 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
- qed (simp_all)
- }
- next
- case False note n_deg_r_l_deg_g = False
- {
- (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
- show ?thesis
- proof (cases "deg R f = 0")
- (*JE: the solutions are different if the degree of f is zero or not*)
- case True
- {
- have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
- have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
- unfolding deg_g apply simp
- unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
- using deg_zero_impl_monom [OF g_in_P deg_g] by simp
- then show ?thesis using f_in_P by blast
- }
- next
- case False note deg_f_nzero = False
- {
- (*JE: now it only remains the case where the induction hypothesis can be used.*)
- (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
- have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
- proof -
- have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
- also have "\<dots> < deg R f"
- proof (rule deg_lcoeff_cancel)
- show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
- using deg_smult_ring [of "lcoeff g" f] using prem
- using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
- show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
- using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
- by simp
- 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)"
- 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"]
- unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
- using R.finsum_cong' [of "{..deg R f}" "{..deg R f}"
- "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))"
- "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
- using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
- unfolding Pi_def using deg_g_le_deg_f by force
- qed (simp_all add: deg_f_nzero)
- finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
- qed
- moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
- moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
- moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
- (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
- ultimately obtain q' r' k'
- 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)"
- and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
- using hypo by blast
- (*JE: we now prove that the new quotient, remainder and exponent can be used to get
- the quotient, remainder and exponent of the long division theorem*)
- show ?thesis
- proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
- 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'"
- proof -
- 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)"
- using smult_assoc1 exist by simp
- 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))"
- using UP_smult_r_distr by simp
- 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')"
- using rem_desc by simp
- 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'"
- 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'"]]
- using q'_in_carrier r'_in_carrier by simp
- 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'"
- using q'_in_carrier by (auto simp add: m_comm)
- 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'"
- using smult_assoc2 q'_in_carrier by auto
- 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'"
- using sym [OF l_distr] and q'_in_carrier by auto
- finally show ?thesis using m_comm q'_in_carrier by auto
- qed
- qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
- }
- qed
- }
- qed
- qed
+ m = deg R x \<longrightarrow>
+ (\<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))"
+ and prem: "n = deg R f" and f_in_P [simp]: "f \<in> carrier P"
+ and deg_g_le_deg_f: "deg R g \<le> deg R f"
+ 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)"
+ and ?q = "monom P (lcoeff f) (deg R f - deg R g)"
+ 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)"
+ proof -
+ (*JE: we first extablish the existence of a triple satisfying the previous equation.
+ Then we will have to prove the second part of the predicate.*)
+ 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"
+ using minus_add
+ 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"]]
+ using r_neg by auto
+ show ?thesis
+ proof (cases "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < deg R g")
+ (*JE: if the degree of the remainder satisfies the statement property we are done*)
+ case True
+ {
+ show ?thesis
+ proof (rule exI3 [of _ ?q "\<ominus>\<^bsub>P\<^esub> ?r" ?k], intro conjI)
+ 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
+ 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
+ qed (simp_all)
+ }
+ next
+ case False note n_deg_r_l_deg_g = False
+ {
+ (*JE: otherwise, we verify the conditions of the induction hypothesis.*)
+ show ?thesis
+ proof (cases "deg R f = 0")
+ (*JE: the solutions are different if the degree of f is zero or not*)
+ case True
+ {
+ have deg_g: "deg R g = 0" using True using deg_g_le_deg_f by simp
+ have "lcoeff g (^) (1::nat) \<odot>\<^bsub>P\<^esub> f = g \<otimes>\<^bsub>P\<^esub> f \<oplus>\<^bsub>P\<^esub> \<zero>\<^bsub>P\<^esub>"
+ unfolding deg_g apply simp
+ unfolding sym [OF monom_mult_is_smult [OF coeff_closed [OF g_in_P, of 0] f_in_P]]
+ using deg_zero_impl_monom [OF g_in_P deg_g] by simp
+ then show ?thesis using f_in_P by blast
+ }
+ next
+ case False note deg_f_nzero = False
+ {
+ (*JE: now it only remains the case where the induction hypothesis can be used.*)
+ (*JE: we first prove that the degree of the remainder is smaller than the one of f*)
+ have deg_remainder_l_f: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n"
+ proof -
+ have "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = deg R ?r" using deg_uminus [of ?r] by simp
+ also have "\<dots> < deg R f"
+ proof (rule deg_lcoeff_cancel)
+ show "deg R (\<ominus>\<^bsub>P\<^esub> (lcoeff g \<odot>\<^bsub>P\<^esub> f)) \<le> deg R f"
+ using deg_smult_ring [of "lcoeff g" f] using prem
+ using lcoeff_nonzero2 [OF g_in_P g_not_zero] by simp
+ show "deg R (g \<otimes>\<^bsub>P\<^esub> ?q) \<le> deg R f"
+ using monom_deg_mult [OF _ g_in_P, of f "lcoeff f"] and deg_g_le_deg_f
+ by simp
+ 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)"
+ 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"]
+ unfolding coeff_monom [OF lcoeff_closed [OF f_in_P], of "(deg R f - deg R g)"]
+ using R.finsum_cong' [of "{..deg R f}" "{..deg R f}"
+ "(\<lambda>i. coeff P g i \<otimes> (if deg R f - deg R g = deg R f - i then lcoeff f else \<zero>))"
+ "(\<lambda>i. if deg R g = i then coeff P g i \<otimes> lcoeff f else \<zero>)"]
+ using R.finsum_singleton [of "deg R g" "{.. deg R f}" "(\<lambda>i. coeff P g i \<otimes> lcoeff f)"]
+ unfolding Pi_def using deg_g_le_deg_f by force
+ qed (simp_all add: deg_f_nzero)
+ finally show "deg R (\<ominus>\<^bsub>P\<^esub> ?r) < n" unfolding prem .
+ qed
+ moreover have "\<ominus>\<^bsub>P\<^esub> ?r \<in> carrier P" by simp
+ moreover obtain m where deg_rem_eq_m: "deg R (\<ominus>\<^bsub>P\<^esub> ?r) = m" by auto
+ moreover have "deg R g \<le> deg R (\<ominus>\<^bsub>P\<^esub> ?r)" using n_deg_r_l_deg_g by simp
+ (*JE: now, by applying the induction hypothesis, we obtain new quotient, remainder and exponent.*)
+ ultimately obtain q' r' k'
+ 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)"
+ and q'_in_carrier: "q' \<in> carrier P" and r'_in_carrier: "r' \<in> carrier P"
+ using hypo by blast
+ (*JE: we now prove that the new quotient, remainder and exponent can be used to get
+ the quotient, remainder and exponent of the long division theorem*)
+ show ?thesis
+ proof (rule exI3 [of _ "((lcoeff g (^) k') \<odot>\<^bsub>P\<^esub> ?q \<oplus>\<^bsub>P\<^esub> q')" r' "Suc k'"], intro conjI)
+ 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'"
+ proof -
+ 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)"
+ using smult_assoc1 exist by simp
+ 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))"
+ using UP_smult_r_distr by simp
+ 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')"
+ using rem_desc by simp
+ 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'"
+ 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'"]]
+ using q'_in_carrier r'_in_carrier by simp
+ 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'"
+ using q'_in_carrier by (auto simp add: m_comm)
+ 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'"
+ using smult_assoc2 q'_in_carrier by auto
+ 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'"
+ using sym [OF l_distr] and q'_in_carrier by auto
+ finally show ?thesis using m_comm q'_in_carrier by auto
+ qed
+ qed (simp_all add: rem_deg q'_in_carrier r'_in_carrier)
+ }
+ qed
+ }
+ qed
+ qed
qed
}
qed
--- a/src/HOL/Algebra/poly/UnivPoly2.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Algebra/poly/UnivPoly2.thy Sat Oct 17 14:43:18 2009 +0200
@@ -100,7 +100,7 @@
begin
definition
- up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
+ up_add_def: "p + q = Abs_UP (%n. Rep_UP p n + Rep_UP q n)"
instance ..
@@ -133,7 +133,7 @@
definition
up_mult_def: "p * q = Abs_UP (%n::nat. setsum
- (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
+ (%i. Rep_UP p i * Rep_UP q (n-i)) {..n})"
instance ..
@@ -201,18 +201,18 @@
have "(%i. f i + g i) : UP"
proof -
from fup obtain n where boundn: "bound n f"
- by (unfold UP_def) fast
+ by (unfold UP_def) fast
from gup obtain m where boundm: "bound m g"
- by (unfold UP_def) fast
+ by (unfold UP_def) fast
have "bound (max n m) (%i. (f i + g i))"
proof
- fix i
- assume "max n m < i"
- with boundn and boundm show "f i + g i = 0"
+ fix i
+ assume "max n m < i"
+ with boundn and boundm show "f i + g i = 0"
by (fastsimp simp add: algebra_simps)
qed
then show "(%i. (f i + g i)) : UP"
- by (unfold UP_def) fast
+ by (unfold UP_def) fast
qed
}
then show ?thesis
@@ -228,30 +228,30 @@
have "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
proof -
from fup obtain n where "bound n f"
- by (unfold UP_def) fast
+ by (unfold UP_def) fast
from gup obtain m where "bound m g"
- by (unfold UP_def) fast
+ by (unfold UP_def) fast
have "bound (n + m) (%n. setsum (%i. f i * g (n-i)) {..n})"
proof
- fix k
- assume bound: "n + m < k"
- {
- fix i
- have "f i * g (k-i) = 0"
- proof cases
- assume "n < i"
- with `bound n f` show ?thesis by (auto simp add: algebra_simps)
- next
- assume "~ (n < i)"
- with bound have "m < k-i" by arith
- with `bound m g` show ?thesis by (auto simp add: algebra_simps)
- qed
- }
- then show "setsum (%i. f i * g (k-i)) {..k} = 0"
- by (simp add: algebra_simps)
+ fix k
+ assume bound: "n + m < k"
+ {
+ fix i
+ have "f i * g (k-i) = 0"
+ proof cases
+ assume "n < i"
+ with `bound n f` show ?thesis by (auto simp add: algebra_simps)
+ next
+ assume "~ (n < i)"
+ with bound have "m < k-i" by arith
+ with `bound m g` show ?thesis by (auto simp add: algebra_simps)
+ qed
+ }
+ then show "setsum (%i. f i * g (k-i)) {..k} = 0"
+ by (simp add: algebra_simps)
qed
then show "(%n. setsum (%i. f i * g (n-i)) {..n}) : UP"
- by (unfold UP_def) fast
+ by (unfold UP_def) fast
qed
}
then show ?thesis
@@ -290,17 +290,17 @@
{
fix k and a b c :: "nat=>'a::ring"
have "k <= n ==>
- setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} =
- setsum (%j. a j * setsum (%i. b i * c (n-j-i)) {..k-j}) {..k}"
- (is "_ ==> ?eq k")
+ setsum (%j. setsum (%i. a i * b (j-i)) {..j} * c (n-j)) {..k} =
+ setsum (%j. a j * setsum (%i. b i * c (n-j-i)) {..k-j}) {..k}"
+ (is "_ ==> ?eq k")
proof (induct k)
- case 0 show ?case by simp
+ case 0 show ?case by simp
next
- case (Suc k)
- then have "k <= n" by arith
- then have "?eq k" by (rule Suc)
- then show ?case
- by (simp add: Suc_diff_le natsum_ldistr)
+ case (Suc k)
+ then have "k <= n" by arith
+ then have "?eq k" by (rule Suc)
+ then show ?case
+ by (simp add: Suc_diff_le natsum_ldistr)
qed
}
then show "coeff ((p * q) * r) n = coeff (p * (q * r)) n"
@@ -326,13 +326,13 @@
fix k
fix a b :: "nat=>'a::ring"
have "k <= n ==>
- setsum (%i. a i * b (n-i)) {..k} =
- setsum (%i. a (k-i) * b (i+n-k)) {..k}"
- (is "_ ==> ?eq k")
+ setsum (%i. a i * b (n-i)) {..k} =
+ setsum (%i. a (k-i) * b (i+n-k)) {..k}"
+ (is "_ ==> ?eq k")
proof (induct k)
- case 0 show ?case by simp
+ case 0 show ?case by simp
next
- case (Suc k) then show ?case by (subst natsum_Suc2) simp
+ case (Suc k) then show ?case by (subst natsum_Suc2) simp
qed
}
then show "coeff (p * q) n = coeff (q * p) n"
--- a/src/HOL/Auth/CertifiedEmail.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/CertifiedEmail.thy Sat Oct 17 14:43:18 2009 +0200
@@ -43,7 +43,7 @@
| FakeSSL: --{*The Spy may open SSL sessions with TTP, who is the only agent
equipped with the necessary credentials to serve as an SSL server.*}
- "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
+ "[| evsfssl \<in> certified_mail; X \<in> synth(analz(spies evsfssl))|]
==> Notes TTP {|Agent Spy, Agent TTP, X|} # evsfssl \<in> certified_mail"
| CM1: --{*The sender approaches the recipient. The message is a number.*}
@@ -54,14 +54,14 @@
hs = Hash{|Number cleartext, Nonce q, response S R q, Crypt K (Number m)|};
S2TTP = Crypt(pubEK TTP) {|Agent S, Number BothAuth, Key K, Agent R, hs|}|]
==> Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number BothAuth,
- Number cleartext, Nonce q, S2TTP|} # evs1
- \<in> certified_mail"
+ Number cleartext, Nonce q, S2TTP|} # evs1
+ \<in> certified_mail"
| CM2: --{*The recipient records @{term S2TTP} while transmitting it and her
password to @{term TTP} over an SSL channel.*}
"[|evs2 \<in> certified_mail;
Gets R {|Agent S, Agent TTP, em, Number BothAuth, Number cleartext,
- Nonce q, S2TTP|} \<in> set evs2;
+ Nonce q, S2TTP|} \<in> set evs2;
TTP \<noteq> R;
hr = Hash {|Number cleartext, Nonce q, response S R q, em|} |]
==>
@@ -289,14 +289,14 @@
"evs \<in> certified_mail ==>
Key K \<notin> analz (spies evs) -->
(\<forall>AO. Crypt (pubEK TTP)
- {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
+ {|Agent S, Number AO, Key K, Agent R, hs|} \<in> used evs -->
(\<exists>m ctxt q.
hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
- Says S R
- {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
- Number ctxt, Nonce q,
- Crypt (pubEK TTP)
- {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))"
+ Says S R
+ {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
+ Number ctxt, Nonce q,
+ Crypt (pubEK TTP)
+ {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs))"
apply (erule certified_mail.induct, analz_mono_contra)
apply (drule_tac [5] CM2_S2TTP_parts_knows_Spy, simp)
apply (simp add: used_Nil Crypt_notin_initState, simp_all)
@@ -322,11 +322,11 @@
evs \<in> certified_mail|]
==> \<exists>m ctxt q.
hs = Hash{|Number ctxt, Nonce q, response S R q, Crypt K (Number m)|} &
- Says S R
- {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
- Number ctxt, Nonce q,
- Crypt (pubEK TTP)
- {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs"
+ Says S R
+ {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
+ Number ctxt, Nonce q,
+ Crypt (pubEK TTP)
+ {|Agent S, Number AO, Key K, Agent R, hs |}|} \<in> set evs"
by (blast intro: S2TTP_sender_lemma)
@@ -401,7 +401,7 @@
Number cleartext, Nonce q, S2TTP|} \<in> set evs;
S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
Key K \<in> analz (spies evs);
- evs \<in> certified_mail;
+ evs \<in> certified_mail;
S\<noteq>Spy|]
==> R \<in> bad & Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
apply (erule rev_mp)
@@ -428,7 +428,7 @@
"[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
Number cleartext, Nonce q, S2TTP|} \<in> set evs;
S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
- evs \<in> certified_mail;
+ evs \<in> certified_mail;
S\<noteq>Spy; R \<notin> bad|]
==> Key K \<notin> analz(spies evs)"
by (blast dest: S_fairness_bad_R)
@@ -438,10 +438,10 @@
until @{term S} has access to the return receipt.*}
theorem S_guarantee:
"[|Says S R {|Agent S, Agent TTP, Crypt K (Number m), Number AO,
- Number cleartext, Nonce q, S2TTP|} \<in> set evs;
- S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
- Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
- S\<noteq>Spy; evs \<in> certified_mail|]
+ Number cleartext, Nonce q, S2TTP|} \<in> set evs;
+ S2TTP = Crypt (pubEK TTP) {|Agent S, Number AO, Key K, Agent R, hs|};
+ Notes R {|Agent TTP, Agent R, Key K, hs|} \<in> set evs;
+ S\<noteq>Spy; evs \<in> certified_mail|]
==> Gets S (Crypt (priSK TTP) S2TTP) \<in> set evs"
apply (erule rev_mp)
apply (erule ssubst)
--- a/src/HOL/Auth/Event.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Event.thy Sat Oct 17 14:43:18 2009 +0200
@@ -22,7 +22,7 @@
| Notes agent msg
consts
- bad :: "agent set" (*compromised agents*)
+ bad :: "agent set" -- {* compromised agents *}
knows :: "agent => event list => msg set"
@@ -43,19 +43,19 @@
knows_Cons:
"knows A (ev # evs) =
(if A = Spy then
- (case ev of
- Says A' B X => insert X (knows Spy evs)
- | Gets A' X => knows Spy evs
- | Notes A' X =>
- if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
- else
- (case ev of
- Says A' B X =>
- if A'=A then insert X (knows A evs) else knows A evs
- | Gets A' X =>
- if A'=A then insert X (knows A evs) else knows A evs
- | Notes A' X =>
- if A'=A then insert X (knows A evs) else knows A evs))"
+ (case ev of
+ Says A' B X => insert X (knows Spy evs)
+ | Gets A' X => knows Spy evs
+ | Notes A' X =>
+ if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs)
+ else
+ (case ev of
+ Says A' B X =>
+ if A'=A then insert X (knows A evs) else knows A evs
+ | Gets A' X =>
+ if A'=A then insert X (knows A evs) else knows A evs
+ | Notes A' X =>
+ if A'=A then insert X (knows A evs) else knows A evs))"
(*
Case A=Spy on the Gets event
@@ -71,10 +71,10 @@
primrec
used_Nil: "used [] = (UN B. parts (initState B))"
used_Cons: "used (ev # evs) =
- (case ev of
- Says A B X => parts {X} \<union> used evs
- | Gets A X => used evs
- | Notes A X => parts {X} \<union> used evs)"
+ (case ev of
+ Says A B X => parts {X} \<union> used evs
+ | Gets A X => used evs
+ | Notes A X => parts {X} \<union> used evs)"
--{*The case for @{term Gets} seems anomalous, but @{term Gets} always
follows @{term Says} in real protocols. Seems difficult to change.
See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
--- a/src/HOL/Auth/KerberosIV.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/KerberosIV.thy Sat Oct 17 14:43:18 2009 +0200
@@ -170,18 +170,18 @@
B \<noteq> Tgs; authK \<in> symKeys;
Says A' Tgs \<lbrace>
(Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
- Number Ta\<rbrace>),
+ Number Ta\<rbrace>),
(Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
- \<in> set evs4;
+ \<in> set evs4;
\<not> expiredAK Ta evs4;
\<not> expiredA T2 evs4;
servKlife + (CT evs4) <= authKlife + Ta
\<rbrakk>
\<Longrightarrow> Says Tgs A
(Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
- Number (CT evs4)\<rbrace> \<rbrace>)
- # evs4 \<in> kerbIV"
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+ Number (CT evs4)\<rbrace> \<rbrace>)
+ # evs4 \<in> kerbIV"
(* Tgs creates a new session key per each request for a service, without
checking if there is still a fresh one for that service.
The cipher under Tgs' key is the authTicket, the cipher under B's key
@@ -196,14 +196,14 @@
| K5: "\<lbrakk> evs5 \<in> kerbIV; authK \<in> symKeys; servK \<in> symKeys;
Says A Tgs
\<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
- Agent B\<rbrace>
+ Agent B\<rbrace>
\<in> set evs5;
Says Tgs' A
(Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
\<in> set evs5;
valid Ts wrt T2 \<rbrakk>
\<Longrightarrow> Says A B \<lbrace>servTicket,
- Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+ Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
# evs5 \<in> kerbIV"
(* Checks similar to those in K3. *)
@@ -609,7 +609,7 @@
evs \<in> kerbIV \<rbrakk>
\<Longrightarrow> servK \<notin> range shrK &
(\<exists>A. servTicket =
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
| servTicket \<in> analz (spies evs)"
apply (frule Says_imp_spies [THEN analz.Inj], auto)
apply (force dest!: servTicket_form)
@@ -1336,15 +1336,15 @@
lemma Confidentiality_lemma [rule_format]:
"\<lbrakk> Says Tgs A
- (Crypt authK
- \<lbrace>Key servK, Agent B, Number Ts,
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
- \<in> set evs;
- Key authK \<notin> analz (spies evs);
+ (Crypt authK
+ \<lbrace>Key servK, Agent B, Number Ts,
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+ \<in> set evs;
+ Key authK \<notin> analz (spies evs);
servK \<in> symKeys;
- A \<notin> bad; B \<notin> bad; evs \<in> kerbIV \<rbrakk>
+ A \<notin> bad; B \<notin> bad; evs \<in> kerbIV \<rbrakk>
\<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
- expiredSK Ts evs"
+ expiredSK Ts evs"
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule kerbIV.induct)
--- a/src/HOL/Auth/KerberosIV_Gets.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/KerberosIV_Gets.thy Sat Oct 17 14:43:18 2009 +0200
@@ -161,18 +161,18 @@
B \<noteq> Tgs; authK \<in> symKeys;
Gets Tgs \<lbrace>
(Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
- Number Ta\<rbrace>),
+ Number Ta\<rbrace>),
(Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
- \<in> set evs4;
+ \<in> set evs4;
\<not> expiredAK Ta evs4;
\<not> expiredA T2 evs4;
servKlife + (CT evs4) <= authKlife + Ta
\<rbrakk>
\<Longrightarrow> Says Tgs A
(Crypt authK \<lbrace>Key servK, Agent B, Number (CT evs4),
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
- Number (CT evs4)\<rbrace> \<rbrace>)
- # evs4 \<in> kerbIV_gets"
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK,
+ Number (CT evs4)\<rbrace> \<rbrace>)
+ # evs4 \<in> kerbIV_gets"
(* Tgs creates a new session key per each request for a service, without
checking if there is still a fresh one for that service.
The cipher under Tgs' key is the authTicket, the cipher under B's key
@@ -187,14 +187,14 @@
| K5: "\<lbrakk> evs5 \<in> kerbIV_gets; authK \<in> symKeys; servK \<in> symKeys;
Says A Tgs
\<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
- Agent B\<rbrace>
+ Agent B\<rbrace>
\<in> set evs5;
Gets A
(Crypt authK \<lbrace>Key servK, Agent B, Number Ts, servTicket\<rbrace>)
\<in> set evs5;
valid Ts wrt T2 \<rbrakk>
\<Longrightarrow> Says A B \<lbrace>servTicket,
- Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+ Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
# evs5 \<in> kerbIV_gets"
(* Checks similar to those in K3. *)
@@ -495,7 +495,7 @@
evs \<in> kerbIV_gets \<rbrakk>
\<Longrightarrow> servK \<notin> range shrK &
(\<exists>A. servTicket =
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Ts\<rbrace>)
| servTicket \<in> analz (spies evs)"
apply (frule Gets_imp_knows_Spy [THEN analz.Inj], auto)
apply (force dest!: servTicket_form)
@@ -1231,15 +1231,15 @@
lemma Confidentiality_lemma [rule_format]:
"\<lbrakk> Says Tgs A
- (Crypt authK
- \<lbrace>Key servK, Agent B, Number Ts,
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
- \<in> set evs;
- Key authK \<notin> analz (spies evs);
+ (Crypt authK
+ \<lbrace>Key servK, Agent B, Number Ts,
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>)
+ \<in> set evs;
+ Key authK \<notin> analz (spies evs);
servK \<in> symKeys;
- A \<notin> bad; B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
+ A \<notin> bad; B \<notin> bad; evs \<in> kerbIV_gets \<rbrakk>
\<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
- expiredSK Ts evs"
+ expiredSK Ts evs"
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule kerbIV_gets.induct)
--- a/src/HOL/Auth/KerberosV.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/KerberosV.thy Sat Oct 17 14:43:18 2009 +0200
@@ -137,9 +137,9 @@
B \<noteq> Tgs; authK \<in> symKeys;
Says A' Tgs \<lbrace>
(Crypt (shrK Tgs) \<lbrace>Agent A, Agent Tgs, Key authK,
- Number Ta\<rbrace>),
+ Number Ta\<rbrace>),
(Crypt authK \<lbrace>Agent A, Number T2\<rbrace>), Agent B\<rbrace>
- \<in> set evs4;
+ \<in> set evs4;
\<not> expiredAK Ta evs4;
\<not> expiredA T2 evs4;
servKlife + (CT evs4) <= authKlife + Ta
@@ -155,14 +155,14 @@
A \<noteq> Kas; A \<noteq> Tgs;
Says A Tgs
\<lbrace>authTicket, Crypt authK \<lbrace>Agent A, Number T2\<rbrace>,
- Agent B\<rbrace>
+ Agent B\<rbrace>
\<in> set evs5;
Says Tgs' A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
servTicket\<rbrace>
\<in> set evs5;
valid Ts wrt T2 \<rbrakk>
\<Longrightarrow> Says A B \<lbrace>servTicket,
- Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
+ Crypt servK \<lbrace>Agent A, Number (CT evs5)\<rbrace> \<rbrace>
# evs5 \<in> kerbV"
| KV6: "\<lbrakk> evs6 \<in> kerbV; B \<noteq> Kas; B \<noteq> Tgs;
@@ -1081,14 +1081,14 @@
lemma Confidentiality_lemma [rule_format]:
"\<lbrakk> Says Tgs A
- \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
- Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
- \<in> set evs;
- Key authK \<notin> analz (spies evs);
+ \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, Number Ts\<rbrace>,
+ Crypt (shrK B) \<lbrace>Agent A, Agent B, Key servK, Number Ts\<rbrace>\<rbrace>
+ \<in> set evs;
+ Key authK \<notin> analz (spies evs);
servK \<in> symKeys;
- A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
+ A \<notin> bad; B \<notin> bad; evs \<in> kerbV \<rbrakk>
\<Longrightarrow> Key servK \<in> analz (spies evs) \<longrightarrow>
- expiredSK Ts evs"
+ expiredSK Ts evs"
apply (erule rev_mp)
apply (erule rev_mp)
apply (erule kerbV.induct)
--- a/src/HOL/Auth/Kerberos_BAN.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Kerberos_BAN.thy Sat Oct 17 14:43:18 2009 +0200
@@ -77,45 +77,45 @@
Nil: "[] \<in> bankerberos"
| Fake: "\<lbrakk> evsf \<in> bankerberos; X \<in> synth (analz (spies evsf)) \<rbrakk>
- \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
+ \<Longrightarrow> Says Spy B X # evsf \<in> bankerberos"
| BK1: "\<lbrakk> evs1 \<in> bankerberos \<rbrakk>
- \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
- \<in> bankerberos"
+ \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+ \<in> bankerberos"
| BK2: "\<lbrakk> evs2 \<in> bankerberos; Key K \<notin> used evs2; K \<in> symKeys;
- Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
- \<Longrightarrow> Says Server A
- (Crypt (shrK A)
- \<lbrace>Number (CT evs2), Agent B, Key K,
- (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
- # evs2 \<in> bankerberos"
+ Says A' Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+ \<Longrightarrow> Says Server A
+ (Crypt (shrK A)
+ \<lbrace>Number (CT evs2), Agent B, Key K,
+ (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+ # evs2 \<in> bankerberos"
| BK3: "\<lbrakk> evs3 \<in> bankerberos;
- Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
- \<in> set evs3;
- Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
- \<not> expiredK Tk evs3 \<rbrakk>
- \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
- # evs3 \<in> bankerberos"
+ Says S A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+ \<in> set evs3;
+ Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+ \<not> expiredK Tk evs3 \<rbrakk>
+ \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+ # evs3 \<in> bankerberos"
| BK4: "\<lbrakk> evs4 \<in> bankerberos;
- Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
- (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
- \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
- \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
- \<in> bankerberos"
+ Says A' B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+ (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+ \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
+ \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+ \<in> bankerberos"
- (*Old session keys may become compromised*)
+ (*Old session keys may become compromised*)
| Oops: "\<lbrakk> evso \<in> bankerberos;
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
- \<in> set evso;
- expiredK Tk evso \<rbrakk>
- \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
+ \<in> set evso;
+ expiredK Tk evso \<rbrakk>
+ \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerberos"
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Sat Oct 17 14:43:18 2009 +0200
@@ -69,47 +69,47 @@
Nil: "[] \<in> bankerb_gets"
| Fake: "\<lbrakk> evsf \<in> bankerb_gets; X \<in> synth (analz (knows Spy evsf)) \<rbrakk>
- \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
+ \<Longrightarrow> Says Spy B X # evsf \<in> bankerb_gets"
| Reception: "\<lbrakk> evsr\<in> bankerb_gets; Says A B X \<in> set evsr \<rbrakk>
\<Longrightarrow> Gets B X # evsr \<in> bankerb_gets"
| BK1: "\<lbrakk> evs1 \<in> bankerb_gets \<rbrakk>
- \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
- \<in> bankerb_gets"
+ \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B\<rbrace> # evs1
+ \<in> bankerb_gets"
| BK2: "\<lbrakk> evs2 \<in> bankerb_gets; Key K \<notin> used evs2; K \<in> symKeys;
- Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
- \<Longrightarrow> Says Server A
- (Crypt (shrK A)
- \<lbrace>Number (CT evs2), Agent B, Key K,
- (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
- # evs2 \<in> bankerb_gets"
+ Gets Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs2 \<rbrakk>
+ \<Longrightarrow> Says Server A
+ (Crypt (shrK A)
+ \<lbrace>Number (CT evs2), Agent B, Key K,
+ (Crypt (shrK B) \<lbrace>Number (CT evs2), Agent A, Key K\<rbrace>)\<rbrace>)
+ # evs2 \<in> bankerb_gets"
| BK3: "\<lbrakk> evs3 \<in> bankerb_gets;
- Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
- \<in> set evs3;
- Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
- \<not> expiredK Tk evs3 \<rbrakk>
- \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
- # evs3 \<in> bankerb_gets"
+ Gets A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
+ \<in> set evs3;
+ Says A Server \<lbrace>Agent A, Agent B\<rbrace> \<in> set evs3;
+ \<not> expiredK Tk evs3 \<rbrakk>
+ \<Longrightarrow> Says A B \<lbrace>Ticket, Crypt K \<lbrace>Agent A, Number (CT evs3)\<rbrace> \<rbrace>
+ # evs3 \<in> bankerb_gets"
| BK4: "\<lbrakk> evs4 \<in> bankerb_gets;
- Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
- (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
- \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
- \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
- \<in> bankerb_gets"
+ Gets B \<lbrace>(Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>),
+ (Crypt K \<lbrace>Agent A, Number Ta\<rbrace>) \<rbrace>: set evs4;
+ \<not> expiredK Tk evs4; \<not> expiredA Ta evs4 \<rbrakk>
+ \<Longrightarrow> Says B A (Crypt K (Number Ta)) # evs4
+ \<in> bankerb_gets"
- (*Old session keys may become compromised*)
+ (*Old session keys may become compromised*)
| Oops: "\<lbrakk> evso \<in> bankerb_gets;
Says Server A (Crypt (shrK A) \<lbrace>Number Tk, Agent B, Key K, Ticket\<rbrace>)
- \<in> set evso;
- expiredK Tk evso \<rbrakk>
- \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
+ \<in> set evso;
+ expiredK Tk evso \<rbrakk>
+ \<Longrightarrow> Notes Spy \<lbrace>Number Tk, Key K\<rbrace> # evso \<in> bankerb_gets"
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -359,7 +359,7 @@
"\<lbrakk> Says B A (Crypt K (Number Ta)) \<in> set evs;
B \<notin> bad; evs \<in> bankerb_gets \<rbrakk>
\<Longrightarrow> \<exists> Tk. Gets B \<lbrace>Crypt (shrK B) \<lbrace>Number Tk, Agent A, Key K\<rbrace>,
- Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
+ Crypt K \<lbrace>Agent A, Number Ta\<rbrace>\<rbrace> \<in> set evs"
apply (erule rev_mp)
apply (erule bankerb_gets.induct)
apply auto
--- a/src/HOL/Auth/Message.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Message.thy Sat Oct 17 14:43:18 2009 +0200
@@ -40,13 +40,13 @@
agent = Server | Friend nat | Spy
datatype
- msg = Agent agent --{*Agent names*}
+ msg = Agent agent --{*Agent names*}
| Number nat --{*Ordinary integers, timestamps, ...*}
| Nonce nat --{*Unguessable nonces*}
| Key key --{*Crypto keys*}
- | Hash msg --{*Hashing*}
- | MPair msg msg --{*Compound messages*}
- | Crypt key msg --{*Encryption, public- or shared-key*}
+ | Hash msg --{*Hashing*}
+ | MPair msg msg --{*Compound messages*}
+ | Crypt key msg --{*Encryption, public- or shared-key*}
text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
@@ -873,8 +873,8 @@
(Fake_insert_simp_tac ss 1
THEN
IF_UNSOLVED (Blast.depth_tac
- (cs addIs [@{thm analz_insertI},
- impOfSubs @{thm analz_subset_parts}]) 4 1))
+ (cs addIs [@{thm analz_insertI},
+ impOfSubs @{thm analz_subset_parts}]) 4 1))
fun spy_analz_tac (cs,ss) i =
DETERM
--- a/src/HOL/Auth/NS_Public.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/NS_Public.thy Sat Oct 17 14:43:18 2009 +0200
@@ -110,8 +110,8 @@
lemma A_trusts_NS2_lemma [rule_format]:
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
\<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace> \<in> parts (spies evs) \<longrightarrow>
- Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
- Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
+ Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+ Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB, Agent B\<rbrace>) \<in> set evs"
apply (erule ns_public.induct, simp_all)
(*Fake, NS1*)
apply (blast dest: Spy_not_see_NA)+
@@ -129,8 +129,8 @@
lemma B_trusts_NS1 [rule_format]:
"evs \<in> ns_public
\<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
- Nonce NA \<notin> analz (spies evs) \<longrightarrow>
- Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+ Nonce NA \<notin> analz (spies evs) \<longrightarrow>
+ Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
apply (erule ns_public.induct, simp_all)
(*Fake*)
apply (blast intro!: analz_insertI)
--- a/src/HOL/Auth/NS_Public_Bad.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/NS_Public_Bad.thy Sat Oct 17 14:43:18 2009 +0200
@@ -116,8 +116,8 @@
lemma A_trusts_NS2_lemma [rule_format]:
"\<lbrakk>A \<notin> bad; B \<notin> bad; evs \<in> ns_public\<rbrakk>
\<Longrightarrow> Crypt (pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
- Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
- Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
+ Says A B (Crypt(pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs \<longrightarrow>
+ Says B A (Crypt(pubEK A) \<lbrace>Nonce NA, Nonce NB\<rbrace>) \<in> set evs"
apply (erule ns_public.induct)
apply (auto dest: Spy_not_see_NA unique_NA)
done
@@ -134,8 +134,8 @@
lemma B_trusts_NS1 [rule_format]:
"evs \<in> ns_public
\<Longrightarrow> Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace> \<in> parts (spies evs) \<longrightarrow>
- Nonce NA \<notin> analz (spies evs) \<longrightarrow>
- Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
+ Nonce NA \<notin> analz (spies evs) \<longrightarrow>
+ Says A B (Crypt (pubEK B) \<lbrace>Nonce NA, Agent A\<rbrace>) \<in> set evs"
apply (erule ns_public.induct, simp_all)
(*Fake*)
apply (blast intro!: analz_insertI)
--- a/src/HOL/Auth/NS_Shared.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/NS_Shared.thy Sat Oct 17 14:43:18 2009 +0200
@@ -27,61 +27,61 @@
inductive_set ns_shared :: "event list set"
where
- (*Initial trace is empty*)
+ (*Initial trace is empty*)
Nil: "[] \<in> ns_shared"
- (*The spy MAY say anything he CAN say. We do not expect him to
- invent new nonces here, but he can also use NS1. Common to
- all similar protocols.*)
+ (*The spy MAY say anything he CAN say. We do not expect him to
+ invent new nonces here, but he can also use NS1. Common to
+ all similar protocols.*)
| Fake: "\<lbrakk>evsf \<in> ns_shared; X \<in> synth (analz (spies evsf))\<rbrakk>
- \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
+ \<Longrightarrow> Says Spy B X # evsf \<in> ns_shared"
- (*Alice initiates a protocol run, requesting to talk to any B*)
+ (*Alice initiates a protocol run, requesting to talk to any B*)
| NS1: "\<lbrakk>evs1 \<in> ns_shared; Nonce NA \<notin> used evs1\<rbrakk>
- \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared"
+ \<Longrightarrow> Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> # evs1 \<in> ns_shared"
- (*Server's response to Alice's message.
- !! It may respond more than once to A's request !!
- Server doesn't know who the true sender is, hence the A' in
- the sender field.*)
+ (*Server's response to Alice's message.
+ !! It may respond more than once to A's request !!
+ Server doesn't know who the true sender is, hence the A' in
+ the sender field.*)
| NS2: "\<lbrakk>evs2 \<in> ns_shared; Key KAB \<notin> used evs2; KAB \<in> symKeys;
- Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
- \<Longrightarrow> Says Server A
- (Crypt (shrK A)
- \<lbrace>Nonce NA, Agent B, Key KAB,
- (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
- # evs2 \<in> ns_shared"
+ Says A' Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs2\<rbrakk>
+ \<Longrightarrow> Says Server A
+ (Crypt (shrK A)
+ \<lbrace>Nonce NA, Agent B, Key KAB,
+ (Crypt (shrK B) \<lbrace>Key KAB, Agent A\<rbrace>)\<rbrace>)
+ # evs2 \<in> ns_shared"
- (*We can't assume S=Server. Agent A "remembers" her nonce.
- Need A \<noteq> Server because we allow messages to self.*)
+ (*We can't assume S=Server. Agent A "remembers" her nonce.
+ Need A \<noteq> Server because we allow messages to self.*)
| NS3: "\<lbrakk>evs3 \<in> ns_shared; A \<noteq> Server;
- Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
- Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
- \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
+ Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>) \<in> set evs3;
+ Says A Server \<lbrace>Agent A, Agent B, Nonce NA\<rbrace> \<in> set evs3\<rbrakk>
+ \<Longrightarrow> Says A B X # evs3 \<in> ns_shared"
- (*Bob's nonce exchange. He does not know who the message came
- from, but responds to A because she is mentioned inside.*)
+ (*Bob's nonce exchange. He does not know who the message came
+ from, but responds to A because she is mentioned inside.*)
| NS4: "\<lbrakk>evs4 \<in> ns_shared; Nonce NB \<notin> used evs4; K \<in> symKeys;
- Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
- \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
+ Says A' B (Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>) \<in> set evs4\<rbrakk>
+ \<Longrightarrow> Says B A (Crypt K (Nonce NB)) # evs4 \<in> ns_shared"
- (*Alice responds with Nonce NB if she has seen the key before.
- Maybe should somehow check Nonce NA again.
- We do NOT send NB-1 or similar as the Spy cannot spoof such things.
- Letting the Spy add or subtract 1 lets him send all nonces.
- Instead we distinguish the messages by sending the nonce twice.*)
+ (*Alice responds with Nonce NB if she has seen the key before.
+ Maybe should somehow check Nonce NA again.
+ We do NOT send NB-1 or similar as the Spy cannot spoof such things.
+ Letting the Spy add or subtract 1 lets him send all nonces.
+ Instead we distinguish the messages by sending the nonce twice.*)
| NS5: "\<lbrakk>evs5 \<in> ns_shared; K \<in> symKeys;
- Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
- Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
- \<in> set evs5\<rbrakk>
- \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
+ Says B' A (Crypt K (Nonce NB)) \<in> set evs5;
+ Says S A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
+ \<in> set evs5\<rbrakk>
+ \<Longrightarrow> Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) # evs5 \<in> ns_shared"
- (*This message models possible leaks of session keys.
- The two Nonces identify the protocol run: the rule insists upon
- the true senders in order to make them accurate.*)
+ (*This message models possible leaks of session keys.
+ The two Nonces identify the protocol run: the rule insists upon
+ the true senders in order to make them accurate.*)
| Oops: "\<lbrakk>evso \<in> ns_shared; Says B A (Crypt K (Nonce NB)) \<in> set evso;
- Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
- \<in> set evso\<rbrakk>
- \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
+ Says Server A (Crypt (shrK A) \<lbrace>Nonce NA, Agent B, Key K, X\<rbrace>)
+ \<in> set evso\<rbrakk>
+ \<Longrightarrow> Notes Spy \<lbrace>Nonce NA, Nonce NB, Key K\<rbrace> # evso \<in> ns_shared"
declare Says_imp_knows_Spy [THEN parts.Inj, dest]
@@ -98,7 +98,7 @@
apply (intro exI bexI)
apply (rule_tac [2] ns_shared.Nil
[THEN ns_shared.NS1, THEN ns_shared.NS2, THEN ns_shared.NS3,
- THEN ns_shared.NS4, THEN ns_shared.NS5])
+ THEN ns_shared.NS4, THEN ns_shared.NS5])
apply (possibility, simp add: used_Cons)
done
@@ -275,7 +275,7 @@
apply blast
txt{*NS3*}
apply (blast dest!: Crypt_Spy_analz_bad A_trusts_NS2
- dest: Says_imp_knows_Spy analz.Inj unique_session_keys)
+ dest: Says_imp_knows_Spy analz.Inj unique_session_keys)
txt{*Oops*}
apply (blast dest: unique_session_keys)
done
@@ -355,8 +355,8 @@
"\<lbrakk>B \<notin> bad; evs \<in> ns_shared\<rbrakk> \<Longrightarrow>
Key K \<notin> analz (spies evs) \<longrightarrow>
Says Server A
- (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
- Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
+ (Crypt (shrK A) \<lbrace>NA, Agent B, Key K,
+ Crypt (shrK B) \<lbrace>Key K, Agent A\<rbrace>\<rbrace>) \<in> set evs \<longrightarrow>
Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace> \<in> parts (spies evs) \<longrightarrow>
Says A B (Crypt K \<lbrace>Nonce NB, Nonce NB\<rbrace>) \<in> set evs"
apply (erule ns_shared.induct, force)
@@ -366,7 +366,7 @@
apply (blast dest!: new_keys_not_used Crypt_imp_keysFor)
txt{*NS5*}
apply (blast dest!: A_trusts_NS2
- dest: Says_imp_knows_Spy [THEN analz.Inj]
+ dest: Says_imp_knows_Spy [THEN analz.Inj]
unique_session_keys Crypt_Spy_analz_bad)
done
--- a/src/HOL/Auth/OtwayRees.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/OtwayRees.thy Sat Oct 17 14:43:18 2009 +0200
@@ -61,7 +61,7 @@
# evs3 : otway"
(*Bob receives the Server's (?) message and compares the Nonces with
- those in the message he previously sent the Server.
+ those in the message he previously sent the Server.
Need B \<noteq> Server because we allow messages to self.*)
| OR4: "[| evs4 \<in> otway; B \<noteq> Server;
Says B Server {|Nonce NA, Agent A, Agent B, X',
--- a/src/HOL/Auth/OtwayReesBella.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/OtwayReesBella.thy Sat Oct 17 14:43:18 2009 +0200
@@ -22,50 +22,50 @@
Nil: "[]\<in> orb"
| Fake: "\<lbrakk>evsa\<in> orb; X\<in> synth (analz (knows Spy evsa))\<rbrakk>
- \<Longrightarrow> Says Spy B X # evsa \<in> orb"
+ \<Longrightarrow> Says Spy B X # evsa \<in> orb"
| Reception: "\<lbrakk>evsr\<in> orb; Says A B X \<in> set evsr\<rbrakk>
- \<Longrightarrow> Gets B X # evsr \<in> orb"
+ \<Longrightarrow> Gets B X # evsr \<in> orb"
| OR1: "\<lbrakk>evs1\<in> orb; Nonce NA \<notin> used evs1\<rbrakk>
- \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B,
- Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
- # evs1 \<in> orb"
+ \<Longrightarrow> Says A B \<lbrace>Nonce M, Agent A, Agent B,
+ Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+ # evs1 \<in> orb"
| OR2: "\<lbrakk>evs2\<in> orb; Nonce NB \<notin> used evs2;
- Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
- \<Longrightarrow> Says B Server
- \<lbrace>Nonce M, Agent A, Agent B, X,
- Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
- # evs2 \<in> orb"
+ Gets B \<lbrace>Nonce M, Agent A, Agent B, X\<rbrace> \<in> set evs2\<rbrakk>
+ \<Longrightarrow> Says B Server
+ \<lbrace>Nonce M, Agent A, Agent B, X,
+ Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+ # evs2 \<in> orb"
| OR3: "\<lbrakk>evs3\<in> orb; Key KAB \<notin> used evs3;
- Gets Server
- \<lbrace>Nonce M, Agent A, Agent B,
- Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>,
- Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
- \<in> set evs3\<rbrakk>
- \<Longrightarrow> Says Server B \<lbrace>Nonce M,
- Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
- Nonce NB, Key KAB\<rbrace>\<rbrace>
- # evs3 \<in> orb"
+ Gets Server
+ \<lbrace>Nonce M, Agent A, Agent B,
+ Crypt (shrK A) \<lbrace>Nonce NA, Nonce M, Agent A, Agent B\<rbrace>,
+ Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+ \<in> set evs3\<rbrakk>
+ \<Longrightarrow> Says Server B \<lbrace>Nonce M,
+ Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+ Nonce NB, Key KAB\<rbrace>\<rbrace>
+ # evs3 \<in> orb"
(*B can only check that the message he is bouncing is a ciphertext*)
(*Sending M back is omitted*)
| OR4: "\<lbrakk>evs4\<in> orb; B \<noteq> Server; \<forall> p q. X \<noteq> \<lbrace>p, q\<rbrace>;
- Says B Server \<lbrace>Nonce M, Agent A, Agent B, X',
- Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
- \<in> set evs4;
- Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
- \<in> set evs4\<rbrakk>
- \<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
+ Says B Server \<lbrace>Nonce M, Agent A, Agent B, X',
+ Crypt (shrK B) \<lbrace>Nonce NB, Nonce M, Nonce M, Agent A, Agent B\<rbrace>\<rbrace>
+ \<in> set evs4;
+ Gets B \<lbrace>Nonce M, Crypt (shrK B) \<lbrace>X, Nonce NB, Key KAB\<rbrace>\<rbrace>
+ \<in> set evs4\<rbrakk>
+ \<Longrightarrow> Says B A \<lbrace>Nonce M, X\<rbrace> # evs4 \<in> orb"
| Oops: "\<lbrakk>evso\<in> orb;
- Says Server B \<lbrace>Nonce M,
- Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
- Nonce NB, Key KAB\<rbrace>\<rbrace>
- \<in> set evso\<rbrakk>
+ Says Server B \<lbrace>Nonce M,
+ Crypt (shrK B) \<lbrace>Crypt (shrK A) \<lbrace>Nonce NA, Key KAB\<rbrace>,
+ Nonce NB, Key KAB\<rbrace>\<rbrace>
+ \<in> set evso\<rbrakk>
\<Longrightarrow> Notes Spy \<lbrace>Agent A, Agent B, Nonce NA, Nonce NB, Key KAB\<rbrace> # evso
\<in> orb"
--- a/src/HOL/Auth/OtwayRees_AN.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/OtwayRees_AN.thy Sat Oct 17 14:43:18 2009 +0200
@@ -35,7 +35,7 @@
| Reception: --{*A message that has been sent can be received by the
intended recipient.*}
- "[| evsr \<in> otway; Says A B X \<in>set evsr |]
+ "[| evsr \<in> otway; Says A B X \<in>set evsr |]
==> Gets B X # evsr \<in> otway"
| OR1: --{*Alice initiates a protocol run*}
@@ -43,14 +43,14 @@
==> Says A B {|Agent A, Agent B, Nonce NA|} # evs1 \<in> otway"
| OR2: --{*Bob's response to Alice's message.*}
- "[| evs2 \<in> otway;
+ "[| evs2 \<in> otway;
Gets B {|Agent A, Agent B, Nonce NA|} \<in>set evs2 |]
==> Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
# evs2 \<in> otway"
| OR3: --{*The Server receives Bob's message. Then he sends a new
session key to Bob with a packet for forwarding to Alice.*}
- "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
+ "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
Gets Server {|Agent A, Agent B, Nonce NA, Nonce NB|}
\<in>set evs3 |]
==> Says Server B
@@ -59,9 +59,9 @@
# evs3 \<in> otway"
| OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
- those in the message he previously sent the Server.
+ those in the message he previously sent the Server.
Need @{term "B \<noteq> Server"} because we allow messages to self.*}
- "[| evs4 \<in> otway; B \<noteq> Server;
+ "[| evs4 \<in> otway; B \<noteq> Server;
Says B Server {|Agent A, Agent B, Nonce NA, Nonce NB|} \<in>set evs4;
Gets B {|X, Crypt(shrK B){|Nonce NB,Agent A,Agent B,Key K|}|}
\<in>set evs4 |]
@@ -69,7 +69,7 @@
| Oops: --{*This message models possible leaks of session keys. The nonces
identify the protocol run.*}
- "[| evso \<in> otway;
+ "[| evso \<in> otway;
Says Server B
{|Crypt (shrK A) {|Nonce NA, Agent A, Agent B, Key K|},
Crypt (shrK B) {|Nonce NB, Agent A, Agent B, Key K|}|}
--- a/src/HOL/Auth/OtwayRees_Bad.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/OtwayRees_Bad.thy Sat Oct 17 14:43:18 2009 +0200
@@ -32,18 +32,18 @@
| Reception: --{*A message that has been sent can be received by the
intended recipient.*}
- "[| evsr \<in> otway; Says A B X \<in>set evsr |]
+ "[| evsr \<in> otway; Says A B X \<in>set evsr |]
==> Gets B X # evsr \<in> otway"
| OR1: --{*Alice initiates a protocol run*}
- "[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
+ "[| evs1 \<in> otway; Nonce NA \<notin> used evs1 |]
==> Says A B {|Nonce NA, Agent A, Agent B,
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|} |}
# evs1 \<in> otway"
| OR2: --{*Bob's response to Alice's message.
This variant of the protocol does NOT encrypt NB.*}
- "[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
+ "[| evs2 \<in> otway; Nonce NB \<notin> used evs2;
Gets B {|Nonce NA, Agent A, Agent B, X|} \<in> set evs2 |]
==> Says B Server
{|Nonce NA, Agent A, Agent B, X, Nonce NB,
@@ -53,7 +53,7 @@
| OR3: --{*The Server receives Bob's message and checks that the three NAs
match. Then he sends a new session key to Bob with a packet for
forwarding to Alice.*}
- "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
+ "[| evs3 \<in> otway; Key KAB \<notin> used evs3;
Gets Server
{|Nonce NA, Agent A, Agent B,
Crypt (shrK A) {|Nonce NA, Agent A, Agent B|},
@@ -67,9 +67,9 @@
# evs3 \<in> otway"
| OR4: --{*Bob receives the Server's (?) message and compares the Nonces with
- those in the message he previously sent the Server.
+ those in the message he previously sent the Server.
Need @{term "B \<noteq> Server"} because we allow messages to self.*}
- "[| evs4 \<in> otway; B \<noteq> Server;
+ "[| evs4 \<in> otway; B \<noteq> Server;
Says B Server {|Nonce NA, Agent A, Agent B, X', Nonce NB,
Crypt (shrK B) {|Nonce NA, Agent A, Agent B|}|}
\<in> set evs4;
@@ -79,7 +79,7 @@
| Oops: --{*This message models possible leaks of session keys. The nonces
identify the protocol run.*}
- "[| evso \<in> otway;
+ "[| evso \<in> otway;
Says Server B {|Nonce NA, X, Crypt (shrK B) {|Nonce NB, Key K|}|}
\<in> set evso |]
==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> otway"
--- a/src/HOL/Auth/Recur.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Recur.thy Sat Oct 17 14:43:18 2009 +0200
@@ -101,7 +101,7 @@
etc.
Oops: "[| evso \<in> recur; Says Server B RB \<in> set evso;
- RB \<in> responses evs'; Key K \<in> parts {RB} |]
+ RB \<in> responses evs'; Key K \<in> parts {RB} |]
==> Notes Spy {|Key K, RB|} # evso \<in> recur"
*)
@@ -140,10 +140,10 @@
apply (rule_tac [2]
recur.Nil
[THEN recur.RA1 [of _ NA],
- THEN recur.RA2 [of _ NB],
- THEN recur.RA3 [OF _ _ respond.One
+ THEN recur.RA2 [of _ NB],
+ THEN recur.RA3 [OF _ _ respond.One
[THEN respond.Cons [of _ _ K _ K']]],
- THEN recur.RA4], possibility)
+ THEN recur.RA4], possibility)
apply (auto simp add: used_Cons)
done
@@ -241,7 +241,7 @@
(K \<in> KK | Key K \<in> analz (insert RB H))"
apply (erule responses.induct)
apply (simp_all del: image_insert
- add: analz_image_freshK_simps, auto)
+ add: analz_image_freshK_simps, auto)
done
--- a/src/HOL/Auth/Smartcard/EventSC.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Smartcard/EventSC.thy Sat Oct 17 14:43:18 2009 +0200
@@ -38,13 +38,13 @@
specification (stolen)
(*The server's card is secure by assumption\<dots>*)
Card_Server_not_stolen [iff]: "Card Server \<notin> stolen"
- Card_Spy_not_stolen [iff]: "Card Spy \<notin> stolen"
+ Card_Spy_not_stolen [iff]: "Card Spy \<notin> stolen"
apply blast done
specification (cloned)
(*\<dots> the spy's card is secure because she already can use it freely*)
Card_Server_not_cloned [iff]: "Card Server \<notin> cloned"
- Card_Spy_not_cloned [iff]: "Card Spy \<notin> cloned"
+ Card_Spy_not_cloned [iff]: "Card Spy \<notin> cloned"
apply blast done
@@ -52,28 +52,28 @@
assumption of secure means*)
knows_Nil: "knows A [] = initState A"
knows_Cons: "knows A (ev # evs) =
- (case ev of
- Says A' B X =>
+ (case ev of
+ Says A' B X =>
if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
- | Notes A' X =>
- if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs)
+ | Notes A' X =>
+ if (A=A' | (A=Spy & A'\<in>bad)) then insert X (knows A evs)
else knows A evs
| Gets A' X =>
- if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
+ if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
else knows A evs
| Inputs A' C X =>
- if secureM then
+ if secureM then
if A=A' then insert X (knows A evs) else knows A evs
- else
- if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
+ else
+ if (A=A' | A=Spy) then insert X (knows A evs) else knows A evs
| C_Gets C X => knows A evs
| Outpts C A' X =>
- if secureM then
+ if secureM then
if A=A' then insert X (knows A evs) else knows A evs
else
- if A=Spy then insert X (knows A evs) else knows A evs
+ if A=Spy then insert X (knows A evs) else knows A evs
| A_Gets A' X =>
- if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
+ if (A=A' & A \<noteq> Spy) then insert X (knows A evs)
else knows A evs)"
@@ -86,14 +86,14 @@
primrec
used_Nil: "used [] = (UN B. parts (initState B))"
used_Cons: "used (ev # evs) =
- (case ev of
- Says A B X => parts {X} \<union> (used evs)
- | Notes A X => parts {X} \<union> (used evs)
- | Gets A X => used evs
+ (case ev of
+ Says A B X => parts {X} \<union> (used evs)
+ | Notes A X => parts {X} \<union> (used evs)
+ | Gets A X => used evs
| Inputs A C X => parts{X} \<union> (used evs)
- | C_Gets C X => used evs
+ | C_Gets C X => used evs
| Outpts C A X => parts{X} \<union> (used evs)
- | A_Gets A X => used evs)"
+ | A_Gets A X => used evs)"
--{*@{term Gets} always follows @{term Says} in real protocols.
Likewise, @{term C_Gets} will always have to follow @{term Inputs}
and @{term A_Gets} will always have to follow @{term Outpts}*}
--- a/src/HOL/Auth/Smartcard/Smartcard.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy Sat Oct 17 14:43:18 2009 +0200
@@ -386,8 +386,8 @@
val analz_image_freshK_ss =
@{simpset} delsimps [image_insert, image_Un]
- delsimps [@{thm imp_disjL}] (*reduces blow-up*)
- addsimps @{thms analz_image_freshK_simps}
+ delsimps [@{thm imp_disjL}] (*reduces blow-up*)
+ addsimps @{thms analz_image_freshK_simps}
end
*}
--- a/src/HOL/Auth/TLS.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/TLS.thy Sat Oct 17 14:43:18 2009 +0200
@@ -36,7 +36,7 @@
Proofs would be simpler if ClientKeyExch included A's name within
Crypt KB (Nonce PMS). As things stand, there is much overlap between proofs
about that message (which B receives) and the stronger event
- Notes A {|Agent B, Nonce PMS|}.
+Notes A {|Agent B, Nonce PMS|}.
*)
header{*The TLS Protocol: Transport Layer Security*}
@@ -112,30 +112,30 @@
| SpyKeys: --{*The spy may apply @{term PRF} and @{term sessionK}
to available nonces*}
"[| evsSK \<in> tls;
- {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
+ {Nonce NA, Nonce NB, Nonce M} <= analz (spies evsSK) |]
==> Notes Spy {| Nonce (PRF(M,NA,NB)),
- Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
+ Key (sessionK((NA,NB,M),role)) |} # evsSK \<in> tls"
| ClientHello:
- --{*(7.4.1.2)
- PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
- It is uninterpreted but will be confirmed in the FINISHED messages.
- NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
+ --{*(7.4.1.2)
+ PA represents @{text CLIENT_VERSION}, @{text CIPHER_SUITES} and @{text COMPRESSION_METHODS}.
+ It is uninterpreted but will be confirmed in the FINISHED messages.
+ NA is CLIENT RANDOM, while SID is @{text SESSION_ID}.
UNIX TIME is omitted because the protocol doesn't use it.
May assume @{term "NA \<notin> range PRF"} because CLIENT RANDOM is
28 bytes while MASTER SECRET is 48 bytes*}
"[| evsCH \<in> tls; Nonce NA \<notin> used evsCH; NA \<notin> range PRF |]
==> Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
- # evsCH \<in> tls"
+ # evsCH \<in> tls"
| ServerHello:
--{*7.4.1.3 of the TLS Internet-Draft
- PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
+ PB represents @{text CLIENT_VERSION}, @{text CIPHER_SUITE} and @{text COMPRESSION_METHOD}.
SERVER CERTIFICATE (7.4.2) is always present.
@{text CERTIFICATE_REQUEST} (7.4.4) is implied.*}
"[| evsSH \<in> tls; Nonce NB \<notin> used evsSH; NB \<notin> range PRF;
Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
- \<in> set evsSH |]
+ \<in> set evsSH |]
==> Says B A {|Nonce NB, Number SID, Number PB|} # evsSH \<in> tls"
| Certificate:
@@ -148,28 +148,28 @@
She encrypts PMS using the supplied KB, which ought to be pubK B.
We assume @{term "PMS \<notin> range PRF"} because a clash betweem the PMS
and another MASTER SECRET is highly unlikely (even though
- both items have the same length, 48 bytes).
+ both items have the same length, 48 bytes).
The Note event records in the trace that she knows PMS
(see REMARK at top). *}
"[| evsCX \<in> tls; Nonce PMS \<notin> used evsCX; PMS \<notin> range PRF;
Says B' A (certificate B KB) \<in> set evsCX |]
==> Says A B (Crypt KB (Nonce PMS))
- # Notes A {|Agent B, Nonce PMS|}
- # evsCX \<in> tls"
+ # Notes A {|Agent B, Nonce PMS|}
+ # evsCX \<in> tls"
| CertVerify:
- --{*The optional Certificate Verify (7.4.8) message contains the
+ --{*The optional Certificate Verify (7.4.8) message contains the
specific components listed in the security analysis, F.1.1.2.
It adds the pre-master-secret, which is also essential!
Checking the signature, which is the only use of A's certificate,
assures B of A's presence*}
"[| evsCV \<in> tls;
Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCV;
- Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
+ Notes A {|Agent B, Nonce PMS|} \<in> set evsCV |]
==> Says A B (Crypt (priK A) (Hash{|Nonce NB, Agent B, Nonce PMS|}))
# evsCV \<in> tls"
- --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
+ --{*Finally come the FINISHED messages (7.4.8), confirming PA and PB
among other things. The master-secret is PRF(PMS,NA,NB).
Either party may send its message first.*}
@@ -181,60 +181,60 @@
could simply put @{term "A\<noteq>Spy"} into the rule, but one should not
expect the spy to be well-behaved.*}
"[| evsCF \<in> tls;
- Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
- \<in> set evsCF;
+ Says A B {|Agent A, Nonce NA, Number SID, Number PA|}
+ \<in> set evsCF;
Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCF;
Notes A {|Agent B, Nonce PMS|} \<in> set evsCF;
- M = PRF(PMS,NA,NB) |]
+ M = PRF(PMS,NA,NB) |]
==> Says A B (Crypt (clientK(NA,NB,M))
- (Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|}))
+ (Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|}))
# evsCF \<in> tls"
| ServerFinished:
- --{*Keeping A' and A'' distinct means B cannot even check that the
+ --{*Keeping A' and A'' distinct means B cannot even check that the
two messages originate from the same source. *}
"[| evsSF \<in> tls;
- Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
- \<in> set evsSF;
- Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
- Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
- M = PRF(PMS,NA,NB) |]
+ Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}
+ \<in> set evsSF;
+ Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSF;
+ Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSF;
+ M = PRF(PMS,NA,NB) |]
==> Says B A (Crypt (serverK(NA,NB,M))
- (Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|}))
+ (Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|}))
# evsSF \<in> tls"
| ClientAccepts:
- --{*Having transmitted ClientFinished and received an identical
+ --{*Having transmitted ClientFinished and received an identical
message encrypted with serverK, the client stores the parameters
needed to resume this session. The "Notes A ..." premise is
used to prove @{text Notes_master_imp_Crypt_PMS}.*}
"[| evsCA \<in> tls;
- Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
- M = PRF(PMS,NA,NB);
- X = Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|};
+ Notes A {|Agent B, Nonce PMS|} \<in> set evsCA;
+ M = PRF(PMS,NA,NB);
+ X = Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|};
Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evsCA;
Says B' A (Crypt (serverK(NA,NB,M)) X) \<in> set evsCA |]
==>
Notes A {|Number SID, Agent A, Agent B, Nonce M|} # evsCA \<in> tls"
| ServerAccepts:
- --{*Having transmitted ServerFinished and received an identical
+ --{*Having transmitted ServerFinished and received an identical
message encrypted with clientK, the server stores the parameters
needed to resume this session. The "Says A'' B ..." premise is
used to prove @{text Notes_master_imp_Crypt_PMS}.*}
"[| evsSA \<in> tls;
- A \<noteq> B;
+ A \<noteq> B;
Says A'' B (Crypt (pubK B) (Nonce PMS)) \<in> set evsSA;
- M = PRF(PMS,NA,NB);
- X = Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|};
+ M = PRF(PMS,NA,NB);
+ X = Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|};
Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evsSA;
Says A' B (Crypt (clientK(NA,NB,M)) X) \<in> set evsSA |]
==>
@@ -244,27 +244,27 @@
--{*If A recalls the @{text SESSION_ID}, then she sends a FINISHED
message using the new nonces and stored MASTER SECRET.*}
"[| evsCR \<in> tls;
- Says A B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
+ Says A B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsCR;
Says B' A {|Nonce NB, Number SID, Number PB|} \<in> set evsCR;
Notes A {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsCR |]
==> Says A B (Crypt (clientK(NA,NB,M))
- (Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|}))
+ (Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|}))
# evsCR \<in> tls"
| ServerResume:
--{*Resumption (7.3): If B finds the @{text SESSION_ID} then he can
send a FINISHED message using the recovered MASTER SECRET*}
"[| evsSR \<in> tls;
- Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
- Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
+ Says A' B {|Agent A, Nonce NA, Number SID, Number PA|}: set evsSR;
+ Says B A {|Nonce NB, Number SID, Number PB|} \<in> set evsSR;
Notes B {|Number SID, Agent A, Agent B, Nonce M|} \<in> set evsSR |]
==> Says B A (Crypt (serverK(NA,NB,M))
- (Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|})) # evsSR
- \<in> tls"
+ (Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|})) # evsSR
+ \<in> tls"
| Oops:
--{*The most plausible compromise is of an old session key. Losing
@@ -273,7 +273,7 @@
otherwise the Spy could learn session keys merely by
replaying messages!*}
"[| evso \<in> tls; A \<noteq> Spy;
- Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
+ Says A B (Crypt (sessionK((NA,NB,M),role)) X) \<in> set evso |]
==> Says A Spy (Key (sessionK((NA,NB,M),role))) # evso \<in> tls"
(*
@@ -328,7 +328,7 @@
(** These proofs assume that the Nonce_supply nonces
- (which have the form @ N. Nonce N \<notin> used evs)
+ (which have the form @ N. Nonce N \<notin> used evs)
lie outside the range of PRF. It seems reasonable, but as it is needed
only for the possibility theorems, it is not taken as an axiom.
**)
@@ -381,11 +381,11 @@
\<forall>evs. (@ N. Nonce N \<notin> used evs) \<notin> range PRF;
A \<noteq> B |]
==> \<exists>NA PA NB PB X. \<exists>evs \<in> tls.
- X = Hash{|Number SID, Nonce M,
- Nonce NA, Number PA, Agent A,
- Nonce NB, Number PB, Agent B|} &
- Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs &
- Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
+ X = Hash{|Number SID, Nonce M,
+ Nonce NA, Number PA, Agent A,
+ Nonce NB, Number PB, Agent B|} &
+ Says A B (Crypt (clientK(NA,NB,M)) X) \<in> set evs &
+ Says B A (Crypt (serverK(NA,NB,M)) X) \<in> set evs"
apply (intro exI bexI)
apply (rule_tac [2] tls.ClientHello
[THEN tls.ServerHello,
@@ -570,7 +570,7 @@
(priK B \<in> KK | B \<in> bad)"
apply (erule tls.induct)
apply (simp_all (no_asm_simp)
- del: image_insert
+ del: image_insert
add: image_Un [THEN sym]
insert_Key_image Un_assoc [THEN sym])
txt{*Fake*}
@@ -598,16 +598,16 @@
lemma analz_image_keys [rule_format]:
"evs \<in> tls ==>
\<forall>KK. KK <= range sessionK -->
- (Nonce N \<in> analz (Key`KK Un (spies evs))) =
- (Nonce N \<in> analz (spies evs))"
+ (Nonce N \<in> analz (Key`KK Un (spies evs))) =
+ (Nonce N \<in> analz (spies evs))"
apply (erule tls.induct, frule_tac [7] CX_KB_is_pubKB)
apply (safe del: iffI)
apply (safe del: impI iffI intro!: analz_image_keys_lemma)
apply (simp_all (no_asm_simp) (*faster*)
del: image_insert imp_disjL (*reduces blow-up*)
- add: image_Un [THEN sym] Un_assoc [THEN sym]
- insert_Key_singleton
- range_sessionkeys_not_priK analz_image_priK)
+ add: image_Un [THEN sym] Un_assoc [THEN sym]
+ insert_Key_singleton
+ range_sessionkeys_not_priK analz_image_priK)
apply (simp_all add: insert_absorb)
txt{*Fake*}
apply spy_analz
@@ -901,11 +901,11 @@
down to 433s on albatross*)
(*5/5/01: conversion to Isar script
- loads in 137s (perch)
+ loads in 137s (perch)
the last ML version loaded in 122s on perch, a 600MHz machine:
- twice as fast as pike. No idea why it's so much slower!
- The Isar script is slower still, perhaps because simp_all simplifies
- the assumptions be default.
+ twice as fast as pike. No idea why it's so much slower!
+ The Isar script is slower still, perhaps because simp_all simplifies
+ the assumptions be default.
*)
end
--- a/src/HOL/Auth/Yahalom.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Yahalom.thy Sat Oct 17 14:43:18 2009 +0200
@@ -58,7 +58,7 @@
uses the new session key to send Bob his Nonce. The premise
@{term "A \<noteq> Server"} is needed for @{text Says_Server_not_range}.
Alice can check that K is symmetric by its length.*}
- "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys;
+ "[| evs4 \<in> yahalom; A \<noteq> Server; K \<in> symKeys;
Gets A {|Crypt(shrK A) {|Agent B, Key K, Nonce NA, Nonce NB|}, X|}
\<in> set evs4;
Says A B {|Agent A, Nonce NA|} \<in> set evs4 |]
@@ -481,9 +481,9 @@
text{*A vital theorem for B, that nonce NB remains secure from the Spy.*}
lemma Spy_not_see_NB :
"[| Says B Server
- {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
- \<in> set evs;
- (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
+ {|Agent B, Crypt (shrK B) {|Agent A, Nonce NA, Nonce NB|}|}
+ \<in> set evs;
+ (\<forall>k. Notes Spy {|Nonce NA, Nonce NB, k|} \<notin> set evs);
A \<notin> bad; B \<notin> bad; evs \<in> yahalom |]
==> Nonce NB \<notin> analz (knows Spy evs)"
apply (erule rev_mp, erule rev_mp)
--- a/src/HOL/Auth/Yahalom2.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/Yahalom2.thy Sat Oct 17 14:43:18 2009 +0200
@@ -51,7 +51,7 @@
Both agents are quoted in the 2nd certificate to prevent attacks!*)
| YM3: "[| evs3 \<in> yahalom; Key KAB \<notin> used evs3;
Gets Server {|Agent B, Nonce NB,
- Crypt (shrK B) {|Agent A, Nonce NA|}|}
+ Crypt (shrK B) {|Agent A, Nonce NA|}|}
\<in> set evs3 |]
==> Says Server A
{|Nonce NB,
--- a/src/HOL/Auth/ZhouGollmann.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Auth/ZhouGollmann.thy Sat Oct 17 14:43:18 2009 +0200
@@ -35,7 +35,7 @@
Nil: "[] \<in> zg"
| Fake: "[| evsf \<in> zg; X \<in> synth (analz (spies evsf)) |]
- ==> Says Spy B X # evsf \<in> zg"
+ ==> Says Spy B X # evsf \<in> zg"
| Reception: "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg"
@@ -44,26 +44,26 @@
We just assume that the protocol's objective is to deliver K fairly,
rather than to keep M secret.*)
| ZG1: "[| evs1 \<in> zg; Nonce L \<notin> used evs1; C = Crypt K (Number m);
- K \<in> symKeys;
- NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
+ K \<in> symKeys;
+ NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg"
(*B must check that NRO is A's signature to learn the sender's name*)
| ZG2: "[| evs2 \<in> zg;
- Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
- NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
- NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
+ Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
+ NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
+ NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} # evs2 \<in> zg"
(*A must check that NRR is B's signature to learn the sender's name;
without spy, the matching label would be enough*)
| ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys;
- Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
- Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
- NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
- sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
+ Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
+ Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
+ NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
+ sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
- # evs3 \<in> zg"
+ # evs3 \<in> zg"
(*TTP checks that sub_K is A's signature to learn who issued K, then
gives credentials to A and B. The Notes event models the availability of
@@ -72,15 +72,15 @@
also allowing lemma @{text Crypt_used_imp_spies} to omit the condition
@{term "K \<noteq> priK TTP"}. *)
| ZG4: "[| evs4 \<in> zg; K \<in> symKeys;
- Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
- \<in> set evs4;
- sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
- con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
- Nonce L, Key K|}|]
+ Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
+ \<in> set evs4;
+ sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
+ con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
+ Nonce L, Key K|}|]
==> Says TTP Spy con_K
#
- Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
- # evs4 \<in> zg"
+ Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
+ # evs4 \<in> zg"
declare Says_imp_knows_Spy [THEN analz.Inj, dest]
--- a/src/HOL/Bali/AxCompl.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/AxCompl.thy Sat Oct 17 14:43:18 2009 +0200
@@ -319,13 +319,13 @@
case 0
with is_cls
show ?thesis
- by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
+ by - (rule ax_impossible [THEN conseq1],fastsimp dest: nyinitcls_emptyD)
next
case (Suc m)
with mgf_hyp have mgf_hyp': "\<And> t. G,A\<turnstile>{=:m} t\<succ> {G\<rightarrow>}"
- by simp
+ by simp
from is_cls obtain c where c: "the (class G C) = c"
- by auto
+ by auto
let ?Q= "(\<lambda>Y s' (x,s) .
G\<turnstile> (x,init_class_obj G C s)
\<midarrow> (if C=Object then Skip else Init (super (the (class G C))))\<rightarrow> s'
@@ -333,45 +333,45 @@
from c
show ?thesis
proof (rule ax_derivs.Init [where ?Q="?Q"])
- let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s
+ let ?P' = "Normal ((\<lambda>Y s' s. s' = supd (init_class_obj G C) s
\<and> normal s \<and> \<not> initd C s) \<and>. G\<turnstile>init\<le>m)"
- show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
+ show "G,A\<turnstile>{Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))}
.(if C = Object then Skip else Init (super c)).
{?Q}"
- proof (rule conseq1 [where ?P'="?P'"])
- show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
- proof (cases "C=Object")
- case True
- have "G,A\<turnstile>{?P'} .Skip. {?Q}"
- by (rule ax_derivs.Skip [THEN conseq1])
- (auto simp add: True intro: eval.Skip)
+ proof (rule conseq1 [where ?P'="?P'"])
+ show "G,A\<turnstile>{?P'} .(if C = Object then Skip else Init (super c)). {?Q}"
+ proof (cases "C=Object")
+ case True
+ have "G,A\<turnstile>{?P'} .Skip. {?Q}"
+ by (rule ax_derivs.Skip [THEN conseq1])
+ (auto simp add: True intro: eval.Skip)
with True show ?thesis
- by simp
- next
- case False
- from mgf_hyp'
- have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
- by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
- with False show ?thesis
- by simp
- qed
- next
- from Suc is_cls
- show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
+ by simp
+ next
+ case False
+ from mgf_hyp'
+ have "G,A\<turnstile>{?P'} .Init (super c). {?Q}"
+ by (rule MGFnD' [THEN conseq12]) (fastsimp simp add: False c)
+ with False show ?thesis
+ by simp
+ qed
+ next
+ from Suc is_cls
+ show "Normal (?P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C))
\<Rightarrow> ?P'"
- by (fastsimp elim: nyinitcls_le_SucD)
- qed
+ by (fastsimp elim: nyinitcls_le_SucD)
+ qed
next
- from mgf_hyp'
- show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty}
+ from mgf_hyp'
+ show "\<forall>l. G,A\<turnstile>{?Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty}
.init c.
{set_lvars l .; ?R}"
- apply (rule MGFnD' [THEN conseq12, THEN allI])
- apply (clarsimp simp add: split_paired_all)
- apply (rule eval.Init [OF c])
- apply (insert c)
- apply auto
- done
+ apply (rule MGFnD' [THEN conseq12, THEN allI])
+ apply (clarsimp simp add: split_paired_all)
+ apply (rule eval.Init [OF c])
+ apply (insert c)
+ apply auto
+ done
qed
qed
thus "G,A\<turnstile>{Normal ?P \<and>. Not \<circ> initd C} .Init C. {?R}"
@@ -399,7 +399,7 @@
mode: "mode = invmode statM e" and
T: "T =Inl (resTy statM)" and
eq_accC_accC': "accC=accC'"
- by cases fastsimp+
+ by cases fastsimp+
let ?Q="(\<lambda>Y s1 (x,s) . x = None \<and>
(\<exists>a. G\<turnstile>Norm s \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
(normal s1 \<longrightarrow> G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT)
@@ -435,29 +435,29 @@
(\<exists>P. \<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P))
\<and> s1\<Colon>\<preceq>(G, L)"
proof -
- from da obtain C where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+ from da obtain C where
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
dom (locals (store ((Norm s0)::state)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
- da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E"
- by cases (simp add: eq_accC_accC')
- from eval_e conf_s0 wt_e da_e wf
- obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
- and "s1\<Colon>\<preceq>(G, L)"
- by (rule eval_type_soundE) simp
- moreover
- {
- assume normal_s1: "normal s1"
- have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
- proof -
- from eval_e wt_e da_e wf normal_s1
- have "nrm C \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approxE') iprover
- with da_ps show ?thesis
- by (rule da_weakenE) iprover
- qed
- }
- ultimately show ?thesis
- using eq_accC_accC' by simp
+ da_ps: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> E"
+ by cases (simp add: eq_accC_accC')
+ from eval_e conf_s0 wt_e da_e wf
+ obtain "(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT)"
+ and "s1\<Colon>\<preceq>(G, L)"
+ by (rule eval_type_soundE) simp
+ moreover
+ {
+ assume normal_s1: "normal s1"
+ have "\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+ proof -
+ from eval_e wt_e da_e wf normal_s1
+ have "nrm C \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approxE') iprover
+ with da_ps show ?thesis
+ by (rule da_weakenE) iprover
+ qed
+ }
+ ultimately show ?thesis
+ using eq_accC_accC' by simp
qed
qed
next
@@ -467,36 +467,36 @@
show "?PS a"
proof (rule MGFnD' [OF mgf_ps, THEN conseq12],
clarsimp simp add: eq_accC_accC' [symmetric])
- fix s0 s1 s2 vs
- assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
- assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
- assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
- assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
- assume da_ps: "abrupt s1 = None \<longrightarrow>
+ fix s0 s1 s2 vs
+ assume conf_s1: "s1\<Colon>\<preceq>(G, L)"
+ assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+ assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+ assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+ assume da_ps: "abrupt s1 = None \<longrightarrow>
(\<exists>P. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P)"
- show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
+ show "(\<exists>s1. G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1 \<and>
(abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT) \<and>
G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2) \<and>
s2\<Colon>\<preceq>(G, L)"
- proof (cases "normal s1")
- case True
- with da_ps obtain P where
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
- by auto
- from eval_ps conf_s1 wt_args this wf
- have "s2\<Colon>\<preceq>(G, L)"
- by (rule eval_type_soundE)
- with eval_e conf_a eval_ps
- show ?thesis
- by auto
- next
- case False
- with eval_ps have "s2=s1" by auto
- with eval_e conf_a eval_ps conf_s1
- show ?thesis
- by auto
- qed
+ proof (cases "normal s1")
+ case True
+ with da_ps obtain P where
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>ps\<rangle>\<^sub>l\<guillemotright> P"
+ by auto
+ from eval_ps conf_s1 wt_args this wf
+ have "s2\<Colon>\<preceq>(G, L)"
+ by (rule eval_type_soundE)
+ with eval_e conf_a eval_ps
+ show ?thesis
+ by auto
+ next
+ case False
+ with eval_ps have "s2=s1" by auto
+ with eval_e conf_a eval_ps conf_s1
+ show ?thesis
+ by auto
+ qed
qed
qed
next
@@ -517,52 +517,52 @@
from mgf_methds [rule_format]
show "?METHD a vs invC declC l"
proof (rule MGFnD' [THEN conseq12],clarsimp)
- fix s4 s2 s1::state
- fix s0 v
- let ?D= "invocation_declclass G mode (store s2) a statT
+ fix s4 s2 s1::state
+ fix s0 v
+ let ?D= "invocation_declclass G mode (store s2) a statT
\<lparr>name=mn,parTs=pTs'\<rparr>"
- let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
- assume inv_prop: "abrupt ?s3=None
+ let ?s3= "init_lvars G ?D \<lparr>name=mn, parTs=pTs'\<rparr> mode a vs s2"
+ assume inv_prop: "abrupt ?s3=None
\<longrightarrow> G\<turnstile>mode\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
- assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
- assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
- assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
- assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
- assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
- show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
+ assume conf_s2: "s2\<Colon>\<preceq>(G, L)"
+ assume conf_a: "abrupt s1 = None \<longrightarrow> G,store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+ assume eval_e: "G\<turnstile>Norm s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
+ assume eval_ps: "G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>vs\<rightarrow> s2"
+ assume eval_mthd: "G\<turnstile>?s3 \<midarrow>Methd ?D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+ show "G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn( {pTs'}ps)-\<succ>v
\<rightarrow> (set_lvars (locals (store s2))) s4"
- proof -
- obtain D where D: "D=?D" by simp
- obtain s3 where s3: "s3=?s3" by simp
- obtain s3' where
- s3': "s3' = check_method_access G accC statT mode
+ proof -
+ obtain D where D: "D=?D" by simp
+ obtain s3 where s3: "s3=?s3" by simp
+ obtain s3' where
+ s3': "s3' = check_method_access G accC statT mode
\<lparr>name=mn,parTs=pTs'\<rparr> a s3"
- by simp
- have eq_s3'_s3: "s3'=s3"
- proof -
- from inv_prop s3 mode
- have "normal s3 \<Longrightarrow>
+ by simp
+ have eq_s3'_s3: "s3'=s3"
+ proof -
+ from inv_prop s3 mode
+ have "normal s3 \<Longrightarrow>
G\<turnstile>invmode statM e\<rightarrow>invocation_class mode (store s2) a statT\<preceq>statT"
- by auto
- with eval_ps wt_e statM conf_s2 conf_a [rule_format]
- have "check_method_access G accC statT (invmode statM e)
+ by auto
+ with eval_ps wt_e statM conf_s2 conf_a [rule_format]
+ have "check_method_access G accC statT (invmode statM e)
\<lparr>name=mn,parTs=pTs'\<rparr> a s3 = s3"
- by (rule error_free_call_access) (auto simp add: s3 mode wf)
- thus ?thesis
- by (simp add: s3' mode)
- qed
- with eval_mthd D s3
- have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
- by simp
- with eval_e eval_ps D _ s3'
- show ?thesis
- by (rule eval_Call) (auto simp add: s3 mode D)
- qed
+ by (rule error_free_call_access) (auto simp add: s3 mode wf)
+ thus ?thesis
+ by (simp add: s3' mode)
+ qed
+ with eval_mthd D s3
+ have "G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<rightarrow> s4"
+ by simp
+ with eval_e eval_ps D _ s3'
+ show ?thesis
+ by (rule eval_Call) (auto simp add: s3 mode D)
+ qed
qed
qed
qed
qed
-
+
lemma eval_expression_no_jump':
assumes eval: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<rightarrow> s1"
and no_jmp: "abrupt s0 \<noteq> Some (Jump j)"
@@ -610,36 +610,36 @@
case True
with normal_t eval_e normal_termination
show ?thesis
- by (auto intro: eval.Loop)
+ by (auto intro: eval.Loop)
next
case False
note abrupt_s' = this
from eval_e _ wt wf
have no_cont: "abrupt s' \<noteq> Some (Jump (Cont l))"
- by (rule eval_expression_no_jump') (insert normal_t,simp)
+ by (rule eval_expression_no_jump') (insert normal_t,simp)
have
- "if the_Bool v
+ "if the_Bool v
then (G\<turnstile>s' \<midarrow>c\<rightarrow> s' \<and>
G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s')
- else s' = s'"
+ else s' = s'"
proof (cases "the_Bool v")
- case False thus ?thesis by simp
+ case False thus ?thesis by simp
next
- case True
- with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
- moreover from abrupt_s' no_cont
- have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
- by (cases s') (simp add: absorb_def split: split_if)
- moreover
- from no_absorb abrupt_s'
- have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
- by auto
- ultimately show ?thesis
- using True by simp
+ case True
+ with abrupt_s' have "G\<turnstile>s' \<midarrow>c\<rightarrow> s'" by auto
+ moreover from abrupt_s' no_cont
+ have no_absorb: "(abupd (absorb (Cont l)) s')=s'"
+ by (cases s') (simp add: absorb_def split: split_if)
+ moreover
+ from no_absorb abrupt_s'
+ have "G\<turnstile>(abupd (absorb (Cont l)) s') \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+ by auto
+ ultimately show ?thesis
+ using True by simp
qed
with eval_e
show ?thesis
- using normal_t by (auto intro: eval.Loop)
+ using normal_t by (auto intro: eval.Loop)
qed
qed
next
@@ -690,49 +690,49 @@
Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s'))
\<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>}"
proof (rule ax_derivs.Loop)
- from mfg_e
- show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n}
+ from mfg_e
+ show "G,A\<turnstile>{(\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n}
e-\<succ>
{(\<lambda>Y s' s. (\<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
Y = In1 b \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s'))
\<and>. G\<turnstile>init\<le>n}"
- proof (rule MGFnD' [THEN conseq12],clarsimp)
- fix s Z s' v
- assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
- moreover
- assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
- ultimately
- show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
- by blast
- qed
+ proof (rule MGFnD' [THEN conseq12],clarsimp)
+ fix s Z s' v
+ assume "(Z, s) \<in> (unroll G l e c)\<^sup>*"
+ moreover
+ assume "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s'"
+ ultimately
+ show "\<exists>t. (Z, t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'"
+ by blast
+ qed
next
- from mfg_c
- show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
+ from mfg_c
+ show "G,A\<turnstile>{Normal (((\<lambda>Y s' s. \<exists>t b. (s, t) \<in> (unroll G l e c)\<^sup>* \<and>
Y = \<lfloor>b\<rfloor>\<^sub>e \<and> G\<turnstile>t \<midarrow>e-\<succ>b\<rightarrow> s')
\<and>. G\<turnstile>init\<le>n)\<leftarrow>=True)}
.c.
{abupd (absorb (Cont l)) .;
((\<lambda>Y s' s. (s, s') \<in> (unroll G l e c)\<^sup>*) \<and>. G\<turnstile>init\<le>n)}"
- proof (rule MGFnD' [THEN conseq12],clarsimp)
- fix Z s' s v t
- assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
- assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s"
- assume true: "the_Bool v"
- assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
- show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
- proof -
- note unroll
- also
- from eval_e true eval_c
- have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c"
- by (unfold unroll_def) force
- ultimately show ?thesis ..
- qed
- qed
+ proof (rule MGFnD' [THEN conseq12],clarsimp)
+ fix Z s' s v t
+ assume unroll: "(Z, t) \<in> (unroll G l e c)\<^sup>*"
+ assume eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> Norm s"
+ assume true: "the_Bool v"
+ assume eval_c: "G\<turnstile>Norm s \<midarrow>c\<rightarrow> s'"
+ show "(Z, abupd (absorb (Cont l)) s') \<in> (unroll G l e c)\<^sup>*"
+ proof -
+ note unroll
+ also
+ from eval_e true eval_c
+ have "(t,abupd (absorb (Cont l)) s') \<in> unroll G l e c"
+ by (unfold unroll_def) force
+ ultimately show ?thesis ..
+ qed
+ qed
qed
next
show
- "\<forall>Y s Z.
+ "\<forall>Y s Z.
(Normal ((\<lambda>Y' s' s. s' = s \<and> normal s) \<and>. G\<turnstile>init\<le>n)) Y s Z
\<longrightarrow> (\<forall>Y' s'.
(\<forall>Y Z'.
@@ -742,28 +742,28 @@
\<and>. G\<turnstile>init\<le>n)\<leftarrow>=False\<down>=\<diamondsuit>) Y' s' Z')
\<longrightarrow> G\<turnstile>Z \<midarrow>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ>\<rightarrow> (Y', s'))"
proof (clarsimp)
- fix Y' s' s
- assume asm:
- "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>*
+ fix Y' s' s
+ assume asm:
+ "\<forall>Z'. (Z', Norm s) \<in> (unroll G l e c)\<^sup>*
\<longrightarrow> card (nyinitcls G s') \<le> n \<and>
(\<exists>v. (\<exists>t. (Z', t) \<in> (unroll G l e c)\<^sup>* \<and> G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s') \<and>
(fst s' = None \<longrightarrow> \<not> the_Bool v)) \<and> Y' = \<diamondsuit>"
- show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
- proof -
- from asm obtain v t where
- -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}
- unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
+ show "Y' = \<diamondsuit> \<and> G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+ proof -
+ from asm obtain v t where
+ -- {* @{term "Z'"} gets instantiated with @{term "Norm s"} *}
+ unroll: "(Norm s, t) \<in> (unroll G l e c)\<^sup>*" and
eval_e: "G\<turnstile>t \<midarrow>e-\<succ>v\<rightarrow> s'" and
normal_termination: "normal s' \<longrightarrow> \<not> the_Bool v" and
- Y': "Y' = \<diamondsuit>"
- by auto
- from unroll eval_e normal_termination wt_e wf
- have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
- by (rule unroll_while)
- with Y'
- show ?thesis
- by simp
- qed
+ Y': "Y' = \<diamondsuit>"
+ by auto
+ from unroll eval_e normal_termination wt_e wf
+ have "G\<turnstile>Norm s \<midarrow>l\<bullet> While(e) c\<rightarrow> s'"
+ by (rule unroll_while)
+ with Y'
+ show ?thesis
+ by simp
+ qed
qed
qed
qed
@@ -810,39 +810,39 @@
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>
s'\<Colon>\<preceq>(G, L)"
proof -
- from da
- obtain E where
- "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp
- moreover
- from eval_init
- have "dom (locals s) \<subseteq> dom (locals (store s'))"
- by (rule dom_locals_eval_mono [elim_format]) simp
- ultimately obtain E' where
- "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
- by (rule da_weakenE)
- moreover
- have "s'\<Colon>\<preceq>(G, L)"
- proof -
- have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
- proof -
- from wf wt_e
- have iscls_statC: "is_class G statC"
- by (auto dest: ty_expr_is_type type_is_class)
- with wf accfield
- have iscls_statDeclC: "is_class G statDeclC"
- by (auto dest!: accfield_fields dest: fields_declC)
- thus ?thesis by simp
- qed
- obtain I where
- da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ from da
+ obtain E where
+ "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals s) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp
+ moreover
+ from eval_init
+ have "dom (locals s) \<subseteq> dom (locals (store s'))"
+ by (rule dom_locals_eval_mono [elim_format]) simp
+ ultimately obtain E' where
+ "\<lparr>prg=G, cls=accC', lcl=L\<rparr>\<turnstile> dom (locals (store s')) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+ by (rule da_weakenE)
+ moreover
+ have "s'\<Colon>\<preceq>(G, L)"
+ proof -
+ have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
+ proof -
+ from wf wt_e
+ have iscls_statC: "is_class G statC"
+ by (auto dest: ty_expr_is_type type_is_class)
+ with wf accfield
+ have iscls_statDeclC: "is_class G statDeclC"
+ by (auto dest!: accfield_fields dest: fields_declC)
+ thus ?thesis by simp
+ qed
+ obtain I where
+ da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store ((Norm s)::state))) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
- by (auto intro: da_Init [simplified] assigned.select_convs)
- from eval_init conf_s wt_init da_init wf
- show ?thesis
- by (rule eval_type_soundE)
- qed
- ultimately show ?thesis by iprover
+ by (auto intro: da_Init [simplified] assigned.select_convs)
+ from eval_init conf_s wt_init da_init wf
+ show ?thesis
+ by (rule eval_type_soundE)
+ qed
+ ultimately show ?thesis by iprover
qed
qed
next
@@ -857,30 +857,30 @@
assume da_e: "\<lparr>prg=G,cls=accC',lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
show "G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>fst ?fvar\<rightarrow> snd ?fvar"
proof -
- obtain v s2' where
- v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
- by simp
- obtain s3 where
- s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
- by simp
- have eq_s3_s2': "s3=s2'"
- proof -
- from eval_e conf_s1 wt_e da_e wf obtain
- conf_s2: "s2\<Colon>\<preceq>(G, L)" and
- conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
- by (rule eval_type_soundE) simp
- from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
- show ?thesis
- by (rule error_free_field_access
+ obtain v s2' where
+ v: "v=fst ?fvar" and s2': "s2'=snd ?fvar"
+ by simp
+ obtain s3 where
+ s3: "s3= check_field_access G accC' statDeclC fn stat a s2'"
+ by simp
+ have eq_s3_s2': "s3=s2'"
+ proof -
+ from eval_e conf_s1 wt_e da_e wf obtain
+ conf_s2: "s2\<Colon>\<preceq>(G, L)" and
+ conf_a: "normal s2 \<Longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+ by (rule eval_type_soundE) simp
+ from accfield wt_e eval_init eval_e conf_s2 conf_a _ wf
+ show ?thesis
+ by (rule error_free_field_access
[where ?v=v and ?s2'=s2',elim_format])
- (simp add: s3 v s2' stat)+
+ (simp add: s3 v s2' stat)+
qed
- from eval_init eval_e
- show ?thesis
- apply (rule eval.FVar [where ?s2'=s2'])
- apply (simp add: s2')
- apply (simp add: s3 [symmetric] eq_s3_s2' eq_accC s2' [symmetric])
- done
+ from eval_init eval_e
+ show ?thesis
+ apply (rule eval.FVar [where ?s2'=s2'])
+ apply (simp add: s2')
+ apply (simp add: s3 [symmetric] eq_s3_s2' eq_accC s2' [symmetric])
+ done
qed
qed
qed
@@ -918,7 +918,7 @@
fix s0
assume "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1 Finally c2\<rangle>\<^sub>s\<guillemotright> C"
thus "\<exists>C1. \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1"
- by cases (auto simp add: inj_term_simps)
+ by cases (auto simp add: inj_term_simps)
qed
next
from mgf_c2
@@ -933,25 +933,25 @@
show "G\<turnstile>Norm s0 \<midarrow>c1 Finally c2
\<rightarrow> abupd (abrupt_if (\<exists>y. abrupt s1 = Some y) (abrupt s1)) s2"
proof -
- obtain abr1 str1 where s1: "s1=(abr1,str1)"
- by (cases s1)
- with eval_c1 eval_c2 obtain
- eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
- eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
- by simp
- obtain s3 where
- s3: "s3 = (if \<exists>err. abr1 = Some (Error err)
- then (abr1, str1)
+ obtain abr1 str1 where s1: "s1=(abr1,str1)"
+ by (cases s1)
+ with eval_c1 eval_c2 obtain
+ eval_c1': "G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (abr1,str1)" and
+ eval_c2': "G\<turnstile>Norm str1 \<midarrow>c2\<rightarrow> s2"
+ by simp
+ obtain s3 where
+ s3: "s3 = (if \<exists>err. abr1 = Some (Error err)
+ then (abr1, str1)
else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
- by simp
- from eval_c1' conf_s0 wt_c1 _ wf
- have "error_free (abr1,str1)"
- by (rule eval_type_soundE) (insert da_c1,auto)
- with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
- by (simp add: error_free_def)
- from eval_c1' eval_c2' s3
- show ?thesis
- by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
+ by simp
+ from eval_c1' conf_s0 wt_c1 _ wf
+ have "error_free (abr1,str1)"
+ by (rule eval_type_soundE) (insert da_c1,auto)
+ with s3 have eq_s3: "s3=abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
+ by (simp add: error_free_def)
+ from eval_c1' eval_c2' s3
+ show ?thesis
+ by (rule eval.Fin [elim_format]) (simp add: s1 eq_s3)
qed
qed
qed
@@ -1009,7 +1009,7 @@
fix s0
assume da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals s0) \<guillemotright>\<langle>Body D c\<rangle>\<^sub>e\<guillemotright> E"
thus "jumpNestingOkS {Ret} c"
- by cases simp
+ by cases simp
qed
next
from mgf_c
@@ -1022,22 +1022,22 @@
show "G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
\<rightarrow> abupd (absorb Ret) s2"
proof -
- from wt obtain d where
+ from wt obtain d where
d: "class G D=Some d" and
wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
- by cases auto
- obtain s3 where
- s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
+ by cases auto
+ obtain s3 where
+ s3: "s3= (if \<exists>l. fst s2 = Some (Jump (Break l)) \<or>
fst s2 = Some (Jump (Cont l))
then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2
else s2)"
- by simp
- from eval_init eval_c nestingOk wt_c d wf
- have eq_s3_s2: "s3=s2"
- by (rule Body_no_break [elim_format]) (simp add: s3)
- from eval_init eval_c s3
- show ?thesis
- by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
+ by simp
+ from eval_init eval_c nestingOk wt_c d wf
+ have eq_s3_s2: "s3=s2"
+ by (rule Body_no_break [elim_format]) (simp add: s3)
+ from eval_init eval_c s3
+ show ?thesis
+ by (rule eval.Body [elim_format]) (simp add: eq_s3_s2)
qed
qed
qed
@@ -1062,312 +1062,312 @@
proof (induct rule: var_expr_stmt.inducts)
case (LVar v)
show "G,A\<turnstile>{=:n} \<langle>LVar v\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.LVar [THEN conseq1])
- apply (clarsimp)
- apply (rule eval.LVar)
- done
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.LVar [THEN conseq1])
+ apply (clarsimp)
+ apply (rule eval.LVar)
+ done
next
case (FVar accC statDeclC stat e fn)
from MGFn_Init [OF hyp] and `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}` and wf
show ?case
- by (rule MGFn_FVar)
+ by (rule MGFn_FVar)
next
case (AVar e1 e2)
note mgf_e1 = `G,A\<turnstile>{=:n} \<langle>e1\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
note mgf_e2 = `G,A\<turnstile>{=:n} \<langle>e2\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
show "G,A\<turnstile>{=:n} \<langle>e1.[e2]\<rangle>\<^sub>v\<succ> {G\<rightarrow>}"
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.AVar)
- apply (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
- apply (rule allI)
- apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
- apply (fastsimp intro: eval.AVar)
- done
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.AVar)
+ apply (rule MGFnD [OF mgf_e1, THEN ax_NormalD])
+ apply (rule allI)
+ apply (rule MGFnD' [OF mgf_e2, THEN conseq12])
+ apply (fastsimp intro: eval.AVar)
+ done
next
case (InsInitV c v)
show ?case
- by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
+ by (rule MGFn_NormalI) (rule ax_derivs.InsInitV)
next
case (NewC C)
show ?case
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.NewC)
- apply (rule MGFn_InitD [OF hyp, THEN conseq2])
- apply (fastsimp intro: eval.NewC)
- done
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.NewC)
+ apply (rule MGFn_InitD [OF hyp, THEN conseq2])
+ apply (fastsimp intro: eval.NewC)
+ done
next
case (NewA T e)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.NewA
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.NewA
[where ?Q = "(\<lambda>Y' s' s. normal s \<and> G\<turnstile>s \<midarrow>In1r (init_comp_ty T)
\<succ>\<rightarrow> (Y',s')) \<and>. G\<turnstile>init\<le>n"])
- apply (simp add: init_comp_ty_def split add: split_if)
- apply (rule conjI, clarsimp)
- apply (rule MGFn_InitD [OF hyp, THEN conseq2])
- apply (clarsimp intro: eval.Init)
- apply clarsimp
- apply (rule ax_derivs.Skip [THEN conseq1])
- apply (clarsimp intro: eval.Skip)
- apply (erule MGFnD' [THEN conseq12])
- apply (fastsimp intro: eval.NewA)
- done
+ apply (simp add: init_comp_ty_def split add: split_if)
+ apply (rule conjI, clarsimp)
+ apply (rule MGFn_InitD [OF hyp, THEN conseq2])
+ apply (clarsimp intro: eval.Init)
+ apply clarsimp
+ apply (rule ax_derivs.Skip [THEN conseq1])
+ apply (clarsimp intro: eval.Skip)
+ apply (erule MGFnD' [THEN conseq12])
+ apply (fastsimp intro: eval.NewA)
+ done
next
case (Cast C e)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
- apply (fastsimp intro: eval.Cast)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Cast])
+ apply (fastsimp intro: eval.Cast)
+ done
next
case (Inst e C)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
- apply (fastsimp intro: eval.Inst)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Inst])
+ apply (fastsimp intro: eval.Inst)
+ done
next
case (Lit v)
show ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Lit [THEN conseq1])
- apply (fastsimp intro: eval.Lit)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Lit [THEN conseq1])
+ apply (fastsimp intro: eval.Lit)
+ done
next
case (UnOp unop e)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.UnOp)
- apply (erule MGFnD' [THEN conseq12])
- apply (fastsimp intro: eval.UnOp)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.UnOp)
+ apply (erule MGFnD' [THEN conseq12])
+ apply (fastsimp intro: eval.UnOp)
+ done
next
case (BinOp binop e1 e2)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.BinOp)
- apply (erule MGFnD [THEN ax_NormalD])
- apply (rule allI)
- apply (case_tac "need_second_arg binop v1")
- apply simp
- apply (erule MGFnD' [THEN conseq12])
- apply (fastsimp intro: eval.BinOp)
- apply simp
- apply (rule ax_Normal_cases)
- apply (rule ax_derivs.Skip [THEN conseq1])
- apply clarsimp
- apply (rule eval_BinOp_arg2_indepI)
- apply simp
- apply simp
- apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
- apply (fastsimp intro: eval.BinOp)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.BinOp)
+ apply (erule MGFnD [THEN ax_NormalD])
+ apply (rule allI)
+ apply (case_tac "need_second_arg binop v1")
+ apply simp
+ apply (erule MGFnD' [THEN conseq12])
+ apply (fastsimp intro: eval.BinOp)
+ apply simp
+ apply (rule ax_Normal_cases)
+ apply (rule ax_derivs.Skip [THEN conseq1])
+ apply clarsimp
+ apply (rule eval_BinOp_arg2_indepI)
+ apply simp
+ apply simp
+ apply (rule ax_derivs.Abrupt [THEN conseq1], clarsimp simp add: Let_def)
+ apply (fastsimp intro: eval.BinOp)
+ done
next
case Super
show ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Super [THEN conseq1])
- apply (fastsimp intro: eval.Super)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Super [THEN conseq1])
+ apply (fastsimp intro: eval.Super)
+ done
next
case (Acc v)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
- apply (fastsimp intro: eval.Acc simp add: split_paired_all)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Acc])
+ apply (fastsimp intro: eval.Acc simp add: split_paired_all)
+ done
next
case (Ass v e)
thus "G,A\<turnstile>{=:n} \<langle>v:=e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Ass)
- apply (erule MGFnD [THEN ax_NormalD])
- apply (rule allI)
- apply (erule MGFnD'[THEN conseq12])
- apply (fastsimp intro: eval.Ass simp add: split_paired_all)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Ass)
+ apply (erule MGFnD [THEN ax_NormalD])
+ apply (rule allI)
+ apply (erule MGFnD'[THEN conseq12])
+ apply (fastsimp intro: eval.Ass simp add: split_paired_all)
+ done
next
case (Cond e1 e2 e3)
thus "G,A\<turnstile>{=:n} \<langle>e1 ? e2 : e3\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Cond)
- apply (erule MGFnD [THEN ax_NormalD])
- apply (rule allI)
- apply (rule ax_Normal_cases)
- prefer 2
- apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
- apply (fastsimp intro: eval.Cond)
- apply (case_tac "b")
- apply simp
- apply (erule MGFnD'[THEN conseq12])
- apply (fastsimp intro: eval.Cond)
- apply simp
- apply (erule MGFnD'[THEN conseq12])
- apply (fastsimp intro: eval.Cond)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Cond)
+ apply (erule MGFnD [THEN ax_NormalD])
+ apply (rule allI)
+ apply (rule ax_Normal_cases)
+ prefer 2
+ apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+ apply (fastsimp intro: eval.Cond)
+ apply (case_tac "b")
+ apply simp
+ apply (erule MGFnD'[THEN conseq12])
+ apply (fastsimp intro: eval.Cond)
+ apply simp
+ apply (erule MGFnD'[THEN conseq12])
+ apply (fastsimp intro: eval.Cond)
+ done
next
case (Call accC statT mode e mn pTs' ps)
note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
note mgf_ps = `G,A\<turnstile>{=:n} \<langle>ps\<rangle>\<^sub>l\<succ> {G\<rightarrow>}`
from mgf_methds mgf_e mgf_ps wf
show "G,A\<turnstile>{=:n} \<langle>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
- by (rule MGFn_Call)
+ by (rule MGFn_Call)
next
case (Methd D mn)
from mgf_methds
show "G,A\<turnstile>{=:n} \<langle>Methd D mn\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
- by simp
+ by simp
next
case (Body D c)
note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
from wf MGFn_Init [OF hyp] mgf_c
show "G,A\<turnstile>{=:n} \<langle>Body D c\<rangle>\<^sub>e\<succ> {G\<rightarrow>}"
- by (rule MGFn_Body)
+ by (rule MGFn_Body)
next
case (InsInitE c e)
show ?case
- by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
+ by (rule MGFn_NormalI) (rule ax_derivs.InsInitE)
next
case (Callee l e)
show ?case
- by (rule MGFn_NormalI) (rule ax_derivs.Callee)
+ by (rule MGFn_NormalI) (rule ax_derivs.Callee)
next
case Skip
show ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Skip [THEN conseq1])
- apply (fastsimp intro: eval.Skip)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Skip [THEN conseq1])
+ apply (fastsimp intro: eval.Skip)
+ done
next
case (Expr e)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
- apply (fastsimp intro: eval.Expr)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (erule MGFnD'[THEN conseq12,THEN ax_derivs.Expr])
+ apply (fastsimp intro: eval.Expr)
+ done
next
case (Lab l c)
thus "G,A\<turnstile>{=:n} \<langle>l\<bullet> c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
- apply (fastsimp intro: eval.Lab)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Lab])
+ apply (fastsimp intro: eval.Lab)
+ done
next
case (Comp c1 c2)
thus "G,A\<turnstile>{=:n} \<langle>c1;; c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Comp)
- apply (erule MGFnD [THEN ax_NormalD])
- apply (erule MGFnD' [THEN conseq12])
- apply (fastsimp intro: eval.Comp)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Comp)
+ apply (erule MGFnD [THEN ax_NormalD])
+ apply (erule MGFnD' [THEN conseq12])
+ apply (fastsimp intro: eval.Comp)
+ done
next
case (If' e c1 c2)
thus "G,A\<turnstile>{=:n} \<langle>If(e) c1 Else c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.If)
- apply (erule MGFnD [THEN ax_NormalD])
- apply (rule allI)
- apply (rule ax_Normal_cases)
- prefer 2
- apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
- apply (fastsimp intro: eval.If)
- apply (case_tac "b")
- apply simp
- apply (erule MGFnD' [THEN conseq12])
- apply (fastsimp intro: eval.If)
- apply simp
- apply (erule MGFnD' [THEN conseq12])
- apply (fastsimp intro: eval.If)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.If)
+ apply (erule MGFnD [THEN ax_NormalD])
+ apply (rule allI)
+ apply (rule ax_Normal_cases)
+ prefer 2
+ apply (rule ax_derivs.Abrupt [THEN conseq1],clarsimp simp add: Let_def)
+ apply (fastsimp intro: eval.If)
+ apply (case_tac "b")
+ apply simp
+ apply (erule MGFnD' [THEN conseq12])
+ apply (fastsimp intro: eval.If)
+ apply simp
+ apply (erule MGFnD' [THEN conseq12])
+ apply (fastsimp intro: eval.If)
+ done
next
case (Loop l e c)
note mgf_e = `G,A\<turnstile>{=:n} \<langle>e\<rangle>\<^sub>e\<succ> {G\<rightarrow>}`
note mgf_c = `G,A\<turnstile>{=:n} \<langle>c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
from mgf_e mgf_c wf
show "G,A\<turnstile>{=:n} \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- by (rule MGFn_Loop)
+ by (rule MGFn_Loop)
next
case (Jmp j)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Jmp [THEN conseq1])
- apply (auto intro: eval.Jmp simp add: abupd_def2)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Jmp [THEN conseq1])
+ apply (auto intro: eval.Jmp simp add: abupd_def2)
+ done
next
case (Throw e)
thus ?case
- apply -
- apply (rule MGFn_NormalI)
- apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
- apply (fastsimp intro: eval.Throw)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (erule MGFnD' [THEN conseq12, THEN ax_derivs.Throw])
+ apply (fastsimp intro: eval.Throw)
+ done
next
case (TryC c1 C vn c2)
thus "G,A\<turnstile>{=:n} \<langle>Try c1 Catch(C vn) c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Try [where
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Try [where
?Q = " (\<lambda>Y' s' s. normal s \<and> (\<exists>s''. G\<turnstile>s \<midarrow>\<langle>c1\<rangle>\<^sub>s\<succ>\<rightarrow> (Y',s'') \<and>
G\<turnstile>s'' \<midarrow>sxalloc\<rightarrow> s')) \<and>. G\<turnstile>init\<le>n"])
- apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
- apply (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
- apply (erule MGFnD'[THEN conseq12])
- apply (fastsimp intro: eval.Try)
- apply (fastsimp intro: eval.Try)
- done
+ apply (erule MGFnD [THEN ax_NormalD, THEN conseq2])
+ apply (fastsimp elim: sxalloc_gext [THEN card_nyinitcls_gext])
+ apply (erule MGFnD'[THEN conseq12])
+ apply (fastsimp intro: eval.Try)
+ apply (fastsimp intro: eval.Try)
+ done
next
case (Fin c1 c2)
note mgf_c1 = `G,A\<turnstile>{=:n} \<langle>c1\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
note mgf_c2 = `G,A\<turnstile>{=:n} \<langle>c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}`
from wf mgf_c1 mgf_c2
show "G,A\<turnstile>{=:n} \<langle>c1 Finally c2\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- by (rule MGFn_Fin)
+ by (rule MGFn_Fin)
next
case (FinA abr c)
show ?case
- by (rule MGFn_NormalI) (rule ax_derivs.FinA)
+ by (rule MGFn_NormalI) (rule ax_derivs.FinA)
next
case (Init C)
from hyp
show "G,A\<turnstile>{=:n} \<langle>Init C\<rangle>\<^sub>s\<succ> {G\<rightarrow>}"
- by (rule MGFn_Init)
+ by (rule MGFn_Init)
next
case Nil_expr
show "G,A\<turnstile>{=:n} \<langle>[]\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Nil [THEN conseq1])
- apply (fastsimp intro: eval.Nil)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Nil [THEN conseq1])
+ apply (fastsimp intro: eval.Nil)
+ done
next
case (Cons_expr e es)
thus "G,A\<turnstile>{=:n} \<langle>e# es\<rangle>\<^sub>l\<succ> {G\<rightarrow>}"
- apply -
- apply (rule MGFn_NormalI)
- apply (rule ax_derivs.Cons)
- apply (erule MGFnD [THEN ax_NormalD])
- apply (rule allI)
- apply (erule MGFnD'[THEN conseq12])
- apply (fastsimp intro: eval.Cons)
- done
+ apply -
+ apply (rule MGFn_NormalI)
+ apply (rule ax_derivs.Cons)
+ apply (erule MGFnD [THEN ax_NormalD])
+ apply (rule allI)
+ apply (erule MGFnD'[THEN conseq12])
+ apply (fastsimp intro: eval.Cons)
+ done
qed
}
thus ?thesis
@@ -1519,15 +1519,15 @@
proof -
from eval_t type_ok wf
obtain n where evaln: "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s')"
- by (rule eval_to_evaln [elim_format]) iprover
+ by (rule eval_to_evaln [elim_format]) iprover
from valid have
- valid_expanded:
- "\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s
+ valid_expanded:
+ "\<forall>n Y s Z. P Y s Z \<longrightarrow> type_ok G t s
\<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y', s') \<longrightarrow> Q Y' s' Z)"
- by (simp add: ax_valids_def triple_valid_def)
+ by (simp add: ax_valids_def triple_valid_def)
from P type_ok evaln
show "Q Y' s' Z"
- by (rule valid_expanded [rule_format])
+ by (rule valid_expanded [rule_format])
qed
qed
qed
--- a/src/HOL/Bali/AxSem.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/AxSem.thy Sat Oct 17 14:43:18 2009 +0200
@@ -519,7 +519,7 @@
G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}"
| NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q}; G,A\<turnstile>{Q} e-\<succ>
- {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
+ {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow>
G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}"
| Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s..
--- a/src/HOL/Bali/AxSound.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/AxSound.thy Sat Oct 17 14:43:18 2009 +0200
@@ -118,14 +118,14 @@
case 0
show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>0\<Colon>t"
proof -
- {
- fix C sig
- assume "(C,sig) \<in> ms"
- have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
- by (rule Methd_triple_valid2_0)
- }
- thus ?thesis
- by (simp add: mtriples_def split_def)
+ {
+ fix C sig
+ assume "(C,sig) \<in> ms"
+ have "G\<Turnstile>0\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+ by (rule Methd_triple_valid2_0)
+ }
+ thus ?thesis
+ by (simp add: mtriples_def split_def)
qed
next
case (Suc m)
@@ -133,28 +133,28 @@
note prem = `\<forall>t\<in>A. G\<Turnstile>Suc m\<Colon>t`
show "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>Suc m\<Colon>t"
proof -
- {
- fix C sig
- assume m: "(C,sig) \<in> ms"
- have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
- proof -
- from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
- by (rule triples_valid2_Suc)
- hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
- by (rule hyp)
- with prem_m
- have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
- by (simp add: ball_Un)
- hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
- by (rule recursive)
- with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
- by (auto simp add: mtriples_def split_def)
- thus ?thesis
- by (rule Methd_triple_valid2_SucI)
- qed
- }
- thus ?thesis
- by (simp add: mtriples_def split_def)
+ {
+ fix C sig
+ assume m: "(C,sig) \<in> ms"
+ have "G\<Turnstile>Suc m\<Colon>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}"
+ proof -
+ from prem have prem_m: "\<forall>t\<in>A. G\<Turnstile>m\<Colon>t"
+ by (rule triples_valid2_Suc)
+ hence "\<forall>t\<in>{{P} Methd-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
+ by (rule hyp)
+ with prem_m
+ have "\<forall>t\<in>(A \<union> {{P} Methd-\<succ> {Q} | ms}). G\<Turnstile>m\<Colon>t"
+ by (simp add: ball_Un)
+ hence "\<forall>t\<in>{{P} body G-\<succ> {Q} | ms}. G\<Turnstile>m\<Colon>t"
+ by (rule recursive)
+ with m have "G\<Turnstile>m\<Colon>{Normal (P C sig)} body G C sig-\<succ> {Q C sig}"
+ by (auto simp add: mtriples_def split_def)
+ thus ?thesis
+ by (rule Methd_triple_valid2_SucI)
+ qed
+ }
+ thus ?thesis
+ by (simp add: mtriples_def split_def)
qed
qed
}
@@ -361,10 +361,10 @@
have "G\<Turnstile>n\<Colon>t" and "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
proof -
from valid_A valid_t show "G\<Turnstile>n\<Colon>t"
- by (simp add: ax_valids2_def)
+ by (simp add: ax_valids2_def)
next
from valid_A valid_ts show "\<forall>t\<in>ts. G\<Turnstile>n\<Colon>t"
- by (unfold ax_valids2_def) blast
+ by (unfold ax_valids2_def) blast
qed
hence "\<forall>t'\<in>insert t ts. G\<Turnstile>n\<Colon>t'"
by simp
@@ -402,27 +402,27 @@
proof -
from valid_A conf wt da eval P con
have "Q v s1 Z"
- apply (simp add: ax_valids2_def triple_valid2_def2)
- apply (tactic "smp_tac 3 1")
- apply clarify
- apply (tactic "smp_tac 1 1")
- apply (erule allE,erule allE, erule mp)
- apply (intro strip)
- apply (tactic "smp_tac 3 1")
- apply (tactic "smp_tac 2 1")
- apply (tactic "smp_tac 1 1")
- by blast
+ apply (simp add: ax_valids2_def triple_valid2_def2)
+ apply (tactic "smp_tac 3 1")
+ apply clarify
+ apply (tactic "smp_tac 1 1")
+ apply (erule allE,erule allE, erule mp)
+ apply (intro strip)
+ apply (tactic "smp_tac 3 1")
+ apply (tactic "smp_tac 2 1")
+ apply (tactic "smp_tac 1 1")
+ by blast
moreover have "s1\<Colon>\<preceq>(G, L)"
proof (cases "normal s0")
- case True
- from eval wt [OF True] da [OF True] conf wf
- show ?thesis
- by (rule evaln_type_sound [elim_format]) simp
+ case True
+ from eval wt [OF True] da [OF True] conf wf
+ show ?thesis
+ by (rule evaln_type_sound [elim_format]) simp
next
- case False
- with eval have "s1=s0"
- by auto
- with conf show ?thesis by simp
+ case False
+ with eval have "s1=s0"
+ by auto
+ with conf show ?thesis by simp
qed
ultimately show ?thesis ..
qed
@@ -461,13 +461,13 @@
show "P (In2 vf) s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof
from eval normal_s0 obtain "s1=s0" "vf=lvar vn (store s0)"
- by (fastsimp elim: evaln_elim_cases)
+ by (fastsimp elim: evaln_elim_cases)
with P show "P (In2 vf) s1 Z"
- by simp
+ by simp
next
from eval wt da conf_s0 wf
show "s1\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
qed
qed
next
@@ -490,78 +490,78 @@
from wt obtain statC f where
wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-Class statC" and
accfield: "accfield G accC statC fn = Some (statDeclC,f)" and
- eq_accC: "accC=accC'" and
+ eq_accC: "accC=accC'" and
stat: "stat=is_static f" and
- T: "T=(type f)"
- by (cases) (auto simp add: member_is_static_simp)
+ T: "T=(type f)"
+ by (cases) (auto simp add: member_is_static_simp)
from da eq_accC
have da_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V"
- by cases simp
+ by cases simp
from eval obtain a s1 s2 s2' where
- eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and
+ eval_init: "G\<turnstile>s0 \<midarrow>Init statDeclC\<midarrow>n\<rightarrow> s1" and
eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s2" and
- fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
- s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ fvar: "(vf,s2')=fvar statDeclC stat fn a s2" and
+ s3: "s3 = check_field_access G accC statDeclC fn stat a s2'"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
have wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>(Init statDeclC)\<Colon>\<surd>"
proof -
- from wf wt_e
- have iscls_statC: "is_class G statC"
- by (auto dest: ty_expr_is_type type_is_class)
- with wf accfield
- have iscls_statDeclC: "is_class G statDeclC"
- by (auto dest!: accfield_fields dest: fields_declC)
- thus ?thesis by simp
+ from wf wt_e
+ have iscls_statC: "is_class G statC"
+ by (auto dest: ty_expr_is_type type_is_class)
+ with wf accfield
+ have iscls_statDeclC: "is_class G statDeclC"
+ by (auto dest!: accfield_fields dest: fields_declC)
+ thus ?thesis by simp
qed
obtain I where
- da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init statDeclC\<rangle>\<^sub>s\<guillemotright> I"
- by (auto intro: da_Init [simplified] assigned.select_convs)
+ by (auto intro: da_Init [simplified] assigned.select_convs)
from valid_init P valid_A conf_s0 eval_init wt_init da_init
obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
- by (rule validE)
+ by (rule validE)
obtain
- R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and
+ R: "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and
conf_s2: "s2\<Colon>\<preceq>(G, L)" and
- conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+ conf_a: "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
proof (cases "normal s1")
- case True
- obtain V' where
- da_e':
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
- proof -
- from eval_init
- have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
- by (rule dom_locals_evaln_mono_elim)
- with da_e show thesis
- by (rule da_weakenE) (rule that)
- qed
- with valid_e Q valid_A conf_s1 eval_e wt_e
- obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
- by (rule validE) (simp add: fvar [symmetric])
- moreover
- from eval_e wt_e da_e' conf_s1 wf
- have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
- by (rule evaln_type_sound [elim_format]) simp
- ultimately show ?thesis ..
+ case True
+ obtain V' where
+ da_e':
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> V'"
+ proof -
+ from eval_init
+ have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
+ by (rule dom_locals_evaln_mono_elim)
+ with da_e show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ with valid_e Q valid_A conf_s1 eval_e wt_e
+ obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
+ by (rule validE) (simp add: fvar [symmetric])
+ moreover
+ from eval_e wt_e da_e' conf_s1 wf
+ have "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+ by (rule evaln_type_sound [elim_format]) simp
+ ultimately show ?thesis ..
next
- case False
- with valid_e Q valid_A conf_s1 eval_e
- obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
- by (cases rule: validE) (simp add: fvar [symmetric])+
- moreover from False eval_e have "\<not> normal s2"
- by auto
- hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
- by auto
- ultimately show ?thesis ..
+ case False
+ with valid_e Q valid_A conf_s1 eval_e
+ obtain "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z" and "s2\<Colon>\<preceq>(G, L)"
+ by (cases rule: validE) (simp add: fvar [symmetric])+
+ moreover from False eval_e have "\<not> normal s2"
+ by auto
+ hence "normal s2 \<longrightarrow> G,store s2\<turnstile>a\<Colon>\<preceq>Class statC"
+ by auto
+ ultimately show ?thesis ..
qed
from accfield wt_e eval_init eval_e conf_s2 conf_a fvar stat s3 wf
have eq_s3_s2': "s3=s2'"
- using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
+ using normal_s0 by (auto dest!: error_free_field_access evaln_eval)
moreover
from eval wt da conf_s0 wf
have "s3\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis using Q R by simp
qed
qed
@@ -584,48 +584,48 @@
show "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z \<and> s2'\<Colon>\<preceq>(G, L)"
proof -
from wt obtain
- wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
+ wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T.[]" and
wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-PrimT Integer"
- by (rule wt_elim_cases) simp
+ by (rule wt_elim_cases) simp
from da obtain E1 where
- da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
- da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
- by (rule da_elim_cases) simp
+ da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
+ da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V"
+ by (rule da_elim_cases) simp
from eval obtain s1 a i s2 where
- eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
- eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
- avar: "avar G i a s2 =(vf, s2')"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>a\<midarrow>n\<rightarrow> s1" and
+ eval_e2: "G\<turnstile>s1 \<midarrow>e2-\<succ>i\<midarrow>n\<rightarrow> s2" and
+ avar: "avar G i a s2 =(vf, s2')"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
obtain Q: "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
- by (rule validE)
+ by (rule validE)
from Q have Q': "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
- by simp
+ by simp
have "R \<lfloor>vf\<rfloor>\<^sub>v s2' Z"
proof (cases "normal s1")
- case True
- obtain V' where
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
- proof -
- from eval_e1 wt_e1 da_e1 wf True
- have "nrm E1 \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approx_evalnE) iprover
- with da_e2 show thesis
- by (rule da_weakenE) (rule that)
- qed
- with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2
- show ?thesis
- by (rule validE) (simp add: avar)
+ case True
+ obtain V' where
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile>dom (locals (store s1))\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> V'"
+ proof -
+ from eval_e1 wt_e1 da_e1 wf True
+ have "nrm E1 \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approx_evalnE) iprover
+ with da_e2 show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ with valid_e2 Q' valid_A conf_s1 eval_e2 wt_e2
+ show ?thesis
+ by (rule validE) (simp add: avar)
next
- case False
- with valid_e2 Q' valid_A conf_s1 eval_e2
- show ?thesis
- by (cases rule: validE) (simp add: avar)+
+ case False
+ with valid_e2 Q' valid_A conf_s1 eval_e2
+ show ?thesis
+ by (cases rule: validE) (simp add: avar)+
qed
moreover
from eval wt da conf_s0 wf
have "s2'\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -646,26 +646,26 @@
show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
proof -
from wt obtain is_cls_C: "is_class G C"
- by (rule wt_elim_cases) (auto dest: is_acc_classD)
+ by (rule wt_elim_cases) (auto dest: is_acc_classD)
hence wt_init: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>Init C\<Colon>\<surd>"
- by auto
+ by auto
obtain I where
- da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
- by (auto intro: da_Init [simplified] assigned.select_convs)
+ da_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init C\<rangle>\<^sub>s\<guillemotright> I"
+ by (auto intro: da_Init [simplified] assigned.select_convs)
from eval obtain s1 a where
- eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and
+ eval_init: "G\<turnstile>s0 \<midarrow>Init C\<midarrow>n\<rightarrow> s1" and
alloc: "G\<turnstile>s1 \<midarrow>halloc CInst C\<succ>a\<rightarrow> s2" and
- v: "v=Addr a"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ v: "v=Addr a"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_init P valid_A conf_s0 eval_init wt_init da_init
obtain "(Alloc G (CInst C) Q) \<diamondsuit> s1 Z"
- by (rule validE)
+ by (rule validE)
with alloc v have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -687,58 +687,58 @@
show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
proof -
from wt obtain
- wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
- wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer"
- by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
+ wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>init_comp_ty T\<Colon>\<surd>" and
+ wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Integer"
+ by (rule wt_elim_cases) (auto intro: wt_init_comp_ty )
from da obtain
- da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp
+ da_e:"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp
from eval obtain s1 i s2 a where
- eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and
+ eval_init: "G\<turnstile>s0 \<midarrow>init_comp_ty T\<midarrow>n\<rightarrow> s1" and
eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>i\<midarrow>n\<rightarrow> s2" and
alloc: "G\<turnstile>abupd (check_neg i) s2 \<midarrow>halloc Arr T (the_Intg i)\<succ>a\<rightarrow> s3" and
v: "v=Addr a"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
obtain I where
- da_init:
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
+ da_init:
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s0)) \<guillemotright>\<langle>init_comp_ty T\<rangle>\<^sub>s\<guillemotright> I"
proof (cases "\<exists>C. T = Class C")
- case True
- thus ?thesis
- by - (rule that, (auto intro: da_Init [simplified]
+ case True
+ thus ?thesis
+ by - (rule that, (auto intro: da_Init [simplified]
assigned.select_convs
simp add: init_comp_ty_def))
- (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
+ (* simplified: to rewrite \<langle>Init C\<rangle> to In1r (Init C) *)
next
- case False
- thus ?thesis
- by - (rule that, (auto intro: da_Skip [simplified]
+ case False
+ thus ?thesis
+ by - (rule that, (auto intro: da_Skip [simplified]
assigned.select_convs
simp add: init_comp_ty_def))
(* simplified: to rewrite \<langle>Skip\<rangle> to In1r (Skip) *)
qed
with valid_init P valid_A conf_s0 eval_init wt_init
obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G, L)"
- by (rule validE)
+ by (rule validE)
obtain E' where
"\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
proof -
- from eval_init
- have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
- by (rule dom_locals_evaln_mono_elim)
- with da_e show thesis
- by (rule da_weakenE) (rule that)
+ from eval_init
+ have "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+ by (rule dom_locals_evaln_mono_elim)
+ with da_e show thesis
+ by (rule da_weakenE) (rule that)
qed
with valid_e Q valid_A conf_s1 eval_e wt_e
have "(\<lambda>Val:i:. abupd (check_neg i) .;
Alloc G (Arr T (the_Intg i)) R) \<lfloor>i\<rfloor>\<^sub>e s2 Z"
- by (rule validE)
+ by (rule validE)
with alloc v have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s3\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -760,25 +760,25 @@
show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
proof -
from wt obtain eT where
- wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
- by cases simp
+ wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+ by cases simp
from da obtain
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp
from eval obtain s1 where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1" and
s2: "s2 = abupd (raise_if (\<not> G,snd s1\<turnstile>v fits T) ClassCast) s1"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e P valid_A conf_s0 eval_e wt_e da_e
have "(\<lambda>Val:v:. \<lambda>s.. abupd (raise_if (\<not> G,s\<turnstile>v fits T) ClassCast) .;
Q\<leftarrow>In1 v) \<lfloor>v\<rfloor>\<^sub>e s1 Z"
- by (rule validE)
+ by (rule validE)
with s2 have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -799,25 +799,25 @@
show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from wt obtain eT where
- wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
- by cases simp
+ wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+ by cases simp
from da obtain
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp
from eval obtain a where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
v: "v = Bool (a \<noteq> Null \<and> G,store s1\<turnstile>a fits RefT T)"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e P valid_A conf_s0 eval_e wt_e da_e
have "(\<lambda>Val:v:. \<lambda>s.. Q\<leftarrow>In1 (Bool (v \<noteq> Null \<and> G,s\<turnstile>v fits RefT T)))
\<lfloor>a\<rfloor>\<^sub>e s1 Z"
- by (rule validE)
+ by (rule validE)
with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s1\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -833,7 +833,7 @@
show "P \<lfloor>v'\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from eval have "s1=s0" and "v'=v"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ using normal_s0 by (auto elim: evaln_elim_cases)
with P conf_s0 show ?thesis by simp
qed
qed
@@ -853,24 +853,24 @@
show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from wt obtain eT where
- wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
- by cases simp
+ wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+ by cases simp
from da obtain
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp
from eval obtain ve where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
v: "v = eval_unop unop ve"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e P valid_A conf_s0 eval_e wt_e da_e
have "(\<lambda>Val:v:. Q\<leftarrow>In1 (eval_unop unop v)) \<lfloor>ve\<rfloor>\<^sub>e s1 Z"
- by (rule validE)
+ by (rule validE)
with v have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s1\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -897,63 +897,63 @@
from wt obtain e1T e2T where
wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-e1T" and
wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-e2T" and
- wt_binop: "wt_binop G binop e1T e2T"
- by cases simp
+ wt_binop: "wt_binop G binop e1T e2T"
+ by cases simp
have wt_Skip: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>Skip\<Colon>\<surd>"
- by simp
+ by simp
(*
obtain S where
- daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ daSkip: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store s1)) \<guillemotright>In1r Skip\<guillemotright> S"
- by (auto intro: da_Skip [simplified] assigned.select_convs) *)
+ by (auto intro: da_Skip [simplified] assigned.select_convs) *)
from da obtain E1 where
- da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
- by cases simp+
+ da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1"
+ by cases simp+
from eval obtain v1 s1 v2 where
- eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
- eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
+ eval_e1: "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<midarrow>n\<rightarrow> s1" and
+ eval_e2: "G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)
\<succ>\<midarrow>n\<rightarrow> (\<lfloor>v2\<rfloor>\<^sub>e, s2)" and
v: "v=eval_binop binop v1 v2"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e1 P valid_A conf_s0 eval_e1 wt_e1 da_e1
obtain Q: "Q \<lfloor>v1\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
from Q have Q': "\<And> v. (Q\<leftarrow>In1 v1) v s1 Z"
- by simp
+ by simp
have "(\<lambda>Val:v2:. R\<leftarrow>In1 (eval_binop binop v1 v2)) \<lfloor>v2\<rfloor>\<^sub>e s2 Z"
proof (cases "normal s1")
- case True
- from eval_e1 wt_e1 da_e1 conf_s0 wf
- have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T"
- by (rule evaln_type_sound [elim_format]) (insert True,simp)
- from eval_e1
- have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
- by (rule evaln_eval)
- from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
- obtain E2 where
- da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1))
+ case True
+ from eval_e1 wt_e1 da_e1 conf_s0 wf
+ have conf_v1: "G,store s1\<turnstile>v1\<Colon>\<preceq>e1T"
+ by (rule evaln_type_sound [elim_format]) (insert True,simp)
+ from eval_e1
+ have "G\<turnstile>s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1"
+ by (rule evaln_eval)
+ from da wt_e1 wt_e2 wt_binop conf_s0 True this conf_v1 wf
+ obtain E2 where
+ da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1))
\<guillemotright>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<guillemotright> E2"
- by (rule da_e2_BinOp [elim_format]) iprover
- from wt_e2 wt_Skip obtain T2
- where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ by (rule da_e2_BinOp [elim_format]) iprover
+ from wt_e2 wt_Skip obtain T2
+ where "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile>(if need_second_arg binop v1 then \<langle>e2\<rangle>\<^sub>e else \<langle>Skip\<rangle>\<^sub>s)\<Colon>T2"
- by (cases "need_second_arg binop v1") auto
- note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2 this da_e2]
- (* chaining Q', without extra OF causes unification error *)
- thus ?thesis
- by (rule ve)
+ by (cases "need_second_arg binop v1") auto
+ note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2 this da_e2]
+ (* chaining Q', without extra OF causes unification error *)
+ thus ?thesis
+ by (rule ve)
next
- case False
- note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
- with False show ?thesis
- by iprover
+ case False
+ note ve=validE [OF valid_e2,OF Q' valid_A conf_s1 eval_e2]
+ with False show ?thesis
+ by iprover
qed
with v have "R \<lfloor>v\<rfloor>\<^sub>e s2 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -969,7 +969,7 @@
show "P \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from eval have "s1=s0" and "v=val_this (store s0)"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ using normal_s0 by (auto elim: evaln_elim_cases)
with P conf_s0 show ?thesis by simp
qed
qed
@@ -989,23 +989,23 @@
show "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from wt obtain
- wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T"
- by cases simp
+ wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=T"
+ by cases simp
from da obtain V where
- da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
- by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
+ da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
+ by (cases "\<exists> n. var=LVar n") (insert da.LVar,auto elim!: da_elim_cases)
from eval obtain w upd where
- eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(v, upd)\<midarrow>n\<rightarrow> s1"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_var P valid_A conf_s0 eval_var wt_var da_var
have "(\<lambda>Var:(v, f):. Q\<leftarrow>In1 v) \<lfloor>(v, upd)\<rfloor>\<^sub>v s1 Z"
- by (rule validE)
+ by (rule validE)
then have "Q \<lfloor>v\<rfloor>\<^sub>e s1 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s1\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1028,96 +1028,96 @@
show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z \<and> s3\<Colon>\<preceq>(G, L)"
proof -
from wt obtain varT where
- wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=varT" and
- wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
- by cases simp
+ wt_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>var\<Colon>=varT" and
+ wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
+ by cases simp
from eval obtain w upd s1 s2 where
- eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
+ eval_var: "G\<turnstile>s0 \<midarrow>var=\<succ>(w, upd)\<midarrow>n\<rightarrow> s1" and
eval_e: "G\<turnstile>s1 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s2" and
- s3: "s3=assign upd v s2"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ s3: "s3=assign upd v s2"
+ using normal_s0 by (auto elim: evaln_elim_cases)
have "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
proof (cases "\<exists> vn. var = LVar vn")
- case False
- with da obtain V where
- da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ case False
+ with da obtain V where
+ da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V" and
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp+
- from valid_var P valid_A conf_s0 eval_var wt_var da_var
- obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
- hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
- by simp
- have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
- proof (cases "normal s1")
- case True
- obtain E' where
- da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
- proof -
- from eval_var wt_var da_var wf True
- have "nrm V \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approx_evalnE) iprover
- with da_e show thesis
- by (rule da_weakenE) (rule that)
- qed
- note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
- show ?thesis
- by (rule ve)
- next
- case False
- note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
- with False show ?thesis
- by iprover
- qed
- with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
- by simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> nrm V \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp+
+ from valid_var P valid_A conf_s0 eval_var wt_var da_var
+ obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+ by (rule validE)
+ hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
+ by simp
+ have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+ proof (cases "normal s1")
+ case True
+ obtain E' where
+ da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
+ proof -
+ from eval_var wt_var da_var wf True
+ have "nrm V \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approx_evalnE) iprover
+ with da_e show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
+ show ?thesis
+ by (rule ve)
+ next
+ case False
+ note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
+ with False show ?thesis
+ by iprover
+ qed
+ with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+ by simp
next
- case True
- then obtain vn where
- vn: "var = LVar vn"
- by auto
- with da obtain E where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp+
- from da.LVar vn obtain V where
- da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ case True
+ then obtain vn where
+ vn: "var = LVar vn"
+ by auto
+ with da obtain E where
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp+
+ from da.LVar vn obtain V where
+ da_var: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>var\<rangle>\<^sub>v\<guillemotright> V"
- by auto
- from valid_var P valid_A conf_s0 eval_var wt_var da_var
- obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
- hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
- by simp
- have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
- proof (cases "normal s1")
- case True
- obtain E' where
- da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ by auto
+ from valid_var P valid_A conf_s0 eval_var wt_var da_var
+ obtain Q: "Q \<lfloor>(w,upd)\<rfloor>\<^sub>v s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
+ by (rule validE)
+ hence Q': "\<And> v. (Q\<leftarrow>In2 (w,upd)) v s1 Z"
+ by simp
+ have "(\<lambda>Val:v:. assign (snd (w,upd)) v .; R) \<lfloor>v\<rfloor>\<^sub>e s2 Z"
+ proof (cases "normal s1")
+ case True
+ obtain E' where
+ da_e': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E'"
- proof -
- from eval_var
- have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
- by (rule dom_locals_evaln_mono_elim)
- with da_e show thesis
- by (rule da_weakenE) (rule that)
- qed
- note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
- show ?thesis
- by (rule ve)
- next
- case False
- note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
- with False show ?thesis
- by iprover
- qed
- with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
- by simp
+ proof -
+ from eval_var
+ have "dom (locals (store s0)) \<subseteq> dom (locals (store (s1)))"
+ by (rule dom_locals_evaln_mono_elim)
+ with da_e show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e wt_e da_e']
+ show ?thesis
+ by (rule ve)
+ next
+ case False
+ note ve=validE [OF valid_e,OF Q' valid_A conf_s1 eval_e]
+ with False show ?thesis
+ by iprover
+ qed
+ with s3 show "R \<lfloor>v\<rfloor>\<^sub>e s3 Z"
+ by simp
qed
moreover
from eval wt da conf_s0 wf
have "s3\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1139,75 +1139,75 @@
show "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
proof -
from wt obtain T1 T2 where
- wt_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
- wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T1" and
- wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-T2"
- by cases simp
+ wt_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e0\<Colon>-PrimT Boolean" and
+ wt_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e1\<Colon>-T1" and
+ wt_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e2\<Colon>-T2"
+ by cases simp
from da obtain E0 E1 E2 where
da_e0: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e0\<rangle>\<^sub>e\<guillemotright> E0" and
da_e1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile>(dom (locals (store s0)) \<union> assigns_if True e0)\<guillemotright>\<langle>e1\<rangle>\<^sub>e\<guillemotright> E1" and
da_e2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile>(dom (locals (store s0)) \<union> assigns_if False e0)\<guillemotright>\<langle>e2\<rangle>\<^sub>e\<guillemotright> E2"
- by cases simp+
+ by cases simp+
from eval obtain b s1 where
- eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
+ eval_e0: "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<midarrow>n\<rightarrow> s1" and
eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<midarrow>n\<rightarrow> s2"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e0 P valid_A conf_s0 eval_e0 wt_e0 da_e0
obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
hence P': "\<And> v. (P'\<leftarrow>=(the_Bool b)) v s1 Z"
- by (cases "normal s1") auto
+ by (cases "normal s1") auto
have "Q \<lfloor>v\<rfloor>\<^sub>e s2 Z"
proof (cases "normal s1")
- case True
- note normal_s1=this
- from wt_e1 wt_e2 obtain T' where
- wt_then_else:
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
- by (cases "the_Bool b") simp+
- have s0_s1: "dom (locals (store s0))
+ case True
+ note normal_s1=this
+ from wt_e1 wt_e2 obtain T' where
+ wt_then_else:
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then e1 else e2)\<Colon>-T'"
+ by (cases "the_Bool b") simp+
+ have s0_s1: "dom (locals (store s0))
\<union> assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
- proof -
- from eval_e0
- have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
- by (rule evaln_eval)
- hence
- "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
- by (rule dom_locals_eval_mono_elim)
+ proof -
+ from eval_e0
+ have eval_e0': "G\<turnstile>s0 \<midarrow>e0-\<succ>b\<rightarrow> s1"
+ by (rule evaln_eval)
+ hence
+ "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+ by (rule dom_locals_eval_mono_elim)
moreover
- from eval_e0' True wt_e0
- have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
- by (rule assigns_if_good_approx')
- ultimately show ?thesis by (rule Un_least)
- qed
- obtain E' where
- da_then_else:
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ from eval_e0' True wt_e0
+ have "assigns_if (the_Bool b) e0 \<subseteq> dom (locals (store s1))"
+ by (rule assigns_if_good_approx')
+ ultimately show ?thesis by (rule Un_least)
+ qed
+ obtain E' where
+ da_then_else:
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then e1 else e2\<rangle>\<^sub>e\<guillemotright> E'"
- proof (cases "the_Bool b")
- case True
- with that da_e1 s0_s1 show ?thesis
- by simp (erule da_weakenE,auto)
- next
- case False
- with that da_e2 s0_s1 show ?thesis
- by simp (erule da_weakenE,auto)
- qed
- with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
- show ?thesis
- by (rule validE)
+ proof (cases "the_Bool b")
+ case True
+ with that da_e1 s0_s1 show ?thesis
+ by simp (erule da_weakenE,auto)
+ next
+ case False
+ with that da_e2 s0_s1 show ?thesis
+ by simp (erule da_weakenE,auto)
+ qed
+ with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
+ show ?thesis
+ by (rule validE)
next
- case False
- with valid_then_else P' valid_A conf_s1 eval_then_else
- show ?thesis
- by (cases rule: validE) iprover+
+ case False
+ with valid_then_else P' valid_A conf_s1 eval_then_else
+ show ?thesis
+ by (cases rule: validE) iprover+
qed
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1248,93 +1248,93 @@
mode: "mode = invmode statM e" and
T: "T =(resTy statM)" and
eq_accC_accC': "accC=accC'"
- by cases fastsimp+
+ by cases fastsimp+
from da obtain C where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
- da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> C" and
+ da_args: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm C \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E"
+ by cases simp
from eval eq_accC_accC' obtain a s1 vs s2 s3 s3' s4 invDeclC where
- evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+ evaln_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
evaln_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
- invDeclC: "invDeclC = invocation_declclass
+ invDeclC: "invDeclC = invocation_declclass
G mode (store s2) a statT \<lparr>name=mn,parTs=pTs'\<rparr>" and
s3: "s3 = init_lvars G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> mode a vs s2" and
check: "s3' = check_method_access G
accC' statT mode \<lparr>name = mn, parTs = pTs'\<rparr> a s3" and
- evaln_methd:
+ evaln_methd:
"G\<turnstile>s3' \<midarrow>Methd invDeclC \<lparr>name=mn,parTs=pTs'\<rparr>-\<succ>v\<midarrow>n\<rightarrow> s4" and
s5: "s5=(set_lvars (locals (store s2))) s4"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ using normal_s0 by (auto elim: evaln_elim_cases)
from evaln_e
have eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<rightarrow> s1"
- by (rule evaln_eval)
+ by (rule evaln_eval)
from eval_e _ wt_e wf
have s1_no_return: "abrupt s1 \<noteq> Some (Jump Ret)"
- by (rule eval_expression_no_jump
+ by (rule eval_expression_no_jump
[where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
- (insert normal_s0,auto)
+ (insert normal_s0,auto)
from valid_e P valid_A conf_s0 evaln_e wt_e da_e
obtain "Q \<lfloor>a\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
hence Q: "\<And> v. (Q\<leftarrow>In1 a) v s1 Z"
- by simp
+ by simp
obtain
- R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and
- conf_s2: "s2\<Colon>\<preceq>(G,L)" and
- s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
+ R: "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" and
+ conf_s2: "s2\<Colon>\<preceq>(G,L)" and
+ s2_no_return: "abrupt s2 \<noteq> Some (Jump Ret)"
proof (cases "normal s1")
- case True
- obtain E' where
- da_args':
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
- proof -
- from evaln_e wt_e da_e wf True
- have "nrm C \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approx_evalnE) iprover
- with da_args show thesis
- by (rule da_weakenE) (rule that)
- qed
- with valid_args Q valid_A conf_s1 evaln_args wt_args
- obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
- by (rule validE)
- moreover
- from evaln_args
- have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
- by (rule evaln_eval)
- from this s1_no_return wt_args wf
- have "abrupt s2 \<noteq> Some (Jump Ret)"
- by (rule eval_expression_list_no_jump
+ case True
+ obtain E' where
+ da_args':
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
+ proof -
+ from evaln_e wt_e da_e wf True
+ have "nrm C \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approx_evalnE) iprover
+ with da_args show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ with valid_args Q valid_A conf_s1 evaln_args wt_args
+ obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
+ by (rule validE)
+ moreover
+ from evaln_args
+ have e: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+ by (rule evaln_eval)
+ from this s1_no_return wt_args wf
+ have "abrupt s2 \<noteq> Some (Jump Ret)"
+ by (rule eval_expression_list_no_jump
[where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
- ultimately show ?thesis ..
+ ultimately show ?thesis ..
next
- case False
- with valid_args Q valid_A conf_s1 evaln_args
- obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
- by (cases rule: validE) iprover+
- moreover
- from False evaln_args have "s2=s1"
- by auto
- with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
- by simp
- ultimately show ?thesis ..
+ case False
+ with valid_args Q valid_A conf_s1 evaln_args
+ obtain "(R a) \<lfloor>vs\<rfloor>\<^sub>l s2 Z" "s2\<Colon>\<preceq>(G,L)"
+ by (cases rule: validE) iprover+
+ moreover
+ from False evaln_args have "s2=s1"
+ by auto
+ with s1_no_return have "abrupt s2 \<noteq> Some (Jump Ret)"
+ by simp
+ ultimately show ?thesis ..
qed
obtain invC where
- invC: "invC = invocation_class mode (store s2) a statT"
- by simp
+ invC: "invC = invocation_class mode (store s2) a statT"
+ by simp
with s3
have invC': "invC = (invocation_class mode (store s3) a statT)"
- by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
+ by (cases s2,cases mode) (auto simp add: init_lvars_def2 )
obtain l where
- l: "l = locals (store s2)"
- by simp
+ l: "l = locals (store s2)"
+ by simp
from eval wt da conf_s0 wf
have conf_s5: "s5\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
let "PROP ?R" = "\<And> v.
(R a\<leftarrow>In3 vs \<and>.
(\<lambda>s. invDeclC = invocation_declclass G mode (store s) a statT
@@ -1345,128 +1345,128 @@
(\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)
) v s3' Z"
{
- assume abrupt_s3: "\<not> normal s3"
- have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
- proof -
- from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
- by (auto simp add: check_method_access_def Let_def)
- with R s3 invDeclC invC l abrupt_s3
- have R': "PROP ?R"
- by auto
- have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
- (* we need an arbirary environment (here empty) that s2' conforms to
+ assume abrupt_s3: "\<not> normal s3"
+ have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
+ proof -
+ from abrupt_s3 check have eq_s3'_s3: "s3'=s3"
+ by (auto simp add: check_method_access_def Let_def)
+ with R s3 invDeclC invC l abrupt_s3
+ have R': "PROP ?R"
+ by auto
+ have conf_s3': "s3'\<Colon>\<preceq>(G, empty)"
+ (* we need an arbirary environment (here empty) that s2' conforms to
to apply validE *)
- proof -
- from s2_no_return s3
- have "abrupt s3 \<noteq> Some (Jump Ret)"
- by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
- moreover
- obtain abr2 str2 where s2: "s2=(abr2,str2)"
- by (cases s2)
- from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
- by (auto simp add: init_lvars_def2 split: split_if_asm)
- ultimately show ?thesis
- using s3 s2 eq_s3'_s3
- apply (simp add: init_lvars_def2)
- apply (rule conforms_set_locals [OF _ wlconf_empty])
- by auto
- qed
- from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
- have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
- by (cases rule: validE) simp+
- with s5 l show ?thesis
- by simp
- qed
+ proof -
+ from s2_no_return s3
+ have "abrupt s3 \<noteq> Some (Jump Ret)"
+ by (cases s2) (auto simp add: init_lvars_def2 split: split_if_asm)
+ moreover
+ obtain abr2 str2 where s2: "s2=(abr2,str2)"
+ by (cases s2)
+ from s3 s2 conf_s2 have "(abrupt s3,str2)\<Colon>\<preceq>(G, L)"
+ by (auto simp add: init_lvars_def2 split: split_if_asm)
+ ultimately show ?thesis
+ using s3 s2 eq_s3'_s3
+ apply (simp add: init_lvars_def2)
+ apply (rule conforms_set_locals [OF _ wlconf_empty])
+ by auto
+ qed
+ from valid_methd R' valid_A conf_s3' evaln_methd abrupt_s3 eq_s3'_s3
+ have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+ by (cases rule: validE) simp+
+ with s5 l show ?thesis
+ by simp
+ qed
} note abrupt_s3_lemma = this
have "S \<lfloor>v\<rfloor>\<^sub>e s5 Z"
proof (cases "normal s2")
- case False
- with s3 have abrupt_s3: "\<not> normal s3"
- by (cases s2) (simp add: init_lvars_def2)
- thus ?thesis
- by (rule abrupt_s3_lemma)
+ case False
+ with s3 have abrupt_s3: "\<not> normal s3"
+ by (cases s2) (simp add: init_lvars_def2)
+ thus ?thesis
+ by (rule abrupt_s3_lemma)
next
- case True
- note normal_s2 = this
- with evaln_args
- have normal_s1: "normal s1"
- by (rule evaln_no_abrupt)
- obtain E' where
- da_args':
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
- proof -
- from evaln_e wt_e da_e wf normal_s1
- have "nrm C \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approx_evalnE) iprover
- with da_args show thesis
- by (rule da_weakenE) (rule that)
- qed
- from evaln_args
- have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
- by (rule evaln_eval)
- from evaln_e wt_e da_e conf_s0 wf
- have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
- by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
- with normal_s1 normal_s2 eval_args
- have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
- by (auto dest: eval_gext intro: conf_gext)
- from evaln_args wt_args da_args' conf_s1 wf
- have conf_args: "list_all2 (conf G (store s2)) vs pTs"
- by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
- from statM
- obtain
- statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>"
- and
- pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
- by (blast dest: max_spec2mheads)
- show ?thesis
- proof (cases "normal s3")
- case False
- thus ?thesis
- by (rule abrupt_s3_lemma)
- next
- case True
- note normal_s3 = this
- with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
- by (cases s2) (auto simp add: init_lvars_def2)
- from conf_s2 conf_a_s2 wf notNull invC
- have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
- by (cases s2) (auto intro: DynT_propI)
+ case True
+ note normal_s2 = this
+ with evaln_args
+ have normal_s1: "normal s1"
+ by (rule evaln_no_abrupt)
+ obtain E' where
+ da_args':
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>args\<rangle>\<^sub>l\<guillemotright> E'"
+ proof -
+ from evaln_e wt_e da_e wf normal_s1
+ have "nrm C \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approx_evalnE) iprover
+ with da_args show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ from evaln_args
+ have eval_args: "G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2"
+ by (rule evaln_eval)
+ from evaln_e wt_e da_e conf_s0 wf
+ have conf_a: "G, store s1\<turnstile>a\<Colon>\<preceq>RefT statT"
+ by (rule evaln_type_sound [elim_format]) (insert normal_s1,simp)
+ with normal_s1 normal_s2 eval_args
+ have conf_a_s2: "G, store s2\<turnstile>a\<Colon>\<preceq>RefT statT"
+ by (auto dest: eval_gext intro: conf_gext)
+ from evaln_args wt_args da_args' conf_s1 wf
+ have conf_args: "list_all2 (conf G (store s2)) vs pTs"
+ by (rule evaln_type_sound [elim_format]) (insert normal_s2,simp)
+ from statM
+ obtain
+ statM': "(statDeclT,statM)\<in>mheads G accC statT \<lparr>name=mn,parTs=pTs'\<rparr>"
+ and
+ pTs_widen: "G\<turnstile>pTs[\<preceq>]pTs'"
+ by (blast dest: max_spec2mheads)
+ show ?thesis
+ proof (cases "normal s3")
+ case False
+ thus ?thesis
+ by (rule abrupt_s3_lemma)
+ next
+ case True
+ note normal_s3 = this
+ with s3 have notNull: "mode = IntVir \<longrightarrow> a \<noteq> Null"
+ by (cases s2) (auto simp add: init_lvars_def2)
+ from conf_s2 conf_a_s2 wf notNull invC
+ have dynT_prop: "G\<turnstile>mode\<rightarrow>invC\<preceq>statT"
+ by (cases s2) (auto intro: DynT_propI)
- with wt_e statM' invC mode wf
- obtain dynM where
+ with wt_e statM' invC mode wf
+ obtain dynM where
dynM: "dynlookup G statT invC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
acc_dynM: "G \<turnstile>Methd \<lparr>name=mn,parTs=pTs'\<rparr> dynM
in invC dyn_accessible_from accC"
- by (force dest!: call_access_ok)
- with invC' check eq_accC_accC'
- have eq_s3'_s3: "s3'=s3"
- by (auto simp add: check_method_access_def Let_def)
-
- with dynT_prop R s3 invDeclC invC l
- have R': "PROP ?R"
- by auto
+ by (force dest!: call_access_ok)
+ with invC' check eq_accC_accC'
+ have eq_s3'_s3: "s3'=s3"
+ by (auto simp add: check_method_access_def Let_def)
+
+ with dynT_prop R s3 invDeclC invC l
+ have R': "PROP ?R"
+ by auto
- from dynT_prop wf wt_e statM' mode invC invDeclC dynM
- obtain
+ from dynT_prop wf wt_e statM' mode invC invDeclC dynM
+ obtain
dynM: "dynlookup G statT invC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
- wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
- dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
+ wf_dynM: "wf_mdecl G invDeclC (\<lparr>name=mn,parTs=pTs'\<rparr>,mthd dynM)" and
+ dynM': "methd G invDeclC \<lparr>name=mn,parTs=pTs'\<rparr> = Some dynM" and
iscls_invDeclC: "is_class G invDeclC" and
- invDeclC': "invDeclC = declclass dynM" and
- invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
- resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
- is_static_eq: "is_static dynM = is_static statM" and
- involved_classes_prop:
+ invDeclC': "invDeclC = declclass dynM" and
+ invC_widen: "G\<turnstile>invC\<preceq>\<^sub>C invDeclC" and
+ resTy_widen: "G\<turnstile>resTy dynM\<preceq>resTy statM" and
+ is_static_eq: "is_static dynM = is_static statM" and
+ involved_classes_prop:
"(if invmode statM e = IntVir
then \<forall>statC. statT = ClassT statC \<longrightarrow> G\<turnstile>invC\<preceq>\<^sub>C statC
else ((\<exists>statC. statT = ClassT statC \<and> G\<turnstile>statC\<preceq>\<^sub>C invDeclC) \<or>
(\<forall>statC. statT \<noteq> ClassT statC \<and> invDeclC = Object)) \<and>
statDeclT = ClassT invDeclC)"
- by (cases rule: DynT_mheadsE) simp
- obtain L' where
- L':"L'=(\<lambda> k.
+ by (cases rule: DynT_mheadsE) simp
+ obtain L' where
+ L':"L'=(\<lambda> k.
(case k of
EName e
\<Rightarrow> (case e of
@@ -1476,121 +1476,121 @@
| Res \<Rightarrow> Some (resTy dynM))
| This \<Rightarrow> if is_static statM
then None else Some (Class invDeclC)))"
- by simp
- from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
+ by simp
+ from wf_dynM [THEN wf_mdeclD1, THEN conjunct1] normal_s2 conf_s2 wt_e
wf eval_args conf_a mode notNull wf_dynM involved_classes_prop
- have conf_s3: "s3\<Colon>\<preceq>(G,L')"
- apply -
+ have conf_s3: "s3\<Colon>\<preceq>(G,L')"
+ apply -
(* FIXME confomrs_init_lvars should be
adjusted to be more directy applicable *)
- apply (drule conforms_init_lvars [of G invDeclC
+ apply (drule conforms_init_lvars [of G invDeclC
"\<lparr>name=mn,parTs=pTs'\<rparr>" dynM "store s2" vs pTs "abrupt s2"
L statT invC a "(statDeclT,statM)" e])
- apply (rule wf)
- apply (rule conf_args)
- apply (simp add: pTs_widen)
- apply (cases s2,simp)
- apply (rule dynM')
- apply (force dest: ty_expr_is_type)
- apply (rule invC_widen)
- apply (force intro: conf_gext dest: eval_gext)
- apply simp
- apply simp
- apply (simp add: invC)
- apply (simp add: invDeclC)
- apply (simp add: normal_s2)
- apply (cases s2, simp add: L' init_lvars_def2 s3
- cong add: lname.case_cong ename.case_cong)
- done
- with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
- from is_static_eq wf_dynM L'
- obtain mthdT where
- "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+ apply (rule wf)
+ apply (rule conf_args)
+ apply (simp add: pTs_widen)
+ apply (cases s2,simp)
+ apply (rule dynM')
+ apply (force dest: ty_expr_is_type)
+ apply (rule invC_widen)
+ apply (force intro: conf_gext dest: eval_gext)
+ apply simp
+ apply simp
+ apply (simp add: invC)
+ apply (simp add: invDeclC)
+ apply (simp add: normal_s2)
+ apply (cases s2, simp add: L' init_lvars_def2 s3
+ cong add: lname.case_cong ename.case_cong)
+ done
+ with eq_s3'_s3 have conf_s3': "s3'\<Colon>\<preceq>(G,L')" by simp
+ from is_static_eq wf_dynM L'
+ obtain mthdT where
+ "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
\<turnstile>Body invDeclC (stmt (mbody (mthd dynM)))\<Colon>-mthdT" and
- mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
- by - (drule wf_mdecl_bodyD,
+ mthdT_widen: "G\<turnstile>mthdT\<preceq>resTy dynM"
+ by - (drule wf_mdecl_bodyD,
auto simp add: callee_lcl_def
cong add: lname.case_cong ename.case_cong)
- with dynM' iscls_invDeclC invDeclC'
- have
- wt_methd:
- "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+ with dynM' iscls_invDeclC invDeclC'
+ have
+ wt_methd:
+ "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
\<turnstile>(Methd invDeclC \<lparr>name = mn, parTs = pTs'\<rparr>)\<Colon>-mthdT"
- by (auto intro: wt.Methd)
- obtain M where
- da_methd:
- "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
- \<turnstile> dom (locals (store s3'))
+ by (auto intro: wt.Methd)
+ obtain M where
+ da_methd:
+ "\<lparr>prg=G,cls=invDeclC,lcl=L'\<rparr>
+ \<turnstile> dom (locals (store s3'))
\<guillemotright>\<langle>Methd invDeclC \<lparr>name=mn,parTs=pTs'\<rparr>\<rangle>\<^sub>e\<guillemotright> M"
- proof -
- from wf_dynM
- obtain M' where
- da_body:
- "\<lparr>prg=G, cls=invDeclC
+ proof -
+ from wf_dynM
+ obtain M' where
+ da_body:
+ "\<lparr>prg=G, cls=invDeclC
,lcl=callee_lcl invDeclC \<lparr>name = mn, parTs = pTs'\<rparr> (mthd dynM)
\<rparr> \<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'" and
res: "Result \<in> nrm M'"
- by (rule wf_mdeclE) iprover
- from da_body is_static_eq L' have
- "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
+ by (rule wf_mdeclE) iprover
+ from da_body is_static_eq L' have
+ "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
\<turnstile> parameters (mthd dynM) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M'"
- by (simp add: callee_lcl_def
+ by (simp add: callee_lcl_def
cong add: lname.case_cong ename.case_cong)
- moreover have "parameters (mthd dynM) \<subseteq> dom (locals (store s3'))"
- proof -
- from is_static_eq
- have "(invmode (mthd dynM) e) = (invmode statM e)"
- by (simp add: invmode_def)
- moreover
- have "length (pars (mthd dynM)) = length vs"
- proof -
- from normal_s2 conf_args
- have "length vs = length pTs"
- by (simp add: list_all2_def)
- also from pTs_widen
- have "\<dots> = length pTs'"
- by (simp add: widens_def list_all2_def)
- also from wf_dynM
- have "\<dots> = length (pars (mthd dynM))"
- by (simp add: wf_mdecl_def wf_mhead_def)
- finally show ?thesis ..
- qed
- moreover note s3 dynM' is_static_eq normal_s2 mode
- ultimately
- have "parameters (mthd dynM) = dom (locals (store s3))"
- using dom_locals_init_lvars
+ moreover have "parameters (mthd dynM) \<subseteq> dom (locals (store s3'))"
+ proof -
+ from is_static_eq
+ have "(invmode (mthd dynM) e) = (invmode statM e)"
+ by (simp add: invmode_def)
+ moreover
+ have "length (pars (mthd dynM)) = length vs"
+ proof -
+ from normal_s2 conf_args
+ have "length vs = length pTs"
+ by (simp add: list_all2_def)
+ also from pTs_widen
+ have "\<dots> = length pTs'"
+ by (simp add: widens_def list_all2_def)
+ also from wf_dynM
+ have "\<dots> = length (pars (mthd dynM))"
+ by (simp add: wf_mdecl_def wf_mhead_def)
+ finally show ?thesis ..
+ qed
+ moreover note s3 dynM' is_static_eq normal_s2 mode
+ ultimately
+ have "parameters (mthd dynM) = dom (locals (store s3))"
+ using dom_locals_init_lvars
[of "mthd dynM" G invDeclC "\<lparr>name=mn,parTs=pTs'\<rparr>" vs e a s2]
- by simp
- thus ?thesis using eq_s3'_s3 by simp
- qed
- ultimately obtain M2 where
- da:
- "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
+ by simp
+ thus ?thesis using eq_s3'_s3 by simp
+ qed
+ ultimately obtain M2 where
+ da:
+ "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr>
\<turnstile> dom (locals (store s3')) \<guillemotright>\<langle>stmt (mbody (mthd dynM))\<rangle>\<guillemotright> M2" and
M2: "nrm M' \<subseteq> nrm M2"
- by (rule da_weakenE)
- from res M2 have "Result \<in> nrm M2"
- by blast
- moreover from wf_dynM
- have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
- by (rule wf_mdeclE)
- ultimately
- obtain M3 where
- "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3'))
+ by (rule da_weakenE)
+ from res M2 have "Result \<in> nrm M2"
+ by blast
+ moreover from wf_dynM
+ have "jumpNestingOkS {Ret} (stmt (mbody (mthd dynM)))"
+ by (rule wf_mdeclE)
+ ultimately
+ obtain M3 where
+ "\<lparr>prg=G, cls=invDeclC,lcl=L'\<rparr> \<turnstile> dom (locals (store s3'))
\<guillemotright>\<langle>Body (declclass dynM) (stmt (mbody (mthd dynM)))\<rangle>\<guillemotright> M3"
- using da
- by (iprover intro: da.Body assigned.select_convs)
- from _ this [simplified]
- show thesis
- by (rule da.Methd [simplified,elim_format])
- (auto intro: dynM' that)
- qed
- from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
- have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
- by (cases rule: validE) iprover+
- with s5 l show ?thesis
- by simp
- qed
+ using da
+ by (iprover intro: da.Body assigned.select_convs)
+ from _ this [simplified]
+ show thesis
+ by (rule da.Methd [simplified,elim_format])
+ (auto intro: dynM' that)
+ qed
+ from valid_methd R' valid_A conf_s3' evaln_methd wt_methd da_methd
+ have "(set_lvars l .; S) \<lfloor>v\<rfloor>\<^sub>e s4 Z"
+ by (cases rule: validE) iprover+
+ with s5 l show ?thesis
+ by simp
+ qed
qed
with conf_s5 show ?thesis by iprover
qed
@@ -1618,61 +1618,61 @@
show "R \<lfloor>v\<rfloor>\<^sub>e s4 Z \<and> s4\<Colon>\<preceq>(G, L)"
proof -
from wt obtain
- iscls_D: "is_class G D" and
+ iscls_D: "is_class G D" and
wt_init: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>Init D\<Colon>\<surd>" and
wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>"
- by cases auto
+ by cases auto
obtain I where
- da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
- by (auto intro: da_Init [simplified] assigned.select_convs)
+ da_init:"\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>Init D\<rangle>\<^sub>s\<guillemotright> I"
+ by (auto intro: da_Init [simplified] assigned.select_convs)
from da obtain C where
- da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
- jmpOk: "jumpNestingOkS {Ret} c"
- by cases simp
+ da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C" and
+ jmpOk: "jumpNestingOkS {Ret} c"
+ by cases simp
from eval obtain s1 s2 s3 where
- eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
+ eval_init: "G\<turnstile>s0 \<midarrow>Init D\<midarrow>n\<rightarrow> s1" and
eval_c: "G\<turnstile>s1 \<midarrow>c\<midarrow>n\<rightarrow> s2" and
- v: "v = the (locals (store s2) Result)" and
+ v: "v = the (locals (store s2) Result)" and
s3: "s3 =(if \<exists>l. abrupt s2 = Some (Jump (Break l)) \<or>
abrupt s2 = Some (Jump (Cont l))
then abupd (\<lambda>x. Some (Error CrossMethodJump)) s2 else s2)"and
s4: "s4 = abupd (absorb Ret) s3"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
obtain C' where
- da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
+ da_c': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s1)))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> C'"
proof -
- from eval_init
- have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
- by (rule dom_locals_evaln_mono_elim)
- with da_c show thesis by (rule da_weakenE) (rule that)
+ from eval_init
+ have "(dom (locals (store s0))) \<subseteq> (dom (locals (store s1)))"
+ by (rule dom_locals_evaln_mono_elim)
+ with da_c show thesis by (rule da_weakenE) (rule that)
qed
from valid_init P valid_A conf_s0 eval_init wt_init da_init
obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
from valid_c Q valid_A conf_s1 eval_c wt_c da_c'
have R: "(\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>In1 (the (locals s Result)))
\<diamondsuit> s2 Z"
- by (rule validE)
+ by (rule validE)
have "s3=s2"
proof -
- from eval_init [THEN evaln_eval] wf
- have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
- by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
+ from eval_init [THEN evaln_eval] wf
+ have s1_no_jmp: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+ by - (rule eval_statement_no_jump [OF _ _ _ wt_init],
insert normal_s0,auto)
- from eval_c [THEN evaln_eval] _ wt_c wf
- have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
- by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
- moreover note s3
- ultimately show ?thesis
- by (force split: split_if)
+ from eval_c [THEN evaln_eval] _ wt_c wf
+ have "\<And> j. abrupt s2 = Some (Jump j) \<Longrightarrow> j=Ret"
+ by (rule jumpNestingOk_evalE) (auto intro: jmpOk simp add: s1_no_jmp)
+ moreover note s3
+ ultimately show ?thesis
+ by (force split: split_if)
qed
with R v s4
have "R \<lfloor>v\<rfloor>\<^sub>e s4 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s4\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1688,9 +1688,9 @@
show "P \<lfloor>vs\<rfloor>\<^sub>l s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from eval obtain "vs=[]" "s1=s0"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ using normal_s0 by (auto elim: evaln_elim_cases)
with P conf_s0 show ?thesis
- by simp
+ by simp
qed
qed
next
@@ -1711,50 +1711,50 @@
show "R \<lfloor>v\<rfloor>\<^sub>l s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
proof -
from wt obtain eT esT where
- wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT" and
- wt_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>esT"
- by cases simp
+ wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-eT" and
+ wt_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>es\<Colon>\<doteq>esT"
+ by cases simp
from da obtain E1 where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
- da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> (dom (locals (store s0)))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E1" and
+ da_es: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> nrm E1 \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E"
+ by cases simp
from eval obtain s1 ve vs where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
- eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
- v: "v=ve#vs"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>ve\<midarrow>n\<rightarrow> s1" and
+ eval_es: "G\<turnstile>s1 \<midarrow>es\<doteq>\<succ>vs\<midarrow>n\<rightarrow> s2" and
+ v: "v=ve#vs"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e P valid_A conf_s0 eval_e wt_e da_e
obtain Q: "Q \<lfloor>ve\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
from Q have Q': "\<And> v. (Q\<leftarrow>\<lfloor>ve\<rfloor>\<^sub>e) v s1 Z"
- by simp
+ by simp
have "(\<lambda>Vals:vs:. R\<leftarrow>\<lfloor>(ve # vs)\<rfloor>\<^sub>l) \<lfloor>vs\<rfloor>\<^sub>l s2 Z"
proof (cases "normal s1")
- case True
- obtain E' where
- da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
- proof -
- from eval_e wt_e da_e wf True
- have "nrm E1 \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approx_evalnE) iprover
- with da_es show thesis
- by (rule da_weakenE) (rule that)
- qed
- from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
- show ?thesis
- by (rule validE)
+ case True
+ obtain E' where
+ da_es': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>es\<rangle>\<^sub>l\<guillemotright> E'"
+ proof -
+ from eval_e wt_e da_e wf True
+ have "nrm E1 \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approx_evalnE) iprover
+ with da_es show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ from valid_es Q' valid_A conf_s1 eval_es wt_es da_es'
+ show ?thesis
+ by (rule validE)
next
- case False
- with valid_es Q' valid_A conf_s1 eval_es
- show ?thesis
- by (cases rule: validE) iprover+
+ case False
+ with valid_es Q' valid_A conf_s1 eval_es
+ show ?thesis
+ by (cases rule: validE) iprover+
qed
with v have "R \<lfloor>v\<rfloor>\<^sub>l s2 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1770,9 +1770,9 @@
show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from eval obtain "s1=s0"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
with P conf_s0 show ?thesis
- by simp
+ by simp
qed
qed
next
@@ -1791,17 +1791,17 @@
show "Q \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G, L)"
proof -
from wt obtain eT where
- wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
- by cases simp
+ wt_e: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>e\<Colon>-eT"
+ by cases simp
from da obtain E where
- da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright>E"
+ by cases simp
from eval obtain v where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>v\<midarrow>n\<rightarrow> s1"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_e P valid_A conf_s0 eval_e wt_e da_e
obtain Q: "(Q\<leftarrow>\<diamondsuit>) \<lfloor>v\<rfloor>\<^sub>e s1 Z" and "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
thus ?thesis by simp
qed
qed
@@ -1821,24 +1821,24 @@
show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G, L)"
proof -
from wt obtain
- wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
- by cases simp
+ wt_c: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c\<Colon>\<surd>"
+ by cases simp
from da obtain E where
- da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
- by cases simp
+ da_c: "\<lparr>prg=G,cls=accC, lcl=L\<rparr>\<turnstile>dom (locals (store s0))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright>E"
+ by cases simp
from eval obtain s1 where
- eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
- s2: "s2 = abupd (absorb l) s1"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ eval_c: "G\<turnstile>s0 \<midarrow>c\<midarrow>n\<rightarrow> s1" and
+ s2: "s2 = abupd (absorb l) s1"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from valid_c P valid_A conf_s0 eval_c wt_c da_c
obtain Q: "(abupd (absorb l) .; Q) \<diamondsuit> s1 Z"
- by (rule validE)
+ by (rule validE)
with s2 have "Q \<diamondsuit> s2 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1859,45 +1859,45 @@
show "R \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
proof -
from eval obtain s1 where
- eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
- eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ eval_c1: "G\<turnstile>s0 \<midarrow>c1 \<midarrow>n\<rightarrow> s1" and
+ eval_c2: "G\<turnstile>s1 \<midarrow>c2 \<midarrow>n\<rightarrow> s2"
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from wt obtain
- wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+ wt_c1: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
wt_c2: "\<lparr>prg = G, cls = accC, lcl = L\<rparr>\<turnstile>c2\<Colon>\<surd>"
- by cases simp
+ by cases simp
from da obtain C1 C2 where
- da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
- da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
- by cases simp
+ da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+ da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>nrm C1 \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
+ by cases simp
from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
have "R \<diamondsuit> s2 Z"
proof (cases "normal s1")
- case True
- obtain C2' where
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
- proof -
- from eval_c1 wt_c1 da_c1 wf True
- have "nrm C1 \<subseteq> dom (locals (store s1))"
- by (cases rule: da_good_approx_evalnE) iprover
- with da_c2 show thesis
- by (rule da_weakenE) (rule that)
- qed
- with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2
- show ?thesis
- by (rule validE)
+ case True
+ obtain C2' where
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s1)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
+ proof -
+ from eval_c1 wt_c1 da_c1 wf True
+ have "nrm C1 \<subseteq> dom (locals (store s1))"
+ by (cases rule: da_good_approx_evalnE) iprover
+ with da_c2 show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ with valid_c2 Q valid_A conf_s1 eval_c2 wt_c2
+ show ?thesis
+ by (rule validE)
next
- case False
- from valid_c2 Q valid_A conf_s1 eval_c2 False
- show ?thesis
- by (cases rule: validE) iprover+
+ case False
+ from valid_c2 Q valid_A conf_s1 eval_c2 False
+ show ?thesis
+ by (cases rule: validE) iprover+
qed
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1920,61 +1920,61 @@
show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
proof -
from eval obtain b s1 where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
- eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>b\<midarrow>n\<rightarrow> s1" and
+ eval_then_else: "G\<turnstile>s1 \<midarrow>(if the_Bool b then c1 else c2)\<midarrow>n\<rightarrow> s2"
+ using normal_s0 by (auto elim: evaln_elim_cases)
from wt obtain
- wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
- wt_then_else: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
- by cases (simp split: split_if)
+ wt_e: "\<lparr>prg=G, cls=accC, lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+ wt_then_else: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(if the_Bool b then c1 else c2)\<Colon>\<surd>"
+ by cases (simp split: split_if)
from da obtain E S where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
- da_then_else:
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
+ da_then_else:
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
(dom (locals (store s0)) \<union> assigns_if (the_Bool b) e)
\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S"
- by cases (cases "the_Bool b",auto)
+ by cases (cases "the_Bool b",auto)
from valid_e P valid_A conf_s0 eval_e wt_e da_e
obtain "P' \<lfloor>b\<rfloor>\<^sub>e s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
hence P': "\<And>v. (P'\<leftarrow>=the_Bool b) v s1 Z"
- by (cases "normal s1") auto
+ by (cases "normal s1") auto
have "Q \<diamondsuit> s2 Z"
proof (cases "normal s1")
- case True
- have s0_s1: "dom (locals (store s0))
+ case True
+ have s0_s1: "dom (locals (store s0))
\<union> assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
- proof -
- from eval_e
- have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
- by (rule evaln_eval)
- hence
- "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
- by (rule dom_locals_eval_mono_elim)
+ proof -
+ from eval_e
+ have eval_e': "G\<turnstile>s0 \<midarrow>e-\<succ>b\<rightarrow> s1"
+ by (rule evaln_eval)
+ hence
+ "dom (locals (store s0)) \<subseteq> dom (locals (store s1))"
+ by (rule dom_locals_eval_mono_elim)
moreover
- from eval_e' True wt_e
- have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
- by (rule assigns_if_good_approx')
- ultimately show ?thesis by (rule Un_least)
- qed
- with da_then_else
- obtain S' where
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ from eval_e' True wt_e
+ have "assigns_if (the_Bool b) e \<subseteq> dom (locals (store s1))"
+ by (rule assigns_if_good_approx')
+ ultimately show ?thesis by (rule Un_least)
+ qed
+ with da_then_else
+ obtain S' where
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile>dom (locals (store s1))\<guillemotright>\<langle>if the_Bool b then c1 else c2\<rangle>\<^sub>s\<guillemotright> S'"
- by (rule da_weakenE)
- with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
- show ?thesis
- by (rule validE)
+ by (rule da_weakenE)
+ with valid_then_else P' valid_A conf_s1 eval_then_else wt_then_else
+ show ?thesis
+ by (rule validE)
next
- case False
- with valid_then_else P' valid_A conf_s1 eval_then_else
- show ?thesis
- by (cases rule: validE) iprover+
+ case False
+ with valid_then_else P' valid_A conf_s1 eval_then_else
+ show ?thesis
+ by (cases rule: validE) iprover+
qed
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G, L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -1998,7 +1998,7 @@
proof -
--{* From the given hypothesises @{text valid_e} and @{text valid_c}
we can only reach the state after unfolding the loop once, i.e.
- @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
+ @{term "P \<diamondsuit> s2 Z"}, where @{term s2} is the state after executing
@{term c}. To gain validity of the further execution of while, to
finally get @{term "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"} we have to get
a hypothesis about the subsequent unfoldings (the whole loop again),
@@ -2008,202 +2008,202 @@
@{text valid_c} in the goal.
*}
{
- fix t s s' v
- assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
- hence "\<And> Y' T E.
+ fix t s s' v
+ assume "G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (v, s')"
+ hence "\<And> Y' T E.
\<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);
normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>t\<Colon>T;
normal s \<Longrightarrow> \<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>E
\<rbrakk>\<Longrightarrow> (P'\<leftarrow>=False\<down>=\<diamondsuit>) v s' Z"
- (is "PROP ?Hyp n t s v s'")
- proof (induct)
- case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
- note while = `(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
+ (is "PROP ?Hyp n t s v s'")
+ proof (induct)
+ case (Loop s0' e' b n' s1' c' s2' l' s3' Y' T E)
+ note while = `(\<langle>l'\<bullet> While(e') c'\<rangle>\<^sub>s::term) = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
hence eqs: "l'=l" "e'=e" "c'=c" by simp_all
- note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
- note P = `P Y' (Norm s0') Z`
- note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
+ note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
+ note P = `P Y' (Norm s0') Z`
+ note conf_s0' = `Norm s0'\<Colon>\<preceq>(G, L)`
have wt: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<Colon>T"
- using Loop.prems eqs by simp
- have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+ using Loop.prems eqs by simp
+ have da: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
dom (locals (store ((Norm s0')::state)))\<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright>E"
- using Loop.prems eqs by simp
- have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'"
- using Loop.hyps eqs by simp
- show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
- proof -
- from wt obtain
- wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
+ using Loop.prems eqs by simp
+ have evaln_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<midarrow>n'\<rightarrow> s1'"
+ using Loop.hyps eqs by simp
+ show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+ proof -
+ from wt obtain
+ wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-PrimT Boolean" and
wt_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c\<Colon>\<surd>"
- by cases (simp add: eqs)
- from da obtain E S where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ by cases (simp add: eqs)
+ from da obtain E S where
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store ((Norm s0')::state))) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E" and
- da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ da_c: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> (dom (locals (store ((Norm s0')::state)))
\<union> assigns_if True e) \<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S"
- by cases (simp add: eqs)
- from evaln_e
- have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
- by (rule evaln_eval)
- from valid_e P valid_A conf_s0' evaln_e wt_e da_e
- obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
- by (rule validE)
- show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
- proof (cases "normal s1'")
- case True
- note normal_s1'=this
- show ?thesis
- proof (cases "the_Bool b")
- case True
- with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
- by auto
- from True Loop.hyps obtain
- eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and
- eval_while:
- "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
- by (simp add: eqs)
- from True Loop.hyps have
- hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s
+ by cases (simp add: eqs)
+ from evaln_e
+ have eval_e: "G\<turnstile>Norm s0' \<midarrow>e-\<succ>b\<rightarrow> s1'"
+ by (rule evaln_eval)
+ from valid_e P valid_A conf_s0' evaln_e wt_e da_e
+ obtain P': "P' \<lfloor>b\<rfloor>\<^sub>e s1' Z" and conf_s1': "s1'\<Colon>\<preceq>(G,L)"
+ by (rule validE)
+ show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+ proof (cases "normal s1'")
+ case True
+ note normal_s1'=this
+ show ?thesis
+ proof (cases "the_Bool b")
+ case True
+ with P' normal_s1' have P'': "(Normal (P'\<leftarrow>=True)) \<lfloor>b\<rfloor>\<^sub>e s1' Z"
+ by auto
+ from True Loop.hyps obtain
+ eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and
+ eval_while:
+ "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
+ by (simp add: eqs)
+ from True Loop.hyps have
+ hyp: "PROP ?Hyp n' \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s
(abupd (absorb (Cont l')) s2') \<diamondsuit> s3'"
- apply (simp only: True if_True eqs)
- apply (elim conjE)
- apply (tactic "smp_tac 3 1")
- apply fast
- done
- from eval_e
- have s0'_s1': "dom (locals (store ((Norm s0')::state)))
+ apply (simp only: True if_True eqs)
+ apply (elim conjE)
+ apply (tactic "smp_tac 3 1")
+ apply fast
+ done
+ from eval_e
+ have s0'_s1': "dom (locals (store ((Norm s0')::state)))
\<subseteq> dom (locals (store s1'))"
- by (rule dom_locals_eval_mono_elim)
- obtain S' where
- da_c':
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(dom (locals (store s1')))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S'"
- proof -
- note s0'_s1'
- moreover
- from eval_e normal_s1' wt_e
- have "assigns_if True e \<subseteq> dom (locals (store s1'))"
- by (rule assigns_if_good_approx' [elim_format])
+ by (rule dom_locals_eval_mono_elim)
+ obtain S' where
+ da_c':
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>(dom (locals (store s1')))\<guillemotright>\<langle>c\<rangle>\<^sub>s\<guillemotright> S'"
+ proof -
+ note s0'_s1'
+ moreover
+ from eval_e normal_s1' wt_e
+ have "assigns_if True e \<subseteq> dom (locals (store s1'))"
+ by (rule assigns_if_good_approx' [elim_format])
(simp add: True)
- ultimately
- have "dom (locals (store ((Norm s0')::state)))
+ ultimately
+ have "dom (locals (store ((Norm s0')::state)))
\<union> assigns_if True e \<subseteq> dom (locals (store s1'))"
- by (rule Un_least)
- with da_c show thesis
- by (rule da_weakenE) (rule that)
- qed
- with valid_c P'' valid_A conf_s1' eval_c wt_c
- obtain "(abupd (absorb (Cont l)) .; P) \<diamondsuit> s2' Z" and
+ by (rule Un_least)
+ with da_c show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ with valid_c P'' valid_A conf_s1' eval_c wt_c
+ obtain "(abupd (absorb (Cont l)) .; P) \<diamondsuit> s2' Z" and
conf_s2': "s2'\<Colon>\<preceq>(G,L)"
- by (rule validE)
- hence P_s2': "P \<diamondsuit> (abupd (absorb (Cont l)) s2') Z"
- by simp
- from conf_s2'
- have conf_absorb: "abupd (absorb (Cont l)) s2' \<Colon>\<preceq>(G, L)"
- by (cases s2') (auto intro: conforms_absorb)
- moreover
- obtain E' where
- da_while':
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
- dom (locals(store (abupd (absorb (Cont l)) s2')))
+ by (rule validE)
+ hence P_s2': "P \<diamondsuit> (abupd (absorb (Cont l)) s2') Z"
+ by simp
+ from conf_s2'
+ have conf_absorb: "abupd (absorb (Cont l)) s2' \<Colon>\<preceq>(G, L)"
+ by (cases s2') (auto intro: conforms_absorb)
+ moreover
+ obtain E' where
+ da_while':
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>
+ dom (locals(store (abupd (absorb (Cont l)) s2')))
\<guillemotright>\<langle>l\<bullet> While(e) c\<rangle>\<^sub>s\<guillemotright> E'"
- proof -
- note s0'_s1'
- also
- from eval_c
- have "G\<turnstile>s1' \<midarrow>c\<rightarrow> s2'"
- by (rule evaln_eval)
- hence "dom (locals (store s1')) \<subseteq> dom (locals (store s2'))"
- by (rule dom_locals_eval_mono_elim)
- also
- have "\<dots>\<subseteq>dom (locals (store (abupd (absorb (Cont l)) s2')))"
- by simp
- finally
- have "dom (locals (store ((Norm s0')::state))) \<subseteq> \<dots>" .
- with da show thesis
- by (rule da_weakenE) (rule that)
- qed
- from valid_A P_s2' conf_absorb wt da_while'
- show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
- using hyp by (simp add: eqs)
- next
- case False
- with Loop.hyps obtain "s3'=s1'"
- by simp
- with P' False show ?thesis
- by auto
- qed
- next
- case False
- note abnormal_s1'=this
- have "s3'=s1'"
- proof -
- from False obtain abr where abr: "abrupt s1' = Some abr"
- by (cases s1') auto
- from eval_e _ wt_e wf
- have no_jmp: "\<And> j. abrupt s1' \<noteq> Some (Jump j)"
- by (rule eval_expression_no_jump
+ proof -
+ note s0'_s1'
+ also
+ from eval_c
+ have "G\<turnstile>s1' \<midarrow>c\<rightarrow> s2'"
+ by (rule evaln_eval)
+ hence "dom (locals (store s1')) \<subseteq> dom (locals (store s2'))"
+ by (rule dom_locals_eval_mono_elim)
+ also
+ have "\<dots>\<subseteq>dom (locals (store (abupd (absorb (Cont l)) s2')))"
+ by simp
+ finally
+ have "dom (locals (store ((Norm s0')::state))) \<subseteq> \<dots>" .
+ with da show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ from valid_A P_s2' conf_absorb wt da_while'
+ show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3' Z"
+ using hyp by (simp add: eqs)
+ next
+ case False
+ with Loop.hyps obtain "s3'=s1'"
+ by simp
+ with P' False show ?thesis
+ by auto
+ qed
+ next
+ case False
+ note abnormal_s1'=this
+ have "s3'=s1'"
+ proof -
+ from False obtain abr where abr: "abrupt s1' = Some abr"
+ by (cases s1') auto
+ from eval_e _ wt_e wf
+ have no_jmp: "\<And> j. abrupt s1' \<noteq> Some (Jump j)"
+ by (rule eval_expression_no_jump
[where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>",simplified])
- simp
- show ?thesis
- proof (cases "the_Bool b")
- case True
- with Loop.hyps obtain
- eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and
- eval_while:
- "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
- by (simp add: eqs)
- from eval_c abr have "s2'=s1'" by auto
- moreover from calculation no_jmp
- have "abupd (absorb (Cont l)) s2'=s2'"
- by (cases s1') (simp add: absorb_def)
- ultimately show ?thesis
- using eval_while abr
- by auto
- next
- case False
- with Loop.hyps show ?thesis by simp
- qed
- qed
- with P' False show ?thesis
- by auto
- qed
- qed
- next
- case (Abrupt abr s t' n' Y' T E)
- note t' = `t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
- note conf = `(Some abr, s)\<Colon>\<preceq>(G, L)`
- note P = `P Y' (Some abr, s) Z`
- note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
- show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) (undefined3 t') (Some abr, s) Z"
- proof -
- have eval_e:
- "G\<turnstile>(Some abr,s) \<midarrow>\<langle>e\<rangle>\<^sub>e\<succ>\<midarrow>n'\<rightarrow> (undefined3 \<langle>e\<rangle>\<^sub>e,(Some abr,s))"
- by auto
- from valid_e P valid_A conf eval_e
- have "P' (undefined3 \<langle>e\<rangle>\<^sub>e) (Some abr,s) Z"
- by (cases rule: validE [where ?P="P"]) simp+
- with t' show ?thesis
- by auto
- qed
- qed simp_all
+ simp
+ show ?thesis
+ proof (cases "the_Bool b")
+ case True
+ with Loop.hyps obtain
+ eval_c: "G\<turnstile>s1' \<midarrow>c\<midarrow>n'\<rightarrow> s2'" and
+ eval_while:
+ "G\<turnstile>abupd (absorb (Cont l)) s2' \<midarrow>l\<bullet> While(e) c\<midarrow>n'\<rightarrow> s3'"
+ by (simp add: eqs)
+ from eval_c abr have "s2'=s1'" by auto
+ moreover from calculation no_jmp
+ have "abupd (absorb (Cont l)) s2'=s2'"
+ by (cases s1') (simp add: absorb_def)
+ ultimately show ?thesis
+ using eval_while abr
+ by auto
+ next
+ case False
+ with Loop.hyps show ?thesis by simp
+ qed
+ qed
+ with P' False show ?thesis
+ by auto
+ qed
+ qed
+ next
+ case (Abrupt abr s t' n' Y' T E)
+ note t' = `t' = \<langle>l\<bullet> While(e) c\<rangle>\<^sub>s`
+ note conf = `(Some abr, s)\<Colon>\<preceq>(G, L)`
+ note P = `P Y' (Some abr, s) Z`
+ note valid_A = `\<forall>t\<in>A. G\<Turnstile>n'\<Colon>t`
+ show "(P'\<leftarrow>=False\<down>=\<diamondsuit>) (undefined3 t') (Some abr, s) Z"
+ proof -
+ have eval_e:
+ "G\<turnstile>(Some abr,s) \<midarrow>\<langle>e\<rangle>\<^sub>e\<succ>\<midarrow>n'\<rightarrow> (undefined3 \<langle>e\<rangle>\<^sub>e,(Some abr,s))"
+ by auto
+ from valid_e P valid_A conf eval_e
+ have "P' (undefined3 \<langle>e\<rangle>\<^sub>e) (Some abr,s) Z"
+ by (cases rule: validE [where ?P="P"]) simp+
+ with t' show ?thesis
+ by auto
+ qed
+ qed simp_all
} note generalized=this
from eval _ valid_A P conf_s0 wt da
have "(P'\<leftarrow>=False\<down>=\<diamondsuit>) \<diamondsuit> s3 Z"
- by (rule generalized) simp_all
+ by (rule generalized) simp_all
moreover
have "s3\<Colon>\<preceq>(G, L)"
proof (cases "normal s0")
- case True
- from eval wt [OF True] da [OF True] conf_s0 wf
- show ?thesis
- by (rule evaln_type_sound [elim_format]) simp
+ case True
+ from eval wt [OF True] da [OF True] conf_s0 wf
+ show ?thesis
+ by (rule evaln_type_sound [elim_format]) simp
next
- case False
- with eval have "s3=s0"
- by auto
- with conf_s0 show ?thesis
- by simp
+ case False
+ with eval have "s3=s0"
+ by auto
+ with conf_s0 show ?thesis
+ by simp
qed
ultimately show ?thesis ..
qed
@@ -2224,14 +2224,14 @@
show "P \<diamondsuit> s1 Z \<and> s1\<Colon>\<preceq>(G,L)"
proof -
from eval obtain s where
- s: "s0=Norm s" "s1=(Some (Jump j), s)"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ s: "s0=Norm s" "s1=(Some (Jump j), s)"
+ using normal_s0 by (auto elim: evaln_elim_cases)
with P have "P \<diamondsuit> s1 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s1\<Colon>\<preceq>(G,L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -2252,24 +2252,24 @@
show "Q \<diamondsuit> s2 Z \<and> s2\<Colon>\<preceq>(G,L)"
proof -
from eval obtain s1 a where
- eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
- s2: "s2 = abupd (throw a) s1"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ eval_e: "G\<turnstile>s0 \<midarrow>e-\<succ>a\<midarrow>n\<rightarrow> s1" and
+ s2: "s2 = abupd (throw a) s1"
+ using normal_s0 by (auto elim: evaln_elim_cases)
from wt obtain T where
- wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
- by cases simp
+ wt_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>e\<Colon>-T"
+ by cases simp
from da obtain E where
- da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
- by cases simp
+ da_e: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>e\<rangle>\<^sub>e\<guillemotright> E"
+ by cases simp
from valid_e P valid_A conf_s0 eval_e wt_e da_e
obtain "(\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>) \<lfloor>a\<rfloor>\<^sub>e s1 Z"
- by (rule validE)
+ by (rule validE)
with s2 have "Q \<diamondsuit> s2 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s2\<Colon>\<preceq>(G,L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -2294,114 +2294,114 @@
show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
proof -
from eval obtain s1 s2 where
- eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1" and
+ eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> s1" and
sxalloc: "G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2" and
s3: "if G,s2\<turnstile>catch C
then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3
else s3 = s2"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from wt obtain
- wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
- wt_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
- by cases simp
+ wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+ wt_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>\<turnstile>c2\<Colon>\<surd>"
+ by cases simp
from da obtain C1 C2 where
- da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
- da_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
+ da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+ da_c2: "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
\<turnstile> (dom (locals (store s0)) \<union> {VName vn}) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
- by cases simp
+ by cases simp
from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
obtain sxQ: "(SXAlloc G Q) \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
from sxalloc sxQ
have Q: "Q \<diamondsuit> s2 Z"
- by auto
+ by auto
have "R \<diamondsuit> s3 Z"
proof (cases "\<exists> x. abrupt s1 = Some (Xcpt x)")
- case False
- from sxalloc wf
- have "s2=s1"
- by (rule sxalloc_type_sound [elim_format])
- (insert False, auto split: option.splits abrupt.splits )
- with False
- have no_catch: "\<not> G,s2\<turnstile>catch C"
- by (simp add: catch_def)
- moreover
- from no_catch s3
- have "s3=s2"
- by simp
- ultimately show ?thesis
- using Q Q_R by simp
+ case False
+ from sxalloc wf
+ have "s2=s1"
+ by (rule sxalloc_type_sound [elim_format])
+ (insert False, auto split: option.splits abrupt.splits )
+ with False
+ have no_catch: "\<not> G,s2\<turnstile>catch C"
+ by (simp add: catch_def)
+ moreover
+ from no_catch s3
+ have "s3=s2"
+ by simp
+ ultimately show ?thesis
+ using Q Q_R by simp
next
- case True
- note exception_s1 = this
- show ?thesis
- proof (cases "G,s2\<turnstile>catch C")
- case False
- with s3
- have "s3=s2"
- by simp
- with False Q Q_R show ?thesis
- by simp
- next
- case True
- with s3 have eval_c2: "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3"
- by simp
- from conf_s1 sxalloc wf
- have conf_s2: "s2\<Colon>\<preceq>(G, L)"
- by (auto dest: sxalloc_type_sound
+ case True
+ note exception_s1 = this
+ show ?thesis
+ proof (cases "G,s2\<turnstile>catch C")
+ case False
+ with s3
+ have "s3=s2"
+ by simp
+ with False Q Q_R show ?thesis
+ by simp
+ next
+ case True
+ with s3 have eval_c2: "G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<midarrow>n\<rightarrow> s3"
+ by simp
+ from conf_s1 sxalloc wf
+ have conf_s2: "s2\<Colon>\<preceq>(G, L)"
+ by (auto dest: sxalloc_type_sound
split: option.splits abrupt.splits)
- from exception_s1 sxalloc wf
- obtain a
- where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
- by (auto dest!: sxalloc_type_sound
+ from exception_s1 sxalloc wf
+ obtain a
+ where xcpt_s2: "abrupt s2 = Some (Xcpt (Loc a))"
+ by (auto dest!: sxalloc_type_sound
split: option.splits abrupt.splits)
- with True
- have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class C"
- by (cases s2) simp
- with xcpt_s2 conf_s2 wf
- have conf_new_xcpt: "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class C))"
- by (auto dest: Try_lemma)
- obtain C2' where
- da_c2':
- "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
+ with True
+ have "G\<turnstile>obj_ty (the (globs (store s2) (Heap a)))\<preceq>Class C"
+ by (cases s2) simp
+ with xcpt_s2 conf_s2 wf
+ have conf_new_xcpt: "new_xcpt_var vn s2 \<Colon>\<preceq>(G, L(VName vn\<mapsto>Class C))"
+ by (auto dest: Try_lemma)
+ obtain C2' where
+ da_c2':
+ "\<lparr>prg=G,cls=accC,lcl=L(VName vn\<mapsto>Class C)\<rparr>
\<turnstile> (dom (locals (store (new_xcpt_var vn s2)))) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
- proof -
- have "(dom (locals (store s0)) \<union> {VName vn})
+ proof -
+ have "(dom (locals (store s0)) \<union> {VName vn})
\<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
proof -
- from eval_c1
+ from eval_c1
have "dom (locals (store s0))
\<subseteq> dom (locals (store s1))"
- by (rule dom_locals_evaln_mono_elim)
+ by (rule dom_locals_evaln_mono_elim)
also
from sxalloc
have "\<dots> \<subseteq> dom (locals (store s2))"
- by (rule dom_locals_sxalloc_mono)
+ by (rule dom_locals_sxalloc_mono)
also
have "\<dots> \<subseteq> dom (locals (store (new_xcpt_var vn s2)))"
- by (cases s2) (simp add: new_xcpt_var_def, blast)
+ by (cases s2) (simp add: new_xcpt_var_def, blast)
also
have "{VName vn} \<subseteq> \<dots>"
- by (cases s2) simp
+ by (cases s2) simp
ultimately show ?thesis
- by (rule Un_least)
+ by (rule Un_least)
qed
- with da_c2 show thesis
- by (rule da_weakenE) (rule that)
- qed
- from Q eval_c2 True
- have "(Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn)
+ with da_c2 show thesis
+ by (rule da_weakenE) (rule that)
+ qed
+ from Q eval_c2 True
+ have "(Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn)
\<diamondsuit> (new_xcpt_var vn s2) Z"
- by auto
- from valid_c2 this valid_A conf_new_xcpt eval_c2 wt_c2 da_c2'
- show "R \<diamondsuit> s3 Z"
- by (rule validE)
- qed
+ by auto
+ from valid_c2 this valid_A conf_new_xcpt eval_c2 wt_c2 da_c2'
+ show "R \<diamondsuit> s3 Z"
+ by (rule validE)
+ qed
qed
moreover
from eval wt da conf_s0 wf
have "s3\<Colon>\<preceq>(G,L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -2426,56 +2426,56 @@
show "R \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
proof -
from eval obtain s1 abr1 s2 where
- eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> (abr1, s1)" and
+ eval_c1: "G\<turnstile>s0 \<midarrow>c1\<midarrow>n\<rightarrow> (abr1, s1)" and
eval_c2: "G\<turnstile>Norm s1 \<midarrow>c2\<midarrow>n\<rightarrow> s2" and
s3: "s3 = (if \<exists>err. abr1 = Some (Error err)
then (abr1, s1)
else abupd (abrupt_if (abr1 \<noteq> None) abr1) s2)"
- using normal_s0 by (fastsimp elim: evaln_elim_cases)
+ using normal_s0 by (fastsimp elim: evaln_elim_cases)
from wt obtain
- wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
- wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c2\<Colon>\<surd>"
- by cases simp
+ wt_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c1\<Colon>\<surd>" and
+ wt_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>\<turnstile>c2\<Colon>\<surd>"
+ by cases simp
from da obtain C1 C2 where
- da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
- da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
- by cases simp
+ da_c1: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c1\<rangle>\<^sub>s\<guillemotright> C1" and
+ da_c2: "\<lparr>prg=G,cls=accC,lcl=L\<rparr> \<turnstile> dom (locals (store s0)) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2"
+ by cases simp
from valid_c1 P valid_A conf_s0 eval_c1 wt_c1 da_c1
obtain Q: "Q \<diamondsuit> (abr1,s1) Z" and conf_s1: "(abr1,s1)\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
from Q
have Q': "(Q \<and>. (\<lambda>s. abr1 = fst s) ;. abupd (\<lambda>x. None)) \<diamondsuit> (Norm s1) Z"
- by auto
+ by auto
from eval_c1 wt_c1 da_c1 conf_s0 wf
have "error_free (abr1,s1)"
- by (rule evaln_type_sound [elim_format]) (insert normal_s0,simp)
+ by (rule evaln_type_sound [elim_format]) (insert normal_s0,simp)
with s3 have s3': "s3 = abupd (abrupt_if (abr1 \<noteq> None) abr1) s2"
- by (simp add: error_free_def)
+ by (simp add: error_free_def)
from conf_s1
have conf_Norm_s1: "Norm s1\<Colon>\<preceq>(G,L)"
- by (rule conforms_NormI)
+ by (rule conforms_NormI)
obtain C2' where
- da_c2': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ da_c2': "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store ((Norm s1)::state))) \<guillemotright>\<langle>c2\<rangle>\<^sub>s\<guillemotright> C2'"
proof -
- from eval_c1
- have "dom (locals (store s0)) \<subseteq> dom (locals (store (abr1,s1)))"
+ from eval_c1
+ have "dom (locals (store s0)) \<subseteq> dom (locals (store (abr1,s1)))"
by (rule dom_locals_evaln_mono_elim)
- hence "dom (locals (store s0))
+ hence "dom (locals (store s0))
\<subseteq> dom (locals (store ((Norm s1)::state)))"
- by simp
- with da_c2 show thesis
- by (rule da_weakenE) (rule that)
+ by simp
+ with da_c2 show thesis
+ by (rule da_weakenE) (rule that)
qed
from valid_c2 Q' valid_A conf_Norm_s1 eval_c2 wt_c2 da_c2'
have "(abupd (abrupt_if (abr1 \<noteq> None) abr1) .; R) \<diamondsuit> s2 Z"
- by (rule validE)
+ by (rule validE)
with s3' have "R \<diamondsuit> s3 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s3\<Colon>\<preceq>(G,L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
@@ -2495,11 +2495,11 @@
show "P \<diamondsuit> s3 Z \<and> s3\<Colon>\<preceq>(G,L)"
proof -
from P have inited: "inited C (globs (store s0))"
- by simp
+ by simp
with eval have "s3=s0"
- using normal_s0 by (auto elim: evaln_elim_cases)
+ using normal_s0 by (auto elim: evaln_elim_cases)
with P conf_s0 show ?thesis
- by simp
+ by simp
qed
qed
next
@@ -2529,88 +2529,88 @@
proof -
from P have not_inited: "\<not> inited C (globs (store s0))" by simp
with eval c obtain s1 s2 where
- eval_super:
- "G\<turnstile>Norm ((init_class_obj G C) (store s0))
+ eval_super:
+ "G\<turnstile>Norm ((init_class_obj G C) (store s0))
\<midarrow>(if C = Object then Skip else Init (super c))\<midarrow>n\<rightarrow> s1" and
- eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2" and
+ eval_init: "G\<turnstile>(set_lvars empty) s1 \<midarrow>init c\<midarrow>n\<rightarrow> s2" and
s3: "s3 = (set_lvars (locals (store s1))) s2"
- using normal_s0 by (auto elim!: evaln_elim_cases)
+ using normal_s0 by (auto elim!: evaln_elim_cases)
from wt c have
- cls_C: "class G C = Some c"
- by cases auto
+ cls_C: "class G C = Some c"
+ by cases auto
from wf cls_C have
- wt_super: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ wt_super: "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile>(if C = Object then Skip else Init (super c))\<Colon>\<surd>"
- by (cases "C=Object")
+ by (cases "C=Object")
(auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD)
obtain S where
- da_super:
- "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
+ da_super:
+ "\<lparr>prg=G,cls=accC,lcl=L\<rparr>
\<turnstile> dom (locals (store ((Norm
((init_class_obj G C) (store s0)))::state)))
\<guillemotright>\<langle>if C = Object then Skip else Init (super c)\<rangle>\<^sub>s\<guillemotright> S"
proof (cases "C=Object")
- case True
- with da_Skip show ?thesis
- using that by (auto intro: assigned.select_convs)
+ case True
+ with da_Skip show ?thesis
+ using that by (auto intro: assigned.select_convs)
next
- case False
- with da_Init show ?thesis
- by - (rule that, auto intro: assigned.select_convs)
+ case False
+ with da_Init show ?thesis
+ by - (rule that, auto intro: assigned.select_convs)
qed
from normal_s0 conf_s0 wf cls_C not_inited
have conf_init_cls: "(Norm ((init_class_obj G C) (store s0)))\<Colon>\<preceq>(G, L)"
- by (auto intro: conforms_init_class_obj)
+ by (auto intro: conforms_init_class_obj)
from P
have P': "(Normal (P \<and>. Not \<circ> initd C ;. supd (init_class_obj G C)))
Y (Norm ((init_class_obj G C) (store s0))) Z"
- by auto
+ by auto
from valid_super P' valid_A conf_init_cls eval_super wt_super da_super
obtain Q: "Q \<diamondsuit> s1 Z" and conf_s1: "s1\<Colon>\<preceq>(G,L)"
- by (rule validE)
+ by (rule validE)
from cls_C wf have wt_init: "\<lparr>prg=G, cls=C,lcl=empty\<rparr>\<turnstile>(init c)\<Colon>\<surd>"
- by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
+ by (rule wf_prog_cdecl [THEN wf_cdecl_wt_init])
from cls_C wf obtain I where
- "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I"
- by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
+ "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile> {} \<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I"
+ by (rule wf_prog_cdecl [THEN wf_cdeclE,simplified]) blast
(* simplified: to rewrite \<langle>init c\<rangle> to In1r (init c) *)
then obtain I' where
- da_init:
- "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1)))
+ da_init:
+ "\<lparr>prg=G,cls=C,lcl=empty\<rparr>\<turnstile>dom (locals (store ((set_lvars empty) s1)))
\<guillemotright>\<langle>init c\<rangle>\<^sub>s\<guillemotright> I'"
- by (rule da_weakenE) simp
+ by (rule da_weakenE) simp
have conf_s1_empty: "(set_lvars empty) s1\<Colon>\<preceq>(G, empty)"
proof -
- from eval_super have
- "G\<turnstile>Norm ((init_class_obj G C) (store s0))
+ from eval_super have
+ "G\<turnstile>Norm ((init_class_obj G C) (store s0))
\<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1"
- by (rule evaln_eval)
- from this wt_super wf
- have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
- by - (rule eval_statement_no_jump
+ by (rule evaln_eval)
+ from this wt_super wf
+ have s1_no_ret: "\<And> j. abrupt s1 \<noteq> Some (Jump j)"
+ by - (rule eval_statement_no_jump
[where ?Env="\<lparr>prg=G,cls=accC,lcl=L\<rparr>"], auto split: split_if)
with conf_s1
- show ?thesis
- by (cases s1) (auto intro: conforms_set_locals)
+ show ?thesis
+ by (cases s1) (auto intro: conforms_set_locals)
qed
obtain l where l: "l = locals (store s1)"
- by simp
+ by simp
with Q
have Q': "(Q \<and>. (\<lambda>s. l = locals (snd s)) ;. set_lvars empty)
\<diamondsuit> ((set_lvars empty) s1) Z"
- by auto
+ by auto
from valid_init Q' valid_A conf_s1_empty eval_init wt_init da_init
have "(set_lvars l .; R) \<diamondsuit> s2 Z"
- by (rule validE)
+ by (rule validE)
with s3 l have "R \<diamondsuit> s3 Z"
- by simp
+ by simp
moreover
from eval wt da conf_s0 wf
have "s3\<Colon>\<preceq>(G,L)"
- by (rule evaln_type_sound [elim_format]) simp
+ by (rule evaln_type_sound [elim_format]) simp
ultimately show ?thesis ..
qed
qed
--- a/src/HOL/Bali/Basis.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/Basis.thy Sat Oct 17 14:43:18 2009 +0200
@@ -86,19 +86,19 @@
proof (cases rule: converse_rtranclE)
assume "a=y"
with a_v_r v_x_rt have "(y,x) \<in> r\<^sup>*"
- by (auto intro: r_into_rtrancl rtrancl_trans)
+ by (auto intro: r_into_rtrancl rtrancl_trans)
then show ?thesis
- by blast
+ by blast
next
fix w
assume a_w_r: "(a, w) \<in> r" and
w_y_rt: "(w, y) \<in> r\<^sup>*"
from a_v_r a_w_r unique
have "v=w"
- by auto
+ by auto
with w_y_rt hyp
show ?thesis
- by blast
+ by blast
qed
qed
qed
@@ -213,11 +213,11 @@
primrec "the_In3 (In3 c) = c"
syntax
- In1l :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
- In1r :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
+ In1l :: "'al \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
+ In1r :: "'ar \<Rightarrow> ('al + 'ar, 'b, 'c) sum3"
translations
- "In1l e" == "In1 (Inl e)"
- "In1r c" == "In1 (Inr c)"
+ "In1l e" == "In1 (Inl e)"
+ "In1r c" == "In1 (Inr c)"
syntax the_In1l :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'al"
the_In1r :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> 'ar"
--- a/src/HOL/Bali/Conform.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/Conform.thy Sat Oct 17 14:43:18 2009 +0200
@@ -17,7 +17,7 @@
\end{itemize}
*}
-types env' = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
+types env' = "prog \<times> (lname, ty) table" (* same as env of WellType.thy *)
section "extension of global store"
@@ -102,7 +102,7 @@
constdefs
conf :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>_" [71,71,71,71] 70)
- "G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. Option.map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
+ "G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. Option.map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
by (auto simp: conf_def)
@@ -250,7 +250,7 @@
done
lemma lconf_init_vals [intro!]:
- " \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<Colon>\<preceq>]fs"
+ " \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<Colon>\<preceq>]fs"
apply (unfold lconf_def)
apply force
done
@@ -338,7 +338,7 @@
by (simp add: wlconf_def)
lemma wlconf_init_vals [intro!]:
- " \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<sim>\<Colon>\<preceq>]fs"
+ " \<forall>n. \<forall>T\<in>fs n:is_type G T \<Longrightarrow> G,s\<turnstile>init_vals fs[\<sim>\<Colon>\<preceq>]fs"
apply (unfold wlconf_def)
apply force
done
@@ -352,9 +352,9 @@
constdefs
oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_" [71,71,71,71] 70)
- "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and>
+ "G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and>
(case r of
- Heap a \<Rightarrow> is_type G (obj_ty obj)
+ Heap a \<Rightarrow> is_type G (obj_ty obj)
| Stat C \<Rightarrow> True)"
@@ -405,11 +405,11 @@
by (auto simp: conforms_def Let_def)
lemma conforms_XcptLocD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Xcpt (Loc a))\<rbrakk> \<Longrightarrow>
- G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
+ G,s\<turnstile>Addr a\<Colon>\<preceq> Class (SXcpt Throwable)"
by (auto simp: conforms_def Let_def)
lemma conforms_RetD: "\<lbrakk>(x, s)\<Colon>\<preceq>(G, L); x = Some (Jump Ret)\<rbrakk> \<Longrightarrow>
- (locals s) Result \<noteq> None"
+ (locals s) Result \<noteq> None"
by (auto simp: conforms_def Let_def)
lemma conforms_RefTD:
--- a/src/HOL/Bali/Decl.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/Decl.thy Sat Oct 17 14:43:18 2009 +0200
@@ -171,7 +171,7 @@
types
fdecl (* field declaration, cf. 8.3 *)
- = "vname \<times> field"
+ = "vname \<times> field"
translations
@@ -294,9 +294,9 @@
record iface = ibody + --{* interface *}
isuperIfs:: "qtname list" --{* superinterface list *}
-types
- idecl --{* interface declaration, cf. 9.1 *}
- = "qtname \<times> iface"
+types
+ idecl --{* interface declaration, cf. 9.1 *}
+ = "qtname \<times> iface"
translations
"ibody" <= (type) "\<lparr>access::acc_modi,imethods::(sig \<times> mhead) list\<rparr>"
@@ -326,9 +326,9 @@
record "class" = cbody + --{* class *}
super :: "qtname" --{* superclass *}
superIfs:: "qtname list" --{* implemented interfaces *}
-types
- cdecl --{* class declaration, cf. 8.1 *}
- = "qtname \<times> class"
+types
+ cdecl --{* class declaration, cf. 8.1 *}
+ = "qtname \<times> class"
translations
"cbody" <= (type) "\<lparr>access::acc_modi,cfields::fdecl list,
@@ -388,8 +388,8 @@
constdefs standard_classes :: "cdecl list"
"standard_classes \<equiv> [ObjectC, SXcptC Throwable,
- SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
- SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
+ SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast,
+ SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]"
section "programs"
@@ -409,10 +409,10 @@
is_class :: "prog \<Rightarrow> qtname \<Rightarrow> bool"
translations
- "iface G I" == "table_of (ifaces G) I"
- "class G C" == "table_of (classes G) C"
- "is_iface G I" == "iface G I \<noteq> None"
- "is_class G C" == "class G C \<noteq> None"
+ "iface G I" == "table_of (ifaces G) I"
+ "class G C" == "table_of (classes G) C"
+ "is_iface G I" == "iface G I \<noteq> None"
+ "is_class G C" == "class G C \<noteq> None"
section "is type"
@@ -421,12 +421,12 @@
is_type :: "prog \<Rightarrow> ty \<Rightarrow> bool"
isrtype :: "prog \<Rightarrow> ref_ty \<Rightarrow> bool"
-primrec "is_type G (PrimT pt) = True"
- "is_type G (RefT rt) = isrtype G rt"
- "isrtype G (NullT ) = True"
- "isrtype G (IfaceT tn) = is_iface G tn"
- "isrtype G (ClassT tn) = is_class G tn"
- "isrtype G (ArrayT T ) = is_type G T"
+primrec "is_type G (PrimT pt) = True"
+ "is_type G (RefT rt) = isrtype G rt"
+ "isrtype G (NullT ) = True"
+ "isrtype G (IfaceT tn) = is_iface G tn"
+ "isrtype G (ClassT tn) = is_class G tn"
+ "isrtype G (ArrayT T ) = is_type G T"
lemma type_is_iface: "is_type G (Iface I) \<Longrightarrow> is_iface G I"
by auto
@@ -457,7 +457,7 @@
translations
"G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" == "(C,D) \<in> subcls1 G"
- "G\<turnstile>C \<preceq>\<^sub>C D" == "(C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *)
+ "G\<turnstile>C \<preceq>\<^sub>C D" == "(C,D) \<in>(subcls1 G)^*" (* cf. 8.1.3 *)
"G\<turnstile>C \<prec>\<^sub>C D" == "(C,D) \<in>(subcls1 G)^+"
@@ -516,7 +516,7 @@
ws_prog :: "prog \<Rightarrow> bool"
"ws_prog G \<equiv> (\<forall>(I,i)\<in>set (ifaces G). ws_idecl G I (isuperIfs i)) \<and>
- (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
+ (\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
lemma ws_progI:
@@ -710,12 +710,12 @@
next
case False
with ws iscls obtain sc where
- sc: "class G (super c) = Some sc"
- by (auto dest: ws_prog_cdeclD)
+ sc: "class G (super c) = Some sc"
+ by (auto dest: ws_prog_cdeclD)
from iscls False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 (super c)" by (rule subcls1I)
with False ws step hyp iscls sc
show "P C c"
- by (auto)
+ by (auto)
qed
qed
with clsC show "P C c" by auto
@@ -742,12 +742,12 @@
case Nil
with if_I hyp_sub
show "P I"
- by auto
+ by auto
next
case (Cons hd tl)
with hyp if_I hyp_sub
show "P I"
- by auto
+ by auto
qed
qed
qed
@@ -802,7 +802,7 @@
\<equiv> iface_rec (G,I)
(\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus>
(Option.set \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
-
+
end
--- a/src/HOL/Bali/DeclConcepts.thy Sat Oct 17 01:05:59 2009 +0200
+++ b/src/HOL/Bali/DeclConcepts.thy Sat Oct 17 14:43:18 2009 +0200
@@ -922,20 +922,20 @@
"memberid n = memberid m"
"G\<turnstile> mbr m declared_in C"
"declclass m = C"
- by auto
+ by auto
with member_n
show ?thesis
- by (cases n, cases m)
+ by (cases n, cases m)
(auto simp add: declared_in_def
cdeclaredmethd_def cdeclaredfield_def
split: memberdecl.splits)
next
case (Inherited m' _ _)
then have "G\<turnstile> memberid m undeclared_in C"
- by simp
+ by simp
with eqid member_n
show ?thesis
- by (cases n) (auto dest: declared_not_undeclared)
+ by (cases n) (auto dest: declared_not_undeclared)
qed
next
case (Inherited n C S)
@@ -950,14 +950,14 @@
then have "G\<turnstile> mbr m declared_in C" by simp
with eqid undecl
show ?thesis
- by (cases m) (auto dest: declared_not_undeclared)
+ by (cases m) (auto dest: declared_not_undeclared)
next
case Inherited
with super have "G \<turnstile> m member_of S"
- by (auto dest!: subcls1D)
+ by (auto dest!: subcls1D)
with eqid hyp
show ?thesis
- by blast
+ by blast
qed
qed
qed
@@ -1227,15 +1227,15 @@
assume "D=C"
with super member_of_D_props
show ?thesis
- by (auto intro: members.Inherited)
+ by (auto intro: members.Inherited)
next
case Subcls
assume "G\<turnstile>D\<prec>\<^sub>C C"
with super
have "G\<turnstile>S\<preceq>\<^sub>C C"
- by (auto dest: subcls1D subcls_superD)
+ by (auto dest: subcls1D subcls_superD)
with subclseq_C_m hyp show ?thesis
- by blast
+ by blast
qed
qed
qed
@@ -1660,34 +1660,34 @@
case (Immediate membr Ca)
then have "Ca=C" "membr = method sig m" and
"G\<turnstile>Methd sig m declared_in C" "declclass m = C"
- by (cases m,auto)
+ by (cases m,auto)
with clsC
have "table_of (map (\<lambda>(s, m). (s, C, m)) (methods c)) sig = Some m"
- by (cases m)
- (auto simp add: declared_in_def cdeclaredmethd_def
- intro: table_of_mapconst_SomeI)
+ by (cases m)
+ (auto simp add: declared_in_def cdeclaredmethd_def
+ intro: table_of_mapconst_SomeI)