diff -r f39f67982854 -r 91713a1915ee src/HOL/UNITY/SubstAx.thy --- a/src/HOL/UNITY/SubstAx.thy Sat Feb 08 14:46:22 2003 +0100 +++ b/src/HOL/UNITY/SubstAx.thy Sat Feb 08 16:05:33 2003 +0100 @@ -21,7 +21,7 @@ "op LeadsTo" :: "['a set, 'a set] => 'a program set" (infixl " \w " 60) -(*Resembles the previous definition of LeadsTo*) +text{*Resembles the previous definition of LeadsTo*} lemma LeadsTo_eq_leadsTo: "A LeadsTo B = {F. F \ (reachable F \ A) leadsTo (reachable F \ B)}" apply (unfold LeadsTo_def) @@ -76,7 +76,7 @@ lemma LeadsTo_UNIV [simp]: "F \ A LeadsTo UNIV" by (simp add: LeadsTo_def) -(*Useful with cancellation, disjunction*) +text{*Useful with cancellation, disjunction*} lemma LeadsTo_Un_duplicate: "F \ A LeadsTo (A' \ A') ==> F \ A LeadsTo A'" by (simp add: Un_ac) @@ -91,14 +91,14 @@ apply (blast intro: LeadsTo_Union) done -(*Binary union introduction rule*) +text{*Binary union introduction rule*} lemma LeadsTo_Un: "[| F \ A LeadsTo C; F \ B LeadsTo C |] ==> F \ (A \ B) LeadsTo C" apply (subst Un_eq_Union) apply (blast intro: LeadsTo_Union) done -(*Lets us look at the starting state*) +text{*Lets us look at the starting state*} lemma single_LeadsTo_I: "(!!s. s \ A ==> F \ {s} LeadsTo B) ==> F \ A LeadsTo B" by (subst UN_singleton [symmetric], rule LeadsTo_UN, blast) @@ -182,8 +182,8 @@ apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen) done -(*Set difference: maybe combine with leadsTo_weaken_L?? - This is the most useful form of the "disjunction" rule*) +text{*Set difference: maybe combine with leadsTo_weaken_L?? + This is the most useful form of the "disjunction" rule*} lemma LeadsTo_Diff: "[| F \ (A-B) LeadsTo C; F \ (A \ B) LeadsTo C |] ==> F \ A LeadsTo C" @@ -198,18 +198,18 @@ done -(*Version with no index set*) +text{*Version with no index set*} lemma LeadsTo_UN_UN_noindex: "(!!i. F \ (A i) LeadsTo (A' i)) ==> F \ (\i. A i) LeadsTo (\i. A' i)" by (blast intro: LeadsTo_UN_UN) -(*Version with no index set*) +text{*Version with no index set*} lemma all_LeadsTo_UN_UN: "\i. F \ (A i) LeadsTo (A' i) ==> F \ (\i. A i) LeadsTo (\i. A' i)" by (blast intro: LeadsTo_UN_UN) -(*Binary union version*) +text{*Binary union version*} lemma LeadsTo_Un_Un: "[| F \ A LeadsTo A'; F \ B LeadsTo B' |] ==> F \ (A \ B) LeadsTo (A' \ B')" @@ -247,18 +247,18 @@ done -(** The impossibility law **) +text{*The impossibility law*} -(*The set "A" may be non-empty, but it contains no reachable states*) -lemma LeadsTo_empty: "F \ A LeadsTo {} ==> F \ Always (-A)" +text{*The set "A" may be non-empty, but it contains no reachable states*} +lemma LeadsTo_empty: "[|F \ A LeadsTo {}; all_total F|] ==> F \ Always (-A)" apply (simp add: LeadsTo_def Always_eq_includes_reachable) apply (drule leadsTo_empty, auto) done -(** PSP: Progress-Safety-Progress **) +subsection{*PSP: Progress-Safety-Progress*} -(*Special case of PSP: Misra's "stable conjunction"*) +text{*Special case of PSP: Misra's "stable conjunction"*} lemma PSP_Stable: "[| F \ A LeadsTo A'; F \ Stable B |] ==> F \ (A \ B) LeadsTo (A' \ B)" @@ -336,7 +336,7 @@ ==> F \ A LeadsTo B" by (rule wf_less_than [THEN LeadsTo_wf_induct], auto) -(*Integer version. Could generalize from 0 to any lower bound*) +text{*Integer version. Could generalize from 0 to any lower bound*} lemma integ_0_le_induct: "[| F \ Always {s. (0::int) \ f s}; !! z. F \ (A \ {s. f s = z}) LeadsTo