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