| author | wenzelm | 
| Thu, 08 Sep 2022 22:19:42 +0200 | |
| changeset 76092 | 282f5e980a67 | 
| parent 69597 | ff784d5a5bfb | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 60587 | 1  | 
(* Title: HOL/TLA/Action.thy  | 
| 35108 | 2  | 
Author: Stephan Merz  | 
3  | 
Copyright: 1998 University of Munich  | 
|
| 21624 | 4  | 
*)  | 
| 3807 | 5  | 
|
| 60592 | 6  | 
section \<open>The action level of TLA as an Isabelle theory\<close>  | 
| 3807 | 7  | 
|
| 17309 | 8  | 
theory Action  | 
9  | 
imports Stfun  | 
|
10  | 
begin  | 
|
11  | 
||
| 62146 | 12  | 
type_synonym 'a trfun = "state \<times> state \<Rightarrow> 'a"  | 
13  | 
type_synonym action = "bool trfun"  | 
|
| 6255 | 14  | 
|
| 
55382
 
9218fa411c15
prefer vacuous definitional type classes over axiomatic ones;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
15  | 
instance prod :: (world, world) world ..  | 
| 3807 | 16  | 
|
| 62146 | 17  | 
definition enabled :: "action \<Rightarrow> stpred"  | 
18  | 
where "enabled A s \<equiv> \<exists>u. (s,u) \<Turnstile> A"  | 
|
19  | 
||
| 6255 | 20  | 
|
| 62146 | 21  | 
consts  | 
22  | 
before :: "'a stfun \<Rightarrow> 'a trfun"  | 
|
23  | 
after :: "'a stfun \<Rightarrow> 'a trfun"  | 
|
24  | 
unch :: "'a stfun \<Rightarrow> action"  | 
|
| 6255 | 25  | 
|
26  | 
syntax  | 
|
27  | 
(* Syntax for writing action expressions in arbitrary contexts *)  | 
|
| 60588 | 28  | 
  "_ACT"        :: "lift \<Rightarrow> 'a"                      ("(ACT _)")
 | 
| 3807 | 29  | 
|
| 60588 | 30  | 
  "_before"     :: "lift \<Rightarrow> lift"                    ("($_)"  [100] 99)
 | 
31  | 
  "_after"      :: "lift \<Rightarrow> lift"                    ("(_$)"  [100] 99)
 | 
|
32  | 
  "_unchanged"  :: "lift \<Rightarrow> lift"                    ("(unchanged _)" [100] 99)
 | 
|
| 6255 | 33  | 
|
34  | 
(*** Priming: same as "after" ***)  | 
|
| 60588 | 35  | 
  "_prime"      :: "lift \<Rightarrow> lift"                    ("(_`)" [100] 99)
 | 
| 6255 | 36  | 
|
| 60588 | 37  | 
  "_Enabled"    :: "lift \<Rightarrow> lift"                    ("(Enabled _)" [100] 100)
 | 
| 3807 | 38  | 
|
| 6255 | 39  | 
translations  | 
| 60588 | 40  | 
"ACT A" => "(A::state*state \<Rightarrow> _)"  | 
| 35108 | 41  | 
"_before" == "CONST before"  | 
42  | 
"_after" == "CONST after"  | 
|
| 
9517
 
f58863b1406a
tuned version by Stephan Merz (unbatchified etc.);
 
wenzelm 
parents: 
6255 
diff
changeset
 | 
43  | 
"_prime" => "_after"  | 
| 35108 | 44  | 
"_unchanged" == "CONST unch"  | 
45  | 
"_Enabled" == "CONST enabled"  | 
|
| 60591 | 46  | 
"s \<Turnstile> Enabled A" <= "_Enabled A s"  | 
47  | 
"w \<Turnstile> unchanged f" <= "_unchanged f w"  | 
|
| 3807 | 48  | 
|
| 47968 | 49  | 
axiomatization where  | 
| 60588 | 50  | 
unl_before: "(ACT $v) (s,t) \<equiv> v s" and  | 
51  | 
unl_after: "(ACT v$) (s,t) \<equiv> v t" and  | 
|
52  | 
unchanged_def: "(s,t) \<Turnstile> unchanged v \<equiv> (v t = v s)"  | 
|
| 47968 | 53  | 
|
| 62146 | 54  | 
|
55  | 
definition SqAct :: "[action, 'a stfun] \<Rightarrow> action"  | 
|
56  | 
where square_def: "SqAct A v \<equiv> ACT (A \<or> unchanged v)"  | 
|
57  | 
||
58  | 
definition AnAct :: "[action, 'a stfun] \<Rightarrow> action"  | 
|
59  | 
where angle_def: "AnAct A v \<equiv> ACT (A \<and> \<not> unchanged v)"  | 
|
| 3807 | 60  | 
|
| 62146 | 61  | 
syntax  | 
62  | 
  "_SqAct" :: "[lift, lift] \<Rightarrow> lift"  ("([_]'_(_))" [0, 1000] 99)
 | 
|
63  | 
  "_AnAct" :: "[lift, lift] \<Rightarrow> lift"  ("(<_>'_(_))" [0, 1000] 99)
 | 
|
64  | 
translations  | 
|
65  | 
"_SqAct" == "CONST SqAct"  | 
|
66  | 
"_AnAct" == "CONST AnAct"  | 
|
67  | 
"w \<Turnstile> [A]_v" \<leftharpoondown> "_SqAct A v w"  | 
|
68  | 
"w \<Turnstile> <A>_v" \<leftharpoondown> "_AnAct A v w"  | 
|
| 17309 | 69  | 
|
| 21624 | 70  | 
|
71  | 
(* The following assertion specializes "intI" for any world type  | 
|
72  | 
which is a pair, not just for "state * state".  | 
|
73  | 
*)  | 
|
74  | 
||
75  | 
lemma actionI [intro!]:  | 
|
| 60588 | 76  | 
assumes "\<And>s t. (s,t) \<Turnstile> A"  | 
77  | 
shows "\<turnstile> A"  | 
|
| 
27104
 
791607529f6d
rep_datatype command now takes list of constructors as input arguments
 
haftmann 
parents: 
24180 
diff
changeset
 | 
78  | 
apply (rule assms intI prod.induct)+  | 
| 21624 | 79  | 
done  | 
80  | 
||
| 60588 | 81  | 
lemma actionD [dest]: "\<turnstile> A \<Longrightarrow> (s,t) \<Turnstile> A"  | 
| 21624 | 82  | 
apply (erule intD)  | 
83  | 
done  | 
|
84  | 
||
85  | 
lemma pr_rews [int_rewrite]:  | 
|
| 60588 | 86  | 
"\<turnstile> (#c)` = #c"  | 
87  | 
"\<And>f. \<turnstile> f<x>` = f<x` >"  | 
|
88  | 
"\<And>f. \<turnstile> f<x,y>` = f<x`,y` >"  | 
|
89  | 
"\<And>f. \<turnstile> f<x,y,z>` = f<x`,y`,z` >"  | 
|
90  | 
"\<turnstile> (\<forall>x. P x)` = (\<forall>x. (P x)`)"  | 
|
91  | 
"\<turnstile> (\<exists>x. P x)` = (\<exists>x. (P x)`)"  | 
|
| 21624 | 92  | 
by (rule actionI, unfold unl_after intensional_rews, rule refl)+  | 
93  | 
||
94  | 
||
95  | 
lemmas act_rews [simp] = unl_before unl_after unchanged_def pr_rews  | 
|
96  | 
||
97  | 
lemmas action_rews = act_rews intensional_rews  | 
|
98  | 
||
99  | 
||
100  | 
(* ================ Functions to "unlift" action theorems into HOL rules ================ *)  | 
|
101  | 
||
| 60592 | 102  | 
ML \<open>  | 
| 21624 | 103  | 
(* The following functions are specialized versions of the corresponding  | 
104  | 
functions defined in Intensional.ML in that they introduce a  | 
|
105  | 
"world" parameter of the form (s,t) and apply additional rewrites.  | 
|
106  | 
*)  | 
|
107  | 
||
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
108  | 
fun action_unlift ctxt th =  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
109  | 
  (rewrite_rule ctxt @{thms action_rews} (th RS @{thm actionD}))
 | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
110  | 
handle THM _ => int_unlift ctxt th;  | 
| 21624 | 111  | 
|
| 60588 | 112  | 
(* Turn \<turnstile> A = B into meta-level rewrite rule A == B *)  | 
| 21624 | 113  | 
val action_rewrite = int_rewrite  | 
114  | 
||
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
115  | 
fun action_use ctxt th =  | 
| 59582 | 116  | 
case Thm.concl_of th of  | 
| 69597 | 117  | 
Const _ $ (Const (\<^const_name>\<open>Valid\<close>, _) $ _) =>  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
118  | 
(flatten (action_unlift ctxt th) handle THM _ => th)  | 
| 21624 | 119  | 
| _ => th;  | 
| 60592 | 120  | 
\<close>  | 
| 21624 | 121  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
122  | 
attribute_setup action_unlift =  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
60592 
diff
changeset
 | 
123  | 
\<open>Scan.succeed (Thm.rule_attribute [] (action_unlift o Context.proof_of))\<close>  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
124  | 
attribute_setup action_rewrite =  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
60592 
diff
changeset
 | 
125  | 
\<open>Scan.succeed (Thm.rule_attribute [] (action_rewrite o Context.proof_of))\<close>  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
126  | 
attribute_setup action_use =  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
60592 
diff
changeset
 | 
127  | 
\<open>Scan.succeed (Thm.rule_attribute [] (action_use o Context.proof_of))\<close>  | 
| 21624 | 128  | 
|
129  | 
||
130  | 
(* =========================== square / angle brackets =========================== *)  | 
|
131  | 
||
| 60588 | 132  | 
lemma idle_squareI: "(s,t) \<Turnstile> unchanged v \<Longrightarrow> (s,t) \<Turnstile> [A]_v"  | 
| 21624 | 133  | 
by (simp add: square_def)  | 
134  | 
||
| 60588 | 135  | 
lemma busy_squareI: "(s,t) \<Turnstile> A \<Longrightarrow> (s,t) \<Turnstile> [A]_v"  | 
| 21624 | 136  | 
by (simp add: square_def)  | 
| 60587 | 137  | 
|
| 21624 | 138  | 
lemma squareE [elim]:  | 
| 60588 | 139  | 
"\<lbrakk> (s,t) \<Turnstile> [A]_v; A (s,t) \<Longrightarrow> B (s,t); v t = v s \<Longrightarrow> B (s,t) \<rbrakk> \<Longrightarrow> B (s,t)"  | 
| 21624 | 140  | 
apply (unfold square_def action_rews)  | 
141  | 
apply (erule disjE)  | 
|
142  | 
apply simp_all  | 
|
143  | 
done  | 
|
144  | 
||
| 60588 | 145  | 
lemma squareCI [intro]: "\<lbrakk> v t \<noteq> v s \<Longrightarrow> A (s,t) \<rbrakk> \<Longrightarrow> (s,t) \<Turnstile> [A]_v"  | 
| 21624 | 146  | 
apply (unfold square_def action_rews)  | 
147  | 
apply (rule disjCI)  | 
|
148  | 
apply (erule (1) meta_mp)  | 
|
149  | 
done  | 
|
150  | 
||
| 60588 | 151  | 
lemma angleI [intro]: "\<And>s t. \<lbrakk> A (s,t); v t \<noteq> v s \<rbrakk> \<Longrightarrow> (s,t) \<Turnstile> <A>_v"  | 
| 21624 | 152  | 
by (simp add: angle_def)  | 
153  | 
||
| 60588 | 154  | 
lemma angleE [elim]: "\<lbrakk> (s,t) \<Turnstile> <A>_v; \<lbrakk> A (s,t); v t \<noteq> v s \<rbrakk> \<Longrightarrow> R \<rbrakk> \<Longrightarrow> R"  | 
| 21624 | 155  | 
apply (unfold angle_def action_rews)  | 
156  | 
apply (erule conjE)  | 
|
157  | 
apply simp  | 
|
158  | 
done  | 
|
159  | 
||
160  | 
lemma square_simulation:  | 
|
| 60591 | 161  | 
"\<And>f. \<lbrakk> \<turnstile> unchanged f \<and> \<not>B \<longrightarrow> unchanged g;  | 
162  | 
\<turnstile> A \<and> \<not>unchanged g \<longrightarrow> B  | 
|
| 60588 | 163  | 
\<rbrakk> \<Longrightarrow> \<turnstile> [A]_f \<longrightarrow> [B]_g"  | 
| 21624 | 164  | 
apply clarsimp  | 
165  | 
apply (erule squareE)  | 
|
166  | 
apply (auto simp add: square_def)  | 
|
167  | 
done  | 
|
168  | 
||
| 60588 | 169  | 
lemma not_square: "\<turnstile> (\<not> [A]_v) = <\<not>A>_v"  | 
| 21624 | 170  | 
by (auto simp: square_def angle_def)  | 
171  | 
||
| 60588 | 172  | 
lemma not_angle: "\<turnstile> (\<not> <A>_v) = [\<not>A]_v"  | 
| 21624 | 173  | 
by (auto simp: square_def angle_def)  | 
174  | 
||
175  | 
||
176  | 
(* ============================== Facts about ENABLED ============================== *)  | 
|
177  | 
||
| 60588 | 178  | 
lemma enabledI: "\<turnstile> A \<longrightarrow> $Enabled A"  | 
| 21624 | 179  | 
by (auto simp add: enabled_def)  | 
180  | 
||
| 60588 | 181  | 
lemma enabledE: "\<lbrakk> s \<Turnstile> Enabled A; \<And>u. A (s,u) \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> Q"  | 
| 21624 | 182  | 
apply (unfold enabled_def)  | 
183  | 
apply (erule exE)  | 
|
184  | 
apply simp  | 
|
185  | 
done  | 
|
186  | 
||
| 60588 | 187  | 
lemma notEnabledD: "\<turnstile> \<not>$Enabled G \<longrightarrow> \<not> G"  | 
| 21624 | 188  | 
by (auto simp add: enabled_def)  | 
189  | 
||
190  | 
(* Monotonicity *)  | 
|
191  | 
lemma enabled_mono:  | 
|
| 60588 | 192  | 
assumes min: "s \<Turnstile> Enabled F"  | 
193  | 
and maj: "\<turnstile> F \<longrightarrow> G"  | 
|
194  | 
shows "s \<Turnstile> Enabled G"  | 
|
| 21624 | 195  | 
apply (rule min [THEN enabledE])  | 
196  | 
apply (rule enabledI [action_use])  | 
|
197  | 
apply (erule maj [action_use])  | 
|
198  | 
done  | 
|
199  | 
||
200  | 
(* stronger variant *)  | 
|
201  | 
lemma enabled_mono2:  | 
|
| 60588 | 202  | 
assumes min: "s \<Turnstile> Enabled F"  | 
203  | 
and maj: "\<And>t. F (s,t) \<Longrightarrow> G (s,t)"  | 
|
204  | 
shows "s \<Turnstile> Enabled G"  | 
|
| 21624 | 205  | 
apply (rule min [THEN enabledE])  | 
206  | 
apply (rule enabledI [action_use])  | 
|
207  | 
apply (erule maj)  | 
|
208  | 
done  | 
|
209  | 
||
| 60591 | 210  | 
lemma enabled_disj1: "\<turnstile> Enabled F \<longrightarrow> Enabled (F \<or> G)"  | 
| 21624 | 211  | 
by (auto elim!: enabled_mono)  | 
212  | 
||
| 60591 | 213  | 
lemma enabled_disj2: "\<turnstile> Enabled G \<longrightarrow> Enabled (F \<or> G)"  | 
| 21624 | 214  | 
by (auto elim!: enabled_mono)  | 
215  | 
||
| 60591 | 216  | 
lemma enabled_conj1: "\<turnstile> Enabled (F \<and> G) \<longrightarrow> Enabled F"  | 
| 21624 | 217  | 
by (auto elim!: enabled_mono)  | 
218  | 
||
| 60591 | 219  | 
lemma enabled_conj2: "\<turnstile> Enabled (F \<and> G) \<longrightarrow> Enabled G"  | 
| 21624 | 220  | 
by (auto elim!: enabled_mono)  | 
221  | 
||
222  | 
lemma enabled_conjE:  | 
|
| 60591 | 223  | 
"\<lbrakk> s \<Turnstile> Enabled (F \<and> G); \<lbrakk> s \<Turnstile> Enabled F; s \<Turnstile> Enabled G \<rbrakk> \<Longrightarrow> Q \<rbrakk> \<Longrightarrow> Q"  | 
| 21624 | 224  | 
apply (frule enabled_conj1 [action_use])  | 
225  | 
apply (drule enabled_conj2 [action_use])  | 
|
226  | 
apply simp  | 
|
227  | 
done  | 
|
228  | 
||
| 60591 | 229  | 
lemma enabled_disjD: "\<turnstile> Enabled (F \<or> G) \<longrightarrow> Enabled F \<or> Enabled G"  | 
| 21624 | 230  | 
by (auto simp add: enabled_def)  | 
231  | 
||
| 60591 | 232  | 
lemma enabled_disj: "\<turnstile> Enabled (F \<or> G) = (Enabled F \<or> Enabled G)"  | 
| 21624 | 233  | 
apply clarsimp  | 
234  | 
apply (rule iffI)  | 
|
235  | 
apply (erule enabled_disjD [action_use])  | 
|
236  | 
apply (erule disjE enabled_disj1 [action_use] enabled_disj2 [action_use])+  | 
|
237  | 
done  | 
|
238  | 
||
| 60588 | 239  | 
lemma enabled_ex: "\<turnstile> Enabled (\<exists>x. F x) = (\<exists>x. Enabled (F x))"  | 
| 21624 | 240  | 
by (force simp add: enabled_def)  | 
241  | 
||
242  | 
||
243  | 
(* A rule that combines enabledI and baseE, but generates fewer instantiations *)  | 
|
244  | 
lemma base_enabled:  | 
|
| 60588 | 245  | 
"\<lbrakk> basevars vs; \<exists>c. \<forall>u. vs u = c \<longrightarrow> A(s,u) \<rbrakk> \<Longrightarrow> s \<Turnstile> Enabled A"  | 
| 21624 | 246  | 
apply (erule exE)  | 
247  | 
apply (erule baseE)  | 
|
248  | 
apply (rule enabledI [action_use])  | 
|
249  | 
apply (erule allE)  | 
|
250  | 
apply (erule mp)  | 
|
251  | 
apply assumption  | 
|
252  | 
done  | 
|
253  | 
||
254  | 
(* ======================= action_simp_tac ============================== *)  | 
|
255  | 
||
| 60592 | 256  | 
ML \<open>  | 
| 21624 | 257  | 
(* A dumb simplification-based tactic with just a little first-order logic:  | 
258  | 
should plug in only "very safe" rules that can be applied blindly.  | 
|
259  | 
Note that it applies whatever simplifications are currently active.  | 
|
260  | 
*)  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52037 
diff
changeset
 | 
261  | 
fun action_simp_tac ctxt intros elims =  | 
| 21624 | 262  | 
asm_full_simp_tac  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
263  | 
(ctxt setloop (fn _ => (resolve_tac ctxt ((map (action_use ctxt) intros)  | 
| 24180 | 264  | 
                                    @ [refl,impI,conjI,@{thm actionI},@{thm intI},allI]))
 | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
265  | 
ORELSE' (eresolve_tac ctxt ((map (action_use ctxt) elims)  | 
| 21624 | 266  | 
@ [conjE,disjE,exE]))));  | 
| 60592 | 267  | 
\<close>  | 
| 21624 | 268  | 
|
269  | 
(* ---------------- enabled_tac: tactic to prove (Enabled A) -------------------- *)  | 
|
270  | 
||
| 60592 | 271  | 
ML \<open>  | 
| 21624 | 272  | 
(* "Enabled A" can be proven as follows:  | 
273  | 
- Assume that we know which state variables are "base variables"  | 
|
274  | 
this should be expressed by a theorem of the form "basevars (x,y,z,...)".  | 
|
275  | 
- Resolve this theorem with baseE to introduce a constant for the value of the  | 
|
276  | 
variables in the successor state, and resolve the goal with the result.  | 
|
277  | 
- Resolve with enabledI and do some rewriting.  | 
|
278  | 
- Solve for the unknowns using standard HOL reasoning.  | 
|
279  | 
The following tactic combines these steps except the final one.  | 
|
280  | 
*)  | 
|
281  | 
||
| 42785 | 282  | 
fun enabled_tac ctxt base_vars =  | 
| 42793 | 283  | 
  clarsimp_tac (ctxt addSIs [base_vars RS @{thm base_enabled}]);
 | 
| 60592 | 284  | 
\<close>  | 
| 21624 | 285  | 
|
| 60592 | 286  | 
method_setup enabled = \<open>  | 
| 42785 | 287  | 
Attrib.thm >> (fn th => fn ctxt => SIMPLE_METHOD' (enabled_tac ctxt th))  | 
| 60592 | 288  | 
\<close>  | 
| 42785 | 289  | 
|
| 21624 | 290  | 
(* Example *)  | 
291  | 
||
292  | 
lemma  | 
|
293  | 
assumes "basevars (x,y,z)"  | 
|
| 60591 | 294  | 
shows "\<turnstile> x \<longrightarrow> Enabled ($x \<and> (y$ = #False))"  | 
| 42785 | 295  | 
apply (enabled assms)  | 
| 21624 | 296  | 
apply auto  | 
297  | 
done  | 
|
298  | 
||
299  | 
end  |