src/HOL/UNITY/SubstAx.thy
changeset 63146 f1ecba0272f9
parent 62343 24106dc44def
child 67613 ce654b0e6d69
equal deleted inserted replaced
63145:703edebd1d92 63146:f1ecba0272f9
     3     Copyright   1998  University of Cambridge
     3     Copyright   1998  University of Cambridge
     4 
     4 
     5 Weak LeadsTo relation (restricted to the set of reachable states)
     5 Weak LeadsTo relation (restricted to the set of reachable states)
     6 *)
     6 *)
     7 
     7 
     8 section{*Weak Progress*}
     8 section\<open>Weak Progress\<close>
     9 
     9 
    10 theory SubstAx imports WFair Constrains begin
    10 theory SubstAx imports WFair Constrains begin
    11 
    11 
    12 definition Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) where
    12 definition Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) where
    13     "A Ensures B == {F. F \<in> (reachable F \<inter> A) ensures B}"
    13     "A Ensures B == {F. F \<in> (reachable F \<inter> A) ensures B}"
    16     "A LeadsTo B == {F. F \<in> (reachable F \<inter> A) leadsTo B}"
    16     "A LeadsTo B == {F. F \<in> (reachable F \<inter> A) leadsTo B}"
    17 
    17 
    18 notation LeadsTo  (infixl "\<longmapsto>w" 60)
    18 notation LeadsTo  (infixl "\<longmapsto>w" 60)
    19 
    19 
    20 
    20 
    21 text{*Resembles the previous definition of LeadsTo*}
    21 text\<open>Resembles the previous definition of LeadsTo\<close>
    22 lemma LeadsTo_eq_leadsTo: 
    22 lemma LeadsTo_eq_leadsTo: 
    23      "A LeadsTo B = {F. F \<in> (reachable F \<inter> A) leadsTo (reachable F \<inter> B)}"
    23      "A LeadsTo B = {F. F \<in> (reachable F \<inter> A) leadsTo (reachable F \<inter> B)}"
    24 apply (unfold LeadsTo_def)
    24 apply (unfold LeadsTo_def)
    25 apply (blast dest: psp_stable2 intro: leadsTo_weaken)
    25 apply (blast dest: psp_stable2 intro: leadsTo_weaken)
    26 done
    26 done
    27 
    27 
    28 
    28 
    29 subsection{*Specialized laws for handling invariants*}
    29 subsection\<open>Specialized laws for handling invariants\<close>
    30 
    30 
    31 (** Conjoining an Always property **)
    31 (** Conjoining an Always property **)
    32 
    32 
    33 lemma Always_LeadsTo_pre:
    33 lemma Always_LeadsTo_pre:
    34      "F \<in> Always INV ==> (F \<in> (INV \<inter> A) LeadsTo A') = (F \<in> A LeadsTo A')"
    34      "F \<in> Always INV ==> (F \<in> (INV \<inter> A) LeadsTo A') = (F \<in> A LeadsTo A')"
    45 
    45 
    46 (* [| F \<in> Always INV;  F \<in> A LeadsTo A' |] ==> F \<in> A LeadsTo (INV \<inter> A') *)
    46 (* [| F \<in> Always INV;  F \<in> A LeadsTo A' |] ==> F \<in> A LeadsTo (INV \<inter> A') *)
    47 lemmas Always_LeadsToD = Always_LeadsTo_post [THEN iffD2]
    47 lemmas Always_LeadsToD = Always_LeadsTo_post [THEN iffD2]
    48 
    48 
    49 
    49 
    50 subsection{*Introduction rules: Basis, Trans, Union*}
    50 subsection\<open>Introduction rules: Basis, Trans, Union\<close>
    51 
    51 
    52 lemma leadsTo_imp_LeadsTo: "F \<in> A leadsTo B ==> F \<in> A LeadsTo B"
    52 lemma leadsTo_imp_LeadsTo: "F \<in> A leadsTo B ==> F \<in> A LeadsTo B"
    53 apply (simp add: LeadsTo_def)
    53 apply (simp add: LeadsTo_def)
    54 apply (blast intro: leadsTo_weaken_L)
    54 apply (blast intro: leadsTo_weaken_L)
    55 done
    55 done
    66 apply (subst Int_Union)
    66 apply (subst Int_Union)
    67 apply (blast intro: leadsTo_UN)
    67 apply (blast intro: leadsTo_UN)
    68 done
    68 done
    69 
    69 
    70 
    70 
    71 subsection{*Derived rules*}
    71 subsection\<open>Derived rules\<close>
    72 
    72 
    73 lemma LeadsTo_UNIV [simp]: "F \<in> A LeadsTo UNIV"
    73 lemma LeadsTo_UNIV [simp]: "F \<in> A LeadsTo UNIV"
    74 by (simp add: LeadsTo_def)
    74 by (simp add: LeadsTo_def)
    75 
    75 
    76 text{*Useful with cancellation, disjunction*}
    76 text\<open>Useful with cancellation, disjunction\<close>
    77 lemma LeadsTo_Un_duplicate:
    77 lemma LeadsTo_Un_duplicate:
    78      "F \<in> A LeadsTo (A' \<union> A') ==> F \<in> A LeadsTo A'"
    78      "F \<in> A LeadsTo (A' \<union> A') ==> F \<in> A LeadsTo A'"
    79 by (simp add: Un_ac)
    79 by (simp add: Un_ac)
    80 
    80 
    81 lemma LeadsTo_Un_duplicate2:
    81 lemma LeadsTo_Un_duplicate2:
    85 lemma LeadsTo_UN: 
    85 lemma LeadsTo_UN: 
    86      "(!!i. i \<in> I ==> F \<in> (A i) LeadsTo B) ==> F \<in> (\<Union>i \<in> I. A i) LeadsTo B"
    86      "(!!i. i \<in> I ==> F \<in> (A i) LeadsTo B) ==> F \<in> (\<Union>i \<in> I. A i) LeadsTo B"
    87 apply (blast intro: LeadsTo_Union)
    87 apply (blast intro: LeadsTo_Union)
    88 done
    88 done
    89 
    89 
    90 text{*Binary union introduction rule*}
    90 text\<open>Binary union introduction rule\<close>
    91 lemma LeadsTo_Un:
    91 lemma LeadsTo_Un:
    92      "[| F \<in> A LeadsTo C; F \<in> B LeadsTo C |] ==> F \<in> (A \<union> B) LeadsTo C"
    92      "[| F \<in> A LeadsTo C; F \<in> B LeadsTo C |] ==> F \<in> (A \<union> B) LeadsTo C"
    93   using LeadsTo_UN [of "{A, B}" F id C] by auto
    93   using LeadsTo_UN [of "{A, B}" F id C] by auto
    94 
    94 
    95 text{*Lets us look at the starting state*}
    95 text\<open>Lets us look at the starting state\<close>
    96 lemma single_LeadsTo_I:
    96 lemma single_LeadsTo_I:
    97      "(!!s. s \<in> A ==> F \<in> {s} LeadsTo B) ==> F \<in> A LeadsTo B"
    97      "(!!s. s \<in> A ==> F \<in> {s} LeadsTo B) ==> F \<in> A LeadsTo B"
    98 by (subst UN_singleton [symmetric], rule LeadsTo_UN, blast)
    98 by (subst UN_singleton [symmetric], rule LeadsTo_UN, blast)
    99 
    99 
   100 lemma subset_imp_LeadsTo: "A \<subseteq> B ==> F \<in> A LeadsTo B"
   100 lemma subset_imp_LeadsTo: "A \<subseteq> B ==> F \<in> A LeadsTo B"
   174   ==> F \<in> A LeadsTo A'"
   174   ==> F \<in> A LeadsTo A'"
   175 apply (rule Always_LeadsToI, assumption)
   175 apply (rule Always_LeadsToI, assumption)
   176 apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen)
   176 apply (blast intro: EnsuresI LeadsTo_Basis Always_ConstrainsD [THEN Constrains_weaken] transient_strengthen)
   177 done
   177 done
   178 
   178 
   179 text{*Set difference: maybe combine with @{text leadsTo_weaken_L}??
   179 text\<open>Set difference: maybe combine with \<open>leadsTo_weaken_L\<close>??
   180   This is the most useful form of the "disjunction" rule*}
   180   This is the most useful form of the "disjunction" rule\<close>
   181 lemma LeadsTo_Diff:
   181 lemma LeadsTo_Diff:
   182      "[| F \<in> (A-B) LeadsTo C;  F \<in> (A \<inter> B) LeadsTo C |]  
   182      "[| F \<in> (A-B) LeadsTo C;  F \<in> (A \<inter> B) LeadsTo C |]  
   183       ==> F \<in> A LeadsTo C"
   183       ==> F \<in> A LeadsTo C"
   184 by (blast intro: LeadsTo_Un LeadsTo_weaken)
   184 by (blast intro: LeadsTo_Un LeadsTo_weaken)
   185 
   185 
   189       ==> F \<in> (\<Union>i \<in> I. A i) LeadsTo (\<Union>i \<in> I. A' i)"
   189       ==> F \<in> (\<Union>i \<in> I. A i) LeadsTo (\<Union>i \<in> I. A' i)"
   190 apply (blast intro: LeadsTo_Union LeadsTo_weaken_R)
   190 apply (blast intro: LeadsTo_Union LeadsTo_weaken_R)
   191 done
   191 done
   192 
   192 
   193 
   193 
   194 text{*Version with no index set*}
   194 text\<open>Version with no index set\<close>
   195 lemma LeadsTo_UN_UN_noindex: 
   195 lemma LeadsTo_UN_UN_noindex: 
   196      "(!!i. F \<in> (A i) LeadsTo (A' i)) ==> F \<in> (\<Union>i. A i) LeadsTo (\<Union>i. A' i)"
   196      "(!!i. F \<in> (A i) LeadsTo (A' i)) ==> F \<in> (\<Union>i. A i) LeadsTo (\<Union>i. A' i)"
   197 by (blast intro: LeadsTo_UN_UN)
   197 by (blast intro: LeadsTo_UN_UN)
   198 
   198 
   199 text{*Version with no index set*}
   199 text\<open>Version with no index set\<close>
   200 lemma all_LeadsTo_UN_UN:
   200 lemma all_LeadsTo_UN_UN:
   201      "\<forall>i. F \<in> (A i) LeadsTo (A' i)  
   201      "\<forall>i. F \<in> (A i) LeadsTo (A' i)  
   202       ==> F \<in> (\<Union>i. A i) LeadsTo (\<Union>i. A' i)"
   202       ==> F \<in> (\<Union>i. A i) LeadsTo (\<Union>i. A' i)"
   203 by (blast intro: LeadsTo_UN_UN)
   203 by (blast intro: LeadsTo_UN_UN)
   204 
   204 
   205 text{*Binary union version*}
   205 text\<open>Binary union version\<close>
   206 lemma LeadsTo_Un_Un:
   206 lemma LeadsTo_Un_Un:
   207      "[| F \<in> A LeadsTo A'; F \<in> B LeadsTo B' |]  
   207      "[| F \<in> A LeadsTo A'; F \<in> B LeadsTo B' |]  
   208             ==> F \<in> (A \<union> B) LeadsTo (A' \<union> B')"
   208             ==> F \<in> (A \<union> B) LeadsTo (A' \<union> B')"
   209 by (blast intro: LeadsTo_Un LeadsTo_weaken_R)
   209 by (blast intro: LeadsTo_Un LeadsTo_weaken_R)
   210 
   210 
   238 prefer 2 apply assumption
   238 prefer 2 apply assumption
   239 apply (simp_all (no_asm_simp))
   239 apply (simp_all (no_asm_simp))
   240 done
   240 done
   241 
   241 
   242 
   242 
   243 text{*The impossibility law*}
   243 text\<open>The impossibility law\<close>
   244 
   244 
   245 text{*The set "A" may be non-empty, but it contains no reachable states*}
   245 text\<open>The set "A" may be non-empty, but it contains no reachable states\<close>
   246 lemma LeadsTo_empty: "[|F \<in> A LeadsTo {}; all_total F|] ==> F \<in> Always (-A)"
   246 lemma LeadsTo_empty: "[|F \<in> A LeadsTo {}; all_total F|] ==> F \<in> Always (-A)"
   247 apply (simp add: LeadsTo_def Always_eq_includes_reachable)
   247 apply (simp add: LeadsTo_def Always_eq_includes_reachable)
   248 apply (drule leadsTo_empty, auto)
   248 apply (drule leadsTo_empty, auto)
   249 done
   249 done
   250 
   250 
   251 
   251 
   252 subsection{*PSP: Progress-Safety-Progress*}
   252 subsection\<open>PSP: Progress-Safety-Progress\<close>
   253 
   253 
   254 text{*Special case of PSP: Misra's "stable conjunction"*}
   254 text\<open>Special case of PSP: Misra's "stable conjunction"\<close>
   255 lemma PSP_Stable:
   255 lemma PSP_Stable:
   256      "[| F \<in> A LeadsTo A';  F \<in> Stable B |]  
   256      "[| F \<in> A LeadsTo A';  F \<in> Stable B |]  
   257       ==> F \<in> (A \<inter> B) LeadsTo (A' \<inter> B)"
   257       ==> F \<in> (A \<inter> B) LeadsTo (A' \<inter> B)"
   258 apply (simp add: LeadsTo_eq_leadsTo Stable_eq_stable)
   258 apply (simp add: LeadsTo_eq_leadsTo Stable_eq_stable)
   259 apply (drule psp_stable, assumption)
   259 apply (drule psp_stable, assumption)
   296           transient_imp_leadsTo [THEN leadsTo_imp_LeadsTo, THEN PSP_Stable2])
   296           transient_imp_leadsTo [THEN leadsTo_imp_LeadsTo, THEN PSP_Stable2])
   297    apply (blast intro: subset_imp_LeadsTo)+
   297    apply (blast intro: subset_imp_LeadsTo)+
   298 done
   298 done
   299 
   299 
   300 
   300 
   301 subsection{*Induction rules*}
   301 subsection\<open>Induction rules\<close>
   302 
   302 
   303 (** Meta or object quantifier ????? **)
   303 (** Meta or object quantifier ????? **)
   304 lemma LeadsTo_wf_induct:
   304 lemma LeadsTo_wf_induct:
   305      "[| wf r;      
   305      "[| wf r;      
   306          \<forall>m. F \<in> (A \<inter> f-`{m}) LeadsTo                      
   306          \<forall>m. F \<in> (A \<inter> f-`{m}) LeadsTo                      
   327 lemma LessThan_induct:
   327 lemma LessThan_induct:
   328      "(!!m::nat. F \<in> (A \<inter> f-`{m}) LeadsTo ((A \<inter> f-`(lessThan m)) \<union> B))
   328      "(!!m::nat. F \<in> (A \<inter> f-`{m}) LeadsTo ((A \<inter> f-`(lessThan m)) \<union> B))
   329       ==> F \<in> A LeadsTo B"
   329       ==> F \<in> A LeadsTo B"
   330 by (rule wf_less_than [THEN LeadsTo_wf_induct], auto)
   330 by (rule wf_less_than [THEN LeadsTo_wf_induct], auto)
   331 
   331 
   332 text{*Integer version.  Could generalize from 0 to any lower bound*}
   332 text\<open>Integer version.  Could generalize from 0 to any lower bound\<close>
   333 lemma integ_0_le_induct:
   333 lemma integ_0_le_induct:
   334      "[| F \<in> Always {s. (0::int) \<le> f s};   
   334      "[| F \<in> Always {s. (0::int) \<le> f s};   
   335          !! z. F \<in> (A \<inter> {s. f s = z}) LeadsTo                      
   335          !! z. F \<in> (A \<inter> {s. f s = z}) LeadsTo                      
   336                    ((A \<inter> {s. f s < z}) \<union> B) |]  
   336                    ((A \<inter> {s. f s < z}) \<union> B) |]  
   337       ==> F \<in> A LeadsTo B"
   337       ==> F \<in> A LeadsTo B"
   361  apply (blast intro: LeadsTo_weaken_R diff_less_mono2)
   361  apply (blast intro: LeadsTo_weaken_R diff_less_mono2)
   362 apply (blast intro: not_le_imp_less subset_imp_LeadsTo)
   362 apply (blast intro: not_le_imp_less subset_imp_LeadsTo)
   363 done
   363 done
   364 
   364 
   365 
   365 
   366 subsection{*Completion: Binary and General Finite versions*}
   366 subsection\<open>Completion: Binary and General Finite versions\<close>
   367 
   367 
   368 lemma Completion:
   368 lemma Completion:
   369      "[| F \<in> A LeadsTo (A' \<union> C);  F \<in> A' Co (A' \<union> C);  
   369      "[| F \<in> A LeadsTo (A' \<union> C);  F \<in> A' Co (A' \<union> C);  
   370          F \<in> B LeadsTo (B' \<union> C);  F \<in> B' Co (B' \<union> C) |]  
   370          F \<in> B LeadsTo (B' \<union> C);  F \<in> B' Co (B' \<union> C) |]  
   371       ==> F \<in> (A \<inter> B) LeadsTo ((A' \<inter> B') \<union> C)"
   371       ==> F \<in> (A \<inter> B) LeadsTo ((A' \<inter> B') \<union> C)"