diff -r f2094906e491 -r e44d86131648 src/ZF/UNITY/SubstAx.thy --- a/src/ZF/UNITY/SubstAx.thy Tue Sep 27 13:34:54 2022 +0200 +++ b/src/ZF/UNITY/SubstAx.thy Tue Sep 27 16:51:35 2022 +0100 @@ -14,18 +14,18 @@ definition (* The definitions below are not `conventional', but yield simpler rules *) Ensures :: "[i,i] => i" (infixl \Ensures\ 60) where - "A Ensures B == {F \ program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" + "A Ensures B \ {F \ program. F \ (reachable(F) \ A) ensures (reachable(F) \ B) }" definition LeadsTo :: "[i, i] => i" (infixl \\w\ 60) where - "A \w B == {F \ program. F:(reachable(F) \ A) \ (reachable(F) \ B)}" + "A \w B \ {F \ program. F:(reachable(F) \ A) \ (reachable(F) \ B)}" (*Resembles the previous definition of LeadsTo*) (* Equivalence with the HOL-like definition *) lemma LeadsTo_eq: -"st_set(B)==> A \w B = {F \ program. F:(reachable(F) \ A) \ B}" +"st_set(B)\ A \w B = {F \ program. F:(reachable(F) \ A) \ B}" apply (unfold LeadsTo_def) apply (blast dest: psp_stable2 leadsToD2 constrainsD2 intro: leadsTo_weaken) done @@ -36,35 +36,35 @@ (*** Specialized laws for handling invariants ***) (** Conjoining an Always property **) -lemma Always_LeadsTo_pre: "F \ Always(I) ==> (F:(I \ A) \w A') \ (F \ A \w A')" +lemma Always_LeadsTo_pre: "F \ Always(I) \ (F:(I \ A) \w A') \ (F \ A \w A')" by (simp add: LeadsTo_def Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2) -lemma Always_LeadsTo_post: "F \ Always(I) ==> (F \ A \w (I \ A')) \ (F \ A \w A')" +lemma Always_LeadsTo_post: "F \ Always(I) \ (F \ A \w (I \ A')) \ (F \ A \w A')" apply (unfold LeadsTo_def) apply (simp add: Always_eq_includes_reachable Int_absorb2 Int_assoc [symmetric] leadsToD2) done (* Like 'Always_LeadsTo_pre RS iffD1', but with premises in the good order *) -lemma Always_LeadsToI: "[| F \ Always(C); F \ (C \ A) \w A' |] ==> F \ A \w A'" +lemma Always_LeadsToI: "\F \ Always(C); F \ (C \ A) \w A'\ \ F \ A \w A'" by (blast intro: Always_LeadsTo_pre [THEN iffD1]) (* Like 'Always_LeadsTo_post RS iffD2', but with premises in the good order *) -lemma Always_LeadsToD: "[| F \ Always(C); F \ A \w A' |] ==> F \ A \w (C \ A')" +lemma Always_LeadsToD: "\F \ Always(C); F \ A \w A'\ \ F \ A \w (C \ A')" by (blast intro: Always_LeadsTo_post [THEN iffD2]) (*** Introduction rules \ Basis, Trans, Union ***) -lemma LeadsTo_Basis: "F \ A Ensures B ==> F \ A \w B" +lemma LeadsTo_Basis: "F \ A Ensures B \ F \ A \w B" by (auto simp add: Ensures_def LeadsTo_def) lemma LeadsTo_Trans: - "[| F \ A \w B; F \ B \w C |] ==> F \ A \w C" + "\F \ A \w B; F \ B \w C\ \ F \ A \w C" apply (simp (no_asm_use) add: LeadsTo_def) apply (blast intro: leadsTo_Trans) done lemma LeadsTo_Union: -"[|(!!A. A \ S ==> F \ A \w B); F \ program|]==>F \ \(S) \w B" +"\(\A. A \ S \ F \ A \w B); F \ program\\F \ \(S) \w B" apply (simp add: LeadsTo_def) apply (subst Int_Union_Union2) apply (rule leadsTo_UN, auto) @@ -72,23 +72,23 @@ (*** Derived rules ***) -lemma leadsTo_imp_LeadsTo: "F \ A \ B ==> F \ A \w B" +lemma leadsTo_imp_LeadsTo: "F \ A \ B \ F \ A \w B" apply (frule leadsToD2, clarify) apply (simp (no_asm_simp) add: LeadsTo_eq) apply (blast intro: leadsTo_weaken_L) done (*Useful with cancellation, disjunction*) -lemma LeadsTo_Un_duplicate: "F \ A \w (A' \ A') ==> F \ A \w A'" +lemma LeadsTo_Un_duplicate: "F \ A \w (A' \ A') \ F \ A \w A'" by (simp add: Un_ac) lemma LeadsTo_Un_duplicate2: - "F \ A \w (A' \ C \ C) ==> F \ A \w (A' \ C)" + "F \ A \w (A' \ C \ C) \ F \ A \w (A' \ C)" by (simp add: Un_ac) lemma LeadsTo_UN: - "[|(!!i. i \ I ==> F \ A(i) \w B); F \ program|] - ==>F:(\i \ I. A(i)) \w B" + "\(\i. i \ I \ F \ A(i) \w B); F \ program\ + \F:(\i \ I. A(i)) \w B" apply (simp add: LeadsTo_def) apply (simp (no_asm_simp) del: UN_simps add: Int_UN_distrib) apply (rule leadsTo_UN, auto) @@ -96,7 +96,7 @@ (*Binary union introduction rule*) lemma LeadsTo_Un: - "[| F \ A \w C; F \ B \w C |] ==> F \ (A \ B) \w C" + "\F \ A \w C; F \ B \w C\ \ F \ (A \ B) \w C" apply (subst Un_eq_Union) apply (rule LeadsTo_Union) apply (auto dest: LeadsTo_type [THEN subsetD]) @@ -104,11 +104,11 @@ (*Lets us look at the starting state*) lemma single_LeadsTo_I: - "[|(!!s. s \ A ==> F:{s} \w B); F \ program|]==>F \ A \w B" + "\(\s. s \ A \ F:{s} \w B); F \ program\\F \ A \w B" apply (subst UN_singleton [symmetric], rule LeadsTo_UN, auto) done -lemma subset_imp_LeadsTo: "[| A \ B; F \ program |] ==> F \ A \w B" +lemma subset_imp_LeadsTo: "\A \ B; F \ program\ \ F \ A \w B" apply (simp (no_asm_simp) add: LeadsTo_def) apply (blast intro: subset_imp_leadsTo) done @@ -122,33 +122,33 @@ by (auto dest: LeadsTo_type [THEN subsetD] simp add: LeadsTo_eq) declare LeadsTo_state [iff] -lemma LeadsTo_weaken_R: "[| F \ A \w A'; A'<=B'|] ==> F \ A \w B'" +lemma LeadsTo_weaken_R: "\F \ A \w A'; A'<=B'\ \ F \ A \w B'" apply (unfold LeadsTo_def) apply (auto intro: leadsTo_weaken_R) done -lemma LeadsTo_weaken_L: "[| F \ A \w A'; B \ A |] ==> F \ B \w A'" +lemma LeadsTo_weaken_L: "\F \ A \w A'; B \ A\ \ F \ B \w A'" apply (unfold LeadsTo_def) apply (auto intro: leadsTo_weaken_L) done -lemma LeadsTo_weaken: "[| F \ A \w A'; B<=A; A'<=B' |] ==> F \ B \w B'" +lemma LeadsTo_weaken: "\F \ A \w A'; B<=A; A'<=B'\ \ F \ B \w B'" by (blast intro: LeadsTo_weaken_R LeadsTo_weaken_L LeadsTo_Trans) lemma Always_LeadsTo_weaken: -"[| F \ Always(C); F \ A \w A'; C \ B \ A; C \ A' \ B' |] - ==> F \ B \w B'" +"\F \ Always(C); F \ A \w A'; C \ B \ A; C \ A' \ B'\ + \ F \ B \w B'" apply (blast dest: Always_LeadsToI intro: LeadsTo_weaken Always_LeadsToD) done (** Two theorems for "proof lattices" **) -lemma LeadsTo_Un_post: "F \ A \w B ==> F:(A \ B) \w B" +lemma LeadsTo_Un_post: "F \ A \w B \ F:(A \ B) \w B" by (blast dest: LeadsTo_type [THEN subsetD] intro: LeadsTo_Un subset_imp_LeadsTo) -lemma LeadsTo_Trans_Un: "[| F \ A \w B; F \ B \w C |] - ==> F \ (A \ B) \w C" +lemma LeadsTo_Trans_Un: "\F \ A \w B; F \ B \w C\ + \ F \ (A \ B) \w C" apply (blast intro: LeadsTo_Un subset_imp_LeadsTo LeadsTo_weaken_L LeadsTo_Trans dest: LeadsTo_type [THEN subsetD]) done @@ -166,14 +166,14 @@ (** More rules using the premise "Always(I)" **) -lemma EnsuresI: "[| F:(A-B) Co (A \ B); F \ transient (A-B) |] ==> F \ A Ensures B" +lemma EnsuresI: "\F:(A-B) Co (A \ B); F \ transient (A-B)\ \ F \ A Ensures B" apply (simp add: Ensures_def Constrains_eq_constrains) apply (blast intro: ensuresI constrains_weaken transient_strengthen dest: constrainsD2) done -lemma Always_LeadsTo_Basis: "[| F \ Always(I); F \ (I \ (A-A')) Co (A \ A'); - F \ transient (I \ (A-A')) |] - ==> F \ A \w A'" +lemma Always_LeadsTo_Basis: "\F \ Always(I); F \ (I \ (A-A')) Co (A \ A'); + F \ transient (I \ (A-A'))\ + \ F \ A \w A'" apply (rule Always_LeadsToI, assumption) apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen) done @@ -181,36 +181,36 @@ (*Set difference: maybe combine with leadsTo_weaken_L?? This is the most useful form of the "disjunction" rule*) lemma LeadsTo_Diff: - "[| F \ (A-B) \w C; F \ (A \ B) \w C |] ==> F \ A \w C" + "\F \ (A-B) \w C; F \ (A \ B) \w C\ \ F \ A \w C" by (blast intro: LeadsTo_Un LeadsTo_weaken) lemma LeadsTo_UN_UN: - "[|(!!i. i \ I ==> F \ A(i) \w A'(i)); F \ program |] - ==> F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" + "\(\i. i \ I \ F \ A(i) \w A'(i)); F \ program\ + \ F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" apply (rule LeadsTo_Union, auto) apply (blast intro: LeadsTo_weaken_R) done (*Binary union version*) lemma LeadsTo_Un_Un: - "[| F \ A \w A'; F \ B \w B' |] ==> F:(A \ B) \w (A' \ B')" + "\F \ A \w A'; F \ B \w B'\ \ F:(A \ B) \w (A' \ B')" by (blast intro: LeadsTo_Un LeadsTo_weaken_R) (** The cancellation law **) -lemma LeadsTo_cancel2: "[| F \ A \w(A' \ B); F \ B \w B' |] ==> F \ A \w (A' \ B')" +lemma LeadsTo_cancel2: "\F \ A \w(A' \ B); F \ B \w B'\ \ F \ A \w (A' \ B')" by (blast intro: LeadsTo_Un_Un subset_imp_LeadsTo LeadsTo_Trans dest: LeadsTo_type [THEN subsetD]) lemma Un_Diff: "A \ (B - A) = A \ B" by auto -lemma LeadsTo_cancel_Diff2: "[| F \ A \w (A' \ B); F \ (B-A') \w B' |] ==> F \ A \w (A' \ B')" +lemma LeadsTo_cancel_Diff2: "\F \ A \w (A' \ B); F \ (B-A') \w B'\ \ F \ A \w (A' \ B')" apply (rule LeadsTo_cancel2) prefer 2 apply assumption apply (simp (no_asm_simp) add: Un_Diff) done -lemma LeadsTo_cancel1: "[| F \ A \w (B \ A'); F \ B \w B' |] ==> F \ A \w (B' \ A')" +lemma LeadsTo_cancel1: "\F \ A \w (B \ A'); F \ B \w B'\ \ F \ A \w (B' \ A')" apply (simp add: Un_commute) apply (blast intro!: LeadsTo_cancel2) done @@ -218,7 +218,7 @@ lemma Diff_Un2: "(B - A) \ A = B \ A" by auto -lemma LeadsTo_cancel_Diff1: "[| F \ A \w (B \ A'); F \ (B-A') \w B' |] ==> F \ A \w (B' \ A')" +lemma LeadsTo_cancel_Diff1: "\F \ A \w (B \ A'); F \ (B-A') \w B'\ \ F \ A \w (B' \ A')" apply (rule LeadsTo_cancel1) prefer 2 apply assumption apply (simp (no_asm_simp) add: Diff_Un2) @@ -227,7 +227,7 @@ (** The impossibility law **) (*The set "A" may be non-empty, but it contains no reachable states*) -lemma LeadsTo_empty: "F \ A \w 0 ==> F \ Always (state -A)" +lemma LeadsTo_empty: "F \ A \w 0 \ F \ Always (state -A)" apply (simp (no_asm_use) add: LeadsTo_def Always_eq_includes_reachable) apply (cut_tac reachable_type) apply (auto dest!: leadsTo_empty) @@ -236,26 +236,26 @@ (** PSP \ Progress-Safety-Progress **) (*Special case of PSP \ Misra's "stable conjunction"*) -lemma PSP_Stable: "[| F \ A \w A'; F \ Stable(B) |]==> F:(A \ B) \w (A' \ B)" +lemma PSP_Stable: "\F \ A \w A'; F \ Stable(B)\\ F:(A \ B) \w (A' \ B)" apply (simp add: LeadsTo_def Stable_eq_stable, clarify) apply (drule psp_stable, assumption) apply (simp add: Int_ac) done -lemma PSP_Stable2: "[| F \ A \w A'; F \ Stable(B) |] ==> F \ (B \ A) \w (B \ A')" +lemma PSP_Stable2: "\F \ A \w A'; F \ Stable(B)\ \ F \ (B \ A) \w (B \ A')" apply (simp (no_asm_simp) add: PSP_Stable Int_ac) done -lemma PSP: "[| F \ A \w A'; F \ B Co B'|]==> F \ (A \ B') \w ((A' \ B) \ (B' - B))" +lemma PSP: "\F \ A \w A'; F \ B Co B'\\ F \ (A \ B') \w ((A' \ B) \ (B' - B))" apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains) apply (blast dest: psp intro: leadsTo_weaken) done -lemma PSP2: "[| F \ A \w A'; F \ B Co B' |]==> F:(B' \ A) \w ((B \ A') \ (B' - B))" +lemma PSP2: "\F \ A \w A'; F \ B Co B'\\ F:(B' \ A) \w ((B \ A') \ (B' - B))" by (simp (no_asm_simp) add: PSP Int_ac) lemma PSP_Unless: -"[| F \ A \w A'; F \ B Unless B'|]==> F:(A \ B) \w ((A' \ B) \ B')" +"\F \ A \w A'; F \ B Unless B'\\ F:(A \ B) \w ((A' \ B) \ B')" apply (unfold op_Unless_def) apply (drule PSP, assumption) apply (blast intro: LeadsTo_Diff LeadsTo_weaken subset_imp_LeadsTo) @@ -264,11 +264,11 @@ (*** Induction rules ***) (** Meta or object quantifier ????? **) -lemma LeadsTo_wf_induct: "[| wf(r); +lemma LeadsTo_wf_induct: "\wf(r); \m \ I. F \ (A \ f-``{m}) \w ((A \ f-``(converse(r) `` {m})) \ B); - field(r)<=I; A<=f-``I; F \ program |] - ==> F \ A \w B" + field(r)<=I; A<=f-``I; F \ program\ + \ F \ A \w B" apply (simp (no_asm_use) add: LeadsTo_def) apply auto apply (erule_tac I = I and f = f in leadsTo_wf_induct, safe) @@ -278,8 +278,8 @@ done -lemma LessThan_induct: "[| \m \ nat. F:(A \ f-``{m}) \w ((A \ f-``m) \ B); - A<=f-``nat; F \ program |] ==> F \ A \w B" +lemma LessThan_induct: "\\m \ nat. F:(A \ f-``{m}) \w ((A \ f-``m) \ B); + A<=f-``nat; F \ program\ \ F \ A \w B" apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN LeadsTo_wf_induct]) apply (simp_all add: nat_measure_field) apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric]) @@ -297,16 +297,16 @@ (*** Completion \ Binary and General Finite versions ***) -lemma Completion: "[| F \ A \w (A' \ C); F \ A' Co (A' \ C); - F \ B \w (B' \ C); F \ B' Co (B' \ C) |] - ==> F \ (A \ B) \w ((A' \ B') \ C)" +lemma Completion: "\F \ A \w (A' \ C); F \ A' Co (A' \ C); + F \ B \w (B' \ C); F \ B' Co (B' \ C)\ + \ F \ (A \ B) \w ((A' \ B') \ C)" apply (simp (no_asm_use) add: LeadsTo_def Constrains_eq_constrains Int_Un_distrib) apply (blast intro: completion leadsTo_weaken) done lemma Finite_completion_aux: - "[| I \ Fin(X);F \ program |] - ==> (\i \ I. F \ (A(i)) \w (A'(i) \ C)) \ + "\I \ Fin(X);F \ program\ + \ (\i \ I. F \ (A(i)) \w (A'(i) \ C)) \ (\i \ I. F \ (A'(i)) Co (A'(i) \ C)) \ F \ (\i \ I. A(i)) \w ((\i \ I. A'(i)) \ C)" apply (erule Fin_induct) @@ -317,16 +317,16 @@ done lemma Finite_completion: - "[| I \ Fin(X); !!i. i \ I ==> F \ A(i) \w (A'(i) \ C); - !!i. i \ I ==> F \ A'(i) Co (A'(i) \ C); - F \ program |] - ==> F \ (\i \ I. A(i)) \w ((\i \ I. A'(i)) \ C)" + "\I \ Fin(X); \i. i \ I \ F \ A(i) \w (A'(i) \ C); + \i. i \ I \ F \ A'(i) Co (A'(i) \ C); + F \ program\ + \ F \ (\i \ I. A(i)) \w ((\i \ I. A'(i)) \ C)" by (blast intro: Finite_completion_aux [THEN mp, THEN mp]) lemma Stable_completion: - "[| F \ A \w A'; F \ Stable(A'); - F \ B \w B'; F \ Stable(B') |] - ==> F \ (A \ B) \w (A' \ B')" + "\F \ A \w A'; F \ Stable(A'); + F \ B \w B'; F \ Stable(B')\ + \ F \ (A \ B) \w (A' \ B')" apply (unfold Stable_def) apply (rule_tac C1 = 0 in Completion [THEN LeadsTo_weaken_R]) prefer 5 @@ -335,10 +335,10 @@ done lemma Finite_stable_completion: - "[| I \ Fin(X); - (!!i. i \ I ==> F \ A(i) \w A'(i)); - (!!i. i \ I ==>F \ Stable(A'(i))); F \ program |] - ==> F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" + "\I \ Fin(X); + (\i. i \ I \ F \ A(i) \w A'(i)); + (\i. i \ I \F \ Stable(A'(i))); F \ program\ + \ F \ (\i \ I. A(i)) \w (\i \ I. A'(i))" apply (unfold Stable_def) apply (rule_tac C1 = 0 in Finite_completion [THEN LeadsTo_weaken_R], simp_all) apply (rule_tac [3] subset_refl, auto)