src/HOL/TLA/TLA.thy
author bulwahn
Thu, 08 Dec 2016 17:22:51 +0100
changeset 64543 6b13586ef1a2
parent 61853 fb7756087101
child 69597 ff784d5a5bfb
permissions -rw-r--r--
remove typo in bij_swap_compose_bij theorem name; tune proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
35108
e384e27c229f modernized syntax/translations;
wenzelm
parents: 35068
diff changeset
     1
(*  Title:      HOL/TLA/TLA.thy
e384e27c229f modernized syntax/translations;
wenzelm
parents: 35068
diff changeset
     2
    Author:     Stephan Merz
e384e27c229f modernized syntax/translations;
wenzelm
parents: 35068
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 temporal level of TLA\<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: 14565
diff changeset
     8
theory TLA
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 14565
diff changeset
     9
imports Init
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 14565
diff changeset
    10
begin
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    11
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    12
consts
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    13
  (** abstract syntax **)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    14
  Box        :: "('w::world) form \<Rightarrow> temporal"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    15
  Dmd        :: "('w::world) form \<Rightarrow> temporal"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    16
  leadsto    :: "['w::world form, 'v::world form] \<Rightarrow> temporal"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    17
  Stable     :: "stpred \<Rightarrow> temporal"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    18
  WF         :: "[action, 'a stfun] \<Rightarrow> temporal"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    19
  SF         :: "[action, 'a stfun] \<Rightarrow> temporal"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    20
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    21
  (* Quantification over (flexible) state variables *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    22
  EEx        :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal"       (binder "Eex " 10)
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    23
  AAll       :: "('a stfun \<Rightarrow> temporal) \<Rightarrow> temporal"       (binder "Aall " 10)
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    24
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    25
  (** concrete syntax **)
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    26
syntax
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    27
  "_Box"     :: "lift \<Rightarrow> lift"                        ("(\<box>_)" [40] 40)
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    28
  "_Dmd"     :: "lift \<Rightarrow> lift"                        ("(\<diamond>_)" [40] 40)
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    29
  "_leadsto" :: "[lift,lift] \<Rightarrow> lift"                 ("(_ \<leadsto> _)" [23,22] 22)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    30
  "_stable"  :: "lift \<Rightarrow> lift"                        ("(stable/ _)")
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    31
  "_WF"      :: "[lift,lift] \<Rightarrow> lift"                 ("(WF'(_')'_(_))" [0,60] 55)
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    32
  "_SF"      :: "[lift,lift] \<Rightarrow> lift"                 ("(SF'(_')'_(_))" [0,60] 55)
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    33
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    34
  "_EEx"     :: "[idts, lift] \<Rightarrow> lift"                ("(3\<exists>\<exists> _./ _)" [0,10] 10)
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    35
  "_AAll"    :: "[idts, lift] \<Rightarrow> lift"                ("(3\<forall>\<forall> _./ _)" [0,10] 10)
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    36
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    37
translations
35068
544867142ea4 modernized translations;
wenzelm
parents: 30528
diff changeset
    38
  "_Box"      ==   "CONST Box"
544867142ea4 modernized translations;
wenzelm
parents: 30528
diff changeset
    39
  "_Dmd"      ==   "CONST Dmd"
544867142ea4 modernized translations;
wenzelm
parents: 30528
diff changeset
    40
  "_leadsto"  ==   "CONST leadsto"
544867142ea4 modernized translations;
wenzelm
parents: 30528
diff changeset
    41
  "_stable"   ==   "CONST Stable"
544867142ea4 modernized translations;
wenzelm
parents: 30528
diff changeset
    42
  "_WF"       ==   "CONST WF"
544867142ea4 modernized translations;
wenzelm
parents: 30528
diff changeset
    43
  "_SF"       ==   "CONST SF"
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    44
  "_EEx v A"  ==   "Eex v. A"
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    45
  "_AAll v A" ==   "Aall v. A"
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    46
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    47
  "sigma \<Turnstile> \<box>F"         <= "_Box F sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    48
  "sigma \<Turnstile> \<diamond>F"         <= "_Dmd F sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    49
  "sigma \<Turnstile> F \<leadsto> G"      <= "_leadsto F G sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    50
  "sigma \<Turnstile> stable P"    <= "_stable P sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    51
  "sigma \<Turnstile> WF(A)_v"     <= "_WF A v sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    52
  "sigma \<Turnstile> SF(A)_v"     <= "_SF A v sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    53
  "sigma \<Turnstile> \<exists>\<exists>x. F"    <= "_EEx x F sigma"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    54
  "sigma \<Turnstile> \<forall>\<forall>x. F"    <= "_AAll x F sigma"
3808
8489375c6198 symbols syntax;
wenzelm
parents: 3807
diff changeset
    55
47968
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 45605
diff changeset
    56
axiomatization where
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    57
  (* Definitions of derived operators *)
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
    58
  dmd_def:      "\<And>F. TEMP \<diamond>F  ==  TEMP \<not>\<box>\<not>F"
47968
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 45605
diff changeset
    59
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 45605
diff changeset
    60
axiomatization where
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
    61
  boxInit:      "\<And>F. TEMP \<box>F  ==  TEMP \<box>Init F" and
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    62
  leadsto_def:  "\<And>F G. TEMP F \<leadsto> G  ==  TEMP \<box>(Init F \<longrightarrow> \<diamond>G)" and
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    63
  stable_def:   "\<And>P. TEMP stable P  ==  TEMP \<box>($P \<longrightarrow> P$)" and
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    64
  WF_def:       "TEMP WF(A)_v  ==  TEMP \<diamond>\<box> Enabled(<A>_v) \<longrightarrow> \<box>\<diamond><A>_v" and
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    65
  SF_def:       "TEMP SF(A)_v  ==  TEMP \<box>\<diamond> Enabled(<A>_v) \<longrightarrow> \<box>\<diamond><A>_v" and
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
    66
  aall_def:     "TEMP (\<forall>\<forall>x. F x)  ==  TEMP \<not> (\<exists>\<exists>x. \<not> F x)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    67
47968
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 45605
diff changeset
    68
axiomatization where
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3808
diff changeset
    69
(* Base axioms for raw TLA. *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    70
  normalT:    "\<And>F G. \<turnstile> \<box>(F \<longrightarrow> G) \<longrightarrow> (\<box>F \<longrightarrow> \<box>G)" and    (* polymorphic *)
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    71
  reflT:      "\<And>F. \<turnstile> \<box>F \<longrightarrow> F" and         (* F::temporal *)
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    72
  transT:     "\<And>F. \<turnstile> \<box>F \<longrightarrow> \<box>\<box>F" and     (* polymorphic *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    73
  linT:       "\<And>F G. \<turnstile> \<diamond>F \<and> \<diamond>G \<longrightarrow> (\<diamond>(F \<and> \<diamond>G)) \<or> (\<diamond>(G \<and> \<diamond>F))" and
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    74
  discT:      "\<And>F. \<turnstile> \<box>(F \<longrightarrow> \<diamond>(\<not>F \<and> \<diamond>F)) \<longrightarrow> (F \<longrightarrow> \<box>\<diamond>F)" and
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    75
  primeI:     "\<And>P. \<turnstile> \<box>P \<longrightarrow> Init P`" and
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    76
  primeE:     "\<And>P F. \<turnstile> \<box>(Init P \<longrightarrow> \<box>F) \<longrightarrow> Init P` \<longrightarrow> (F \<longrightarrow> \<box>F)" and
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    77
  indT:       "\<And>P F. \<turnstile> \<box>(Init P \<and> \<not>\<box>F \<longrightarrow> Init P` \<and> F) \<longrightarrow> Init P \<longrightarrow> \<box>F" and
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    78
  allT:       "\<And>F. \<turnstile> (\<forall>x. \<box>(F x)) = (\<box>(\<forall> x. F x))"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    79
47968
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 45605
diff changeset
    80
axiomatization where
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    81
  necT:       "\<And>F. \<turnstile> F \<Longrightarrow> \<turnstile> \<box>F"      (* polymorphic *)
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    82
47968
3119ad2b5ad3 eliminated old 'axioms';
wenzelm
parents: 45605
diff changeset
    83
axiomatization where
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    84
(* Flexible quantification: refinement mappings, history variables *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    85
  eexI:       "\<turnstile> F x \<longrightarrow> (\<exists>\<exists>x. F x)" and
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    86
  eexE:       "\<lbrakk> sigma \<Turnstile> (\<exists>\<exists>x. F x); basevars vs;
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    87
                 (\<And>x. \<lbrakk> basevars (x, vs); sigma \<Turnstile> F x \<rbrakk> \<Longrightarrow> (G sigma)::bool)
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    88
              \<rbrakk> \<Longrightarrow> G sigma" and
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
    89
  history:    "\<turnstile> \<exists>\<exists>h. Init(h = ha) \<and> \<box>(\<forall>x. $h = #x \<longrightarrow> h` = hb x)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 14565
diff changeset
    90
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    91
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    92
(* Specialize intensional introduction/elimination rules for temporal formulas *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    93
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    94
lemma tempI [intro!]: "(\<And>sigma. sigma \<Turnstile> (F::temporal)) \<Longrightarrow> \<turnstile> F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    95
  apply (rule intI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    96
  apply (erule meta_spec)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    97
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    98
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
    99
lemma tempD [dest]: "\<turnstile> (F::temporal) \<Longrightarrow> sigma \<Turnstile> F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   100
  by (erule intD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   101
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   102
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   103
(* ======== Functions to "unlift" temporal theorems ====== *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   104
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   105
ML \<open>
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   106
(* The following functions are specialized versions of the corresponding
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   107
   functions defined in theory Intensional in that they introduce a
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   108
   "world" parameter of type "behavior".
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   109
*)
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   110
fun temp_unlift ctxt th =
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   111
  (rewrite_rule ctxt @{thms action_rews} (th RS @{thm tempD}))
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   112
    handle THM _ => action_unlift ctxt th;
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   113
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   114
(* Turn  \<turnstile> F = G  into meta-level rewrite rule  F == G *)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   115
val temp_rewrite = int_rewrite
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   116
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   117
fun temp_use ctxt th =
59582
0fbed69ff081 tuned signature -- prefer qualified names;
wenzelm
parents: 59498
diff changeset
   118
  case Thm.concl_of th of
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   119
    Const _ $ (Const (@{const_name Intensional.Valid}, _) $ _) =>
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   120
            ((flatten (temp_unlift ctxt th)) handle THM _ => th)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   121
  | _ => th;
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   122
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   123
fun try_rewrite ctxt th = temp_rewrite ctxt th handle THM _ => temp_use ctxt th;
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   124
\<close>
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   125
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   126
attribute_setup temp_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: 60754
diff changeset
   127
  \<open>Scan.succeed (Thm.rule_attribute [] (temp_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: 51717
diff changeset
   128
attribute_setup temp_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: 60754
diff changeset
   129
  \<open>Scan.succeed (Thm.rule_attribute [] (temp_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: 51717
diff changeset
   130
attribute_setup temp_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: 60754
diff changeset
   131
  \<open>Scan.succeed (Thm.rule_attribute [] (temp_use o Context.proof_of))\<close>
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 51717
diff changeset
   132
attribute_setup try_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: 60754
diff changeset
   133
  \<open>Scan.succeed (Thm.rule_attribute [] (try_rewrite o Context.proof_of))\<close>
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 27239
diff changeset
   134
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   135
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   136
(* ------------------------------------------------------------------------- *)
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   137
(***           "Simple temporal logic": only \<box> and \<diamond>                     ***)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   138
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   139
section "Simple temporal logic"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   140
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   141
(* \<box>\<not>F == \<box>\<not>Init F *)
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   142
lemmas boxNotInit = boxInit [of "LIFT \<not>F", unfolded Init_simps] for F
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   143
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   144
lemma dmdInit: "TEMP \<diamond>F == TEMP \<diamond> Init F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   145
  apply (unfold dmd_def)
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   146
  apply (unfold boxInit [of "LIFT \<not>F"])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   147
  apply (simp (no_asm) add: Init_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   148
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   149
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   150
lemmas dmdNotInit = dmdInit [of "LIFT \<not>F", unfolded Init_simps] for F
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   151
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   152
(* boxInit and dmdInit cannot be used as rewrites, because they loop.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   153
   Non-looping instances for state predicates and actions are occasionally useful.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   154
*)
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   155
lemmas boxInit_stp = boxInit [where 'a = state]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   156
lemmas boxInit_act = boxInit [where 'a = "state * state"]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   157
lemmas dmdInit_stp = dmdInit [where 'a = state]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   158
lemmas dmdInit_act = dmdInit [where 'a = "state * state"]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   159
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   160
(* The symmetric equations can be used to get rid of Init *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   161
lemmas boxInitD = boxInit [symmetric]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   162
lemmas dmdInitD = dmdInit [symmetric]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   163
lemmas boxNotInitD = boxNotInit [symmetric]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   164
lemmas dmdNotInitD = dmdNotInit [symmetric]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   165
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   166
lemmas Init_simps = Init_simps boxInitD dmdInitD boxNotInitD dmdNotInitD
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   167
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   168
(* ------------------------ STL2 ------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   169
lemmas STL2 = reflT
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   170
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   171
(* The "polymorphic" (generic) variant *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   172
lemma STL2_gen: "\<turnstile> \<box>F \<longrightarrow> Init F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   173
  apply (unfold boxInit [of F])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   174
  apply (rule STL2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   175
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   176
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   177
(* see also STL2_pr below: "\<turnstile> \<box>P \<longrightarrow> Init P & Init (P`)" *)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   178
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   179
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   180
(* Dual versions for \<diamond> *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   181
lemma InitDmd: "\<turnstile> F \<longrightarrow> \<diamond> F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   182
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   183
  apply (auto dest!: STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   184
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   185
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   186
lemma InitDmd_gen: "\<turnstile> Init F \<longrightarrow> \<diamond>F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   187
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   188
  apply (drule InitDmd [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   189
  apply (simp add: dmdInitD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   190
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   191
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   192
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   193
(* ------------------------ STL3 ------------------------------------------- *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   194
lemma STL3: "\<turnstile> (\<box>\<box>F) = (\<box>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   195
  by (auto elim: transT [temp_use] STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   196
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   197
(* corresponding elimination rule introduces double boxes:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   198
   \<lbrakk> (sigma \<Turnstile> \<box>F); (sigma \<Turnstile> \<box>\<box>F) \<Longrightarrow> PROP W \<rbrakk> \<Longrightarrow> PROP W
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   199
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   200
lemmas dup_boxE = STL3 [temp_unlift, THEN iffD2, elim_format]
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   201
lemmas dup_boxD = STL3 [temp_unlift, THEN iffD1]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   202
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   203
(* dual versions for \<diamond> *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   204
lemma DmdDmd: "\<turnstile> (\<diamond>\<diamond>F) = (\<diamond>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   205
  by (auto simp add: dmd_def [try_rewrite] STL3 [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   206
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   207
lemmas dup_dmdE = DmdDmd [temp_unlift, THEN iffD2, elim_format]
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   208
lemmas dup_dmdD = DmdDmd [temp_unlift, THEN iffD1]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   209
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   210
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   211
(* ------------------------ STL4 ------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   212
lemma STL4:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   213
  assumes "\<turnstile> F \<longrightarrow> G"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   214
  shows "\<turnstile> \<box>F \<longrightarrow> \<box>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   215
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   216
  apply (rule normalT [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   217
   apply (rule assms [THEN necT, temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   218
  apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   219
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   220
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   221
(* Unlifted version as an elimination rule *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   222
lemma STL4E: "\<lbrakk> sigma \<Turnstile> \<box>F; \<turnstile> F \<longrightarrow> G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   223
  by (erule (1) STL4 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   224
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   225
lemma STL4_gen: "\<turnstile> Init F \<longrightarrow> Init G \<Longrightarrow> \<turnstile> \<box>F \<longrightarrow> \<box>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   226
  apply (drule STL4)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   227
  apply (simp add: boxInitD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   228
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   229
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   230
lemma STL4E_gen: "\<lbrakk> sigma \<Turnstile> \<box>F; \<turnstile> Init F \<longrightarrow> Init G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   231
  by (erule (1) STL4_gen [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   232
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   233
(* see also STL4Edup below, which allows an auxiliary boxed formula:
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   234
       \<box>A /\ F => G
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   235
     -----------------
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   236
     \<box>A /\ \<box>F => \<box>G
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   237
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   238
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   239
(* The dual versions for \<diamond> *)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   240
lemma DmdImpl:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   241
  assumes prem: "\<turnstile> F \<longrightarrow> G"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   242
  shows "\<turnstile> \<diamond>F \<longrightarrow> \<diamond>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   243
  apply (unfold dmd_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   244
  apply (fastforce intro!: prem [temp_use] elim!: STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   245
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   246
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   247
lemma DmdImplE: "\<lbrakk> sigma \<Turnstile> \<diamond>F; \<turnstile> F \<longrightarrow> G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<diamond>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   248
  by (erule (1) DmdImpl [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   249
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   250
(* ------------------------ STL5 ------------------------------------------- *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   251
lemma STL5: "\<turnstile> (\<box>F \<and> \<box>G) = (\<box>(F \<and> G))"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   252
  apply auto
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   253
  apply (subgoal_tac "sigma \<Turnstile> \<box> (G \<longrightarrow> (F \<and> G))")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   254
     apply (erule normalT [temp_use])
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   255
     apply (fastforce elim!: STL4E [temp_use])+
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   256
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   257
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   258
(* rewrite rule to split conjunctions under boxes *)
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   259
lemmas split_box_conj = STL5 [temp_unlift, symmetric]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   260
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   261
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   262
(* the corresponding elimination rule allows to combine boxes in the hypotheses
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   263
   (NB: F and G must have the same type, i.e., both actions or temporals.)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   264
   Use "addSE2" etc. if you want to add this to a claset, otherwise it will loop!
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   265
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   266
lemma box_conjE:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   267
  assumes "sigma \<Turnstile> \<box>F"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   268
     and "sigma \<Turnstile> \<box>G"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   269
  and "sigma \<Turnstile> \<box>(F\<and>G) \<Longrightarrow> PROP R"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   270
  shows "PROP R"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   271
  by (rule assms STL5 [temp_unlift, THEN iffD1] conjI)+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   272
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   273
(* Instances of box_conjE for state predicates, actions, and temporals
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   274
   in case the general rule is "too polymorphic".
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   275
*)
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   276
lemmas box_conjE_temp = box_conjE [where 'a = behavior]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   277
lemmas box_conjE_stp = box_conjE [where 'a = state]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   278
lemmas box_conjE_act = box_conjE [where 'a = "state * state"]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   279
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   280
(* Define a tactic that tries to merge all boxes in an antecedent. The definition is
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   281
   a bit kludgy in order to simulate "double elim-resolution".
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   282
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   283
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   284
lemma box_thin: "\<lbrakk> sigma \<Turnstile> \<box>F; PROP W \<rbrakk> \<Longrightarrow> PROP W" .
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   285
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   286
ML \<open>
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   287
fun merge_box_tac ctxt i =
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   288
   REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE} i, assume_tac ctxt i,
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   289
    eresolve_tac ctxt @{thms box_thin} i])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   290
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26305
diff changeset
   291
fun merge_temp_box_tac ctxt i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   292
  REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_temp} i, assume_tac ctxt i,
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   293
    Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "behavior")] [] @{thm box_thin} i])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   294
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26305
diff changeset
   295
fun merge_stp_box_tac ctxt i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   296
  REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_stp} i, assume_tac ctxt i,
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   297
    Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "state")] [] @{thm box_thin} i])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   298
27208
5fe899199f85 proper context for tactics derived from res_inst_tac;
wenzelm
parents: 26305
diff changeset
   299
fun merge_act_box_tac ctxt i =
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   300
  REPEAT_DETERM (EVERY [eresolve_tac ctxt @{thms box_conjE_act} i, assume_tac ctxt i,
59780
23b67731f4f0 support 'for' fixes in rule_tac etc.;
wenzelm
parents: 59763
diff changeset
   301
    Rule_Insts.eres_inst_tac ctxt [((("'a", 0), Position.none), "state * state")] [] @{thm box_thin} i])
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   302
\<close>
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   303
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   304
method_setup merge_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_box_tac)\<close>
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   305
method_setup merge_temp_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_temp_box_tac)\<close>
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   306
method_setup merge_stp_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_stp_box_tac)\<close>
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   307
method_setup merge_act_box = \<open>Scan.succeed (SIMPLE_METHOD' o merge_act_box_tac)\<close>
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   308
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   309
(* rewrite rule to push universal quantification through box:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   310
      (sigma \<Turnstile> \<box>(\<forall>x. F x)) = (\<forall>x. (sigma \<Turnstile> \<box>F x))
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   311
*)
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   312
lemmas all_box = allT [temp_unlift, symmetric]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   313
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   314
lemma DmdOr: "\<turnstile> (\<diamond>(F \<or> G)) = (\<diamond>F \<or> \<diamond>G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   315
  apply (auto simp add: dmd_def split_box_conj [try_rewrite])
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   316
  apply (erule contrapos_np, merge_box, fastforce elim!: STL4E [temp_use])+
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   317
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   318
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   319
lemma exT: "\<turnstile> (\<exists>x. \<diamond>(F x)) = (\<diamond>(\<exists>x. F x))"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   320
  by (auto simp: dmd_def Not_Rex [try_rewrite] all_box [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   321
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   322
lemmas ex_dmd = exT [temp_unlift, symmetric]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   323
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   324
lemma STL4Edup: "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<box>A; sigma \<Turnstile> \<box>F; \<turnstile> F \<and> \<box>A \<longrightarrow> G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   325
  apply (erule dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   326
  apply merge_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   327
  apply (erule STL4E)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   328
  apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   329
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   330
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   331
lemma DmdImpl2:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   332
    "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<diamond>F; sigma \<Turnstile> \<box>(F \<longrightarrow> G) \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<diamond>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   333
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   334
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   335
  apply (erule notE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   336
  apply merge_box
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   337
  apply (fastforce elim!: STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   338
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   339
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   340
lemma InfImpl:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   341
  assumes 1: "sigma \<Turnstile> \<box>\<diamond>F"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   342
    and 2: "sigma \<Turnstile> \<box>G"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   343
    and 3: "\<turnstile> F \<and> G \<longrightarrow> H"
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   344
  shows "sigma \<Turnstile> \<box>\<diamond>H"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   345
  apply (insert 1 2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   346
  apply (erule_tac F = G in dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   347
  apply merge_box
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   348
  apply (fastforce elim!: STL4E [temp_use] DmdImpl2 [temp_use] intro!: 3 [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   349
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   350
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   351
(* ------------------------ STL6 ------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   352
(* Used in the proof of STL6, but useful in itself. *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   353
lemma BoxDmd: "\<turnstile> \<box>F \<and> \<diamond>G \<longrightarrow> \<diamond>(\<box>F \<and> G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   354
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   355
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   356
  apply (erule dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   357
  apply merge_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   358
  apply (erule contrapos_np)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   359
  apply (fastforce elim!: STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   360
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   361
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   362
(* weaker than BoxDmd, but more polymorphic (and often just right) *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   363
lemma BoxDmd_simple: "\<turnstile> \<box>F \<and> \<diamond>G \<longrightarrow> \<diamond>(F \<and> G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   364
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   365
  apply clarsimp
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   366
  apply merge_box
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   367
  apply (fastforce elim!: notE STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   368
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   369
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   370
lemma BoxDmd2_simple: "\<turnstile> \<box>F \<and> \<diamond>G \<longrightarrow> \<diamond>(G \<and> F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   371
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   372
  apply clarsimp
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   373
  apply merge_box
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   374
  apply (fastforce elim!: notE STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   375
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   376
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   377
lemma DmdImpldup:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   378
  assumes 1: "sigma \<Turnstile> \<box>A"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   379
    and 2: "sigma \<Turnstile> \<diamond>F"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   380
    and 3: "\<turnstile> \<box>A \<and> F \<longrightarrow> G"
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   381
  shows "sigma \<Turnstile> \<diamond>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   382
  apply (rule 2 [THEN 1 [THEN BoxDmd [temp_use]], THEN DmdImplE])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   383
  apply (rule 3)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   384
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   385
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   386
lemma STL6: "\<turnstile> \<diamond>\<box>F \<and> \<diamond>\<box>G \<longrightarrow> \<diamond>\<box>(F \<and> G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   387
  apply (auto simp: STL5 [temp_rewrite, symmetric])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   388
  apply (drule linT [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   389
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   390
  apply (erule thin_rl)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   391
  apply (rule DmdDmd [temp_unlift, THEN iffD1])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   392
  apply (erule disjE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   393
   apply (erule DmdImplE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   394
   apply (rule BoxDmd)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   395
  apply (erule DmdImplE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   396
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   397
  apply (drule BoxDmd [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   398
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   399
  apply (erule thin_rl)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   400
  apply (fastforce elim!: DmdImplE [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   401
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   402
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   403
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   404
(* ------------------------ True / False ----------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   405
section "Simplification of constants"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   406
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   407
lemma BoxConst: "\<turnstile> (\<box>#P) = #P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   408
  apply (rule tempI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   409
  apply (cases P)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   410
   apply (auto intro!: necT [temp_use] dest: STL2_gen [temp_use] simp: Init_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   411
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   412
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   413
lemma DmdConst: "\<turnstile> (\<diamond>#P) = #P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   414
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   415
  apply (cases P)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   416
  apply (simp_all add: BoxConst [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   417
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   418
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   419
lemmas temp_simps [temp_rewrite, simp] = BoxConst DmdConst
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   420
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   421
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   422
(* ------------------------ Further rewrites ----------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   423
section "Further rewrites"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   424
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   425
lemma NotBox: "\<turnstile> (\<not>\<box>F) = (\<diamond>\<not>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   426
  by (simp add: dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   427
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   428
lemma NotDmd: "\<turnstile> (\<not>\<diamond>F) = (\<box>\<not>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   429
  by (simp add: dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   430
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   431
(* These are not declared by default, because they could be harmful,
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   432
   e.g. \<box>F & \<not>\<box>F becomes \<box>F & \<diamond>\<not>F !! *)
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   433
lemmas more_temp_simps1 =
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   434
  STL3 [temp_rewrite] DmdDmd [temp_rewrite] NotBox [temp_rewrite] NotDmd [temp_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   435
  NotBox [temp_unlift, THEN eq_reflection]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   436
  NotDmd [temp_unlift, THEN eq_reflection]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   437
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   438
lemma BoxDmdBox: "\<turnstile> (\<box>\<diamond>\<box>F) = (\<diamond>\<box>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   439
  apply (auto dest!: STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   440
  apply (rule ccontr)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   441
  apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box>\<box>F \<and> \<diamond>\<box>\<not>\<box>F")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   442
   apply (erule thin_rl)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   443
   apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   444
    apply (drule STL6 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   445
     apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   446
    apply simp
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   447
   apply (simp_all add: more_temp_simps1)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   448
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   449
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   450
lemma DmdBoxDmd: "\<turnstile> (\<diamond>\<box>\<diamond>F) = (\<box>\<diamond>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   451
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   452
  apply (auto simp: BoxDmdBox [unfolded dmd_def, try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   453
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   454
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   455
lemmas more_temp_simps2 = more_temp_simps1 BoxDmdBox [temp_rewrite] DmdBoxDmd [temp_rewrite]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   456
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   457
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   458
(* ------------------------ Miscellaneous ----------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   459
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   460
lemma BoxOr: "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<box>F \<or> \<box>G \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>(F \<or> G)"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   461
  by (fastforce elim!: STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   462
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   463
(* "persistently implies infinitely often" *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   464
lemma DBImplBD: "\<turnstile> \<diamond>\<box>F \<longrightarrow> \<box>\<diamond>F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   465
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   466
  apply (rule ccontr)
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   467
  apply (simp add: more_temp_simps2)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   468
  apply (drule STL6 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   469
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   470
  apply simp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   471
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   472
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   473
lemma BoxDmdDmdBox: "\<turnstile> \<box>\<diamond>F \<and> \<diamond>\<box>G \<longrightarrow> \<box>\<diamond>(F \<and> G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   474
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   475
  apply (rule ccontr)
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   476
  apply (unfold more_temp_simps2)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   477
  apply (drule STL6 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   478
   apply assumption
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   479
  apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box>\<not>F")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   480
   apply (force simp: dmd_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   481
  apply (fastforce elim: DmdImplE [temp_use] STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   482
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   483
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   484
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   485
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   486
(***          TLA-specific theorems: primed formulas                       ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   487
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   488
section "priming"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   489
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   490
(* ------------------------ TLA2 ------------------------------------------- *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   491
lemma STL2_pr: "\<turnstile> \<box>P \<longrightarrow> Init P \<and> Init P`"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   492
  by (fastforce intro!: STL2_gen [temp_use] primeI [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   493
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   494
(* Auxiliary lemma allows priming of boxed actions *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   495
lemma BoxPrime: "\<turnstile> \<box>P \<longrightarrow> \<box>($P \<and> P$)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   496
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   497
  apply (erule dup_boxE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   498
  apply (unfold boxInit_act)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   499
  apply (erule STL4E)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   500
  apply (auto simp: Init_simps dest!: STL2_pr [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   501
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   502
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   503
lemma TLA2:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   504
  assumes "\<turnstile> $P \<and> P$ \<longrightarrow> A"
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   505
  shows "\<turnstile> \<box>P \<longrightarrow> \<box>A"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   506
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   507
  apply (drule BoxPrime [temp_use])
41529
ba60efa2fd08 eliminated global prems;
wenzelm
parents: 35108
diff changeset
   508
  apply (auto simp: Init_stp_act_rev [try_rewrite] intro!: assms [temp_use]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   509
    elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   510
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   511
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   512
lemma TLA2E: "\<lbrakk> sigma \<Turnstile> \<box>P; \<turnstile> $P \<and> P$ \<longrightarrow> A \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>A"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   513
  by (erule (1) TLA2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   514
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   515
lemma DmdPrime: "\<turnstile> (\<diamond>P`) \<longrightarrow> (\<diamond>P)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   516
  apply (unfold dmd_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   517
  apply (fastforce elim!: TLA2E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   518
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   519
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   520
lemmas PrimeDmd = InitDmd_gen [temp_use, THEN DmdPrime [temp_use]]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   521
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   522
(* ------------------------ INV1, stable --------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   523
section "stable, invariant"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   524
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   525
lemma ind_rule:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   526
   "\<lbrakk> sigma \<Turnstile> \<box>H; sigma \<Turnstile> Init P; \<turnstile> H \<longrightarrow> (Init P \<and> \<not>\<box>F \<longrightarrow> Init(P`) \<and> F) \<rbrakk>
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   527
    \<Longrightarrow> sigma \<Turnstile> \<box>F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   528
  apply (rule indT [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   529
   apply (erule (2) STL4E)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   530
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   531
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   532
lemma box_stp_act: "\<turnstile> (\<box>$P) = (\<box>P)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   533
  by (simp add: boxInit_act Init_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   534
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   535
lemmas box_stp_actI = box_stp_act [temp_use, THEN iffD2]
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   536
lemmas box_stp_actD = box_stp_act [temp_use, THEN iffD1]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   537
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   538
lemmas more_temp_simps3 = box_stp_act [temp_rewrite] more_temp_simps2
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   539
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   540
lemma INV1:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   541
  "\<turnstile> (Init P) \<longrightarrow> (stable P) \<longrightarrow> \<box>P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   542
  apply (unfold stable_def boxInit_stp boxInit_act)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   543
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   544
  apply (erule ind_rule)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   545
   apply (auto simp: Init_simps elim: ind_rule)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   546
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   547
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   548
lemma StableT:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   549
    "\<And>P. \<turnstile> $P \<and> A \<longrightarrow> P` \<Longrightarrow> \<turnstile> \<box>A \<longrightarrow> stable P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   550
  apply (unfold stable_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   551
  apply (fastforce elim!: STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   552
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   553
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   554
lemma Stable: "\<lbrakk> sigma \<Turnstile> \<box>A; \<turnstile> $P \<and> A \<longrightarrow> P` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> stable P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   555
  by (erule (1) StableT [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   556
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   557
(* Generalization of INV1 *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   558
lemma StableBox: "\<turnstile> (stable P) \<longrightarrow> \<box>(Init P \<longrightarrow> \<box>P)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   559
  apply (unfold stable_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   560
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   561
  apply (erule dup_boxE)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   562
  apply (force simp: stable_def elim: STL4E [temp_use] INV1 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   563
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   564
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   565
lemma DmdStable: "\<turnstile> (stable P) \<and> \<diamond>P \<longrightarrow> \<diamond>\<box>P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   566
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   567
  apply (rule DmdImpl2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   568
   prefer 2
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   569
   apply (erule StableBox [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   570
  apply (simp add: dmdInitD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   571
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   572
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   573
(* ---------------- (Semi-)automatic invariant tactics ---------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   574
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   575
ML \<open>
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   576
(* inv_tac reduces goals of the form ... \<Longrightarrow> sigma \<Turnstile> \<box>P *)
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   577
fun inv_tac ctxt =
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   578
  SELECT_GOAL
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   579
    (EVERY
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   580
     [auto_tac ctxt,
60754
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   581
      TRY (merge_box_tac ctxt 1),
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   582
      resolve_tac ctxt [temp_use ctxt @{thm INV1}] 1, (* fail if the goal is not a box *)
02924903a6fd prefer tactics with explicit context;
wenzelm
parents: 60592
diff changeset
   583
      TRYALL (eresolve_tac ctxt @{thms Stable})]);
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   584
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   585
(* auto_inv_tac applies inv_tac and then tries to attack the subgoals
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   586
   in simple cases it may be able to handle goals like \<turnstile> MyProg \<longrightarrow> \<box>Inv.
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   587
   In these simple cases the simplifier seems to be more useful than the
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   588
   auto-tactic, which applies too much propositional logic and simplifies
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   589
   too late.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   590
*)
42803
7ed59879b1b6 proper runtime context for auto_inv_tac;
wenzelm
parents: 42795
diff changeset
   591
fun auto_inv_tac ctxt =
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   592
  SELECT_GOAL
42803
7ed59879b1b6 proper runtime context for auto_inv_tac;
wenzelm
parents: 42795
diff changeset
   593
    (inv_tac ctxt 1 THEN
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   594
      (TRYALL (action_simp_tac
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 51668
diff changeset
   595
        (ctxt addsimps [@{thm Init_stp}, @{thm Init_act}]) [] [@{thm squareE}])));
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   596
\<close>
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   597
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   598
method_setup invariant = \<open>
42793
88bee9f6eec7 proper Proof.context for classical tactics;
wenzelm
parents: 42788
diff changeset
   599
  Method.sections Clasimp.clasimp_modifiers >> (K (SIMPLE_METHOD' o inv_tac))
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   600
\<close>
42769
053b4b487085 proper method_setup;
wenzelm
parents: 41529
diff changeset
   601
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   602
method_setup auto_invariant = \<open>
42803
7ed59879b1b6 proper runtime context for auto_inv_tac;
wenzelm
parents: 42795
diff changeset
   603
  Method.sections Clasimp.clasimp_modifiers >> (K (SIMPLE_METHOD' o auto_inv_tac))
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   604
\<close>
42769
053b4b487085 proper method_setup;
wenzelm
parents: 41529
diff changeset
   605
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   606
lemma unless: "\<turnstile> \<box>($P \<longrightarrow> P` \<or> Q`) \<longrightarrow> (stable P) \<or> \<diamond>Q"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   607
  apply (unfold dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   608
  apply (clarsimp dest!: BoxPrime [temp_use])
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   609
  apply merge_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   610
  apply (erule contrapos_np)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   611
  apply (fastforce elim!: Stable [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   612
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   613
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   614
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   615
(* --------------------- Recursive expansions --------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   616
section "recursive expansions"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   617
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   618
(* Recursive expansions of \<box> and \<diamond> for state predicates *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   619
lemma BoxRec: "\<turnstile> (\<box>P) = (Init P \<and> \<box>P`)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   620
  apply (auto intro!: STL2_gen [temp_use])
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   621
   apply (fastforce elim!: TLA2E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   622
  apply (auto simp: stable_def elim!: INV1 [temp_use] STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   623
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   624
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   625
lemma DmdRec: "\<turnstile> (\<diamond>P) = (Init P \<or> \<diamond>P`)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   626
  apply (unfold dmd_def BoxRec [temp_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   627
  apply (auto simp: Init_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   628
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   629
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   630
lemma DmdRec2: "\<And>sigma. \<lbrakk> sigma \<Turnstile> \<diamond>P; sigma \<Turnstile> \<box>\<not>P` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> Init P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   631
  apply (force simp: DmdRec [temp_rewrite] dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   632
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   633
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   634
lemma InfinitePrime: "\<turnstile> (\<box>\<diamond>P) = (\<box>\<diamond>P`)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   635
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   636
   apply (rule classical)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   637
   apply (rule DBImplBD [temp_use])
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   638
   apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box>P")
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   639
    apply (fastforce elim!: DmdImplE [temp_use] TLA2E [temp_use])
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   640
   apply (subgoal_tac "sigma \<Turnstile> \<diamond>\<box> (\<diamond>P \<and> \<box>\<not>P`)")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   641
    apply (force simp: boxInit_stp [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   642
      elim!: DmdImplE [temp_use] STL4E [temp_use] DmdRec2 [temp_use])
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   643
   apply (force intro!: STL6 [temp_use] simp: more_temp_simps3)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   644
  apply (fastforce intro: DmdPrime [temp_use] elim!: STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   645
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   646
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   647
lemma InfiniteEnsures:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   648
  "\<lbrakk> sigma \<Turnstile> \<box>N; sigma \<Turnstile> \<box>\<diamond>A; \<turnstile> A \<and> N \<longrightarrow> P` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>\<diamond>P"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   649
  apply (unfold InfinitePrime [temp_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   650
  apply (rule InfImpl)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   651
    apply assumption+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   652
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   653
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   654
(* ------------------------ fairness ------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   655
section "fairness"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   656
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   657
(* alternative definitions of fairness *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   658
lemma WF_alt: "\<turnstile> WF(A)_v = (\<box>\<diamond>\<not>Enabled(<A>_v) \<or> \<box>\<diamond><A>_v)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   659
  apply (unfold WF_def dmd_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   660
  apply fastforce
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   661
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   662
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   663
lemma SF_alt: "\<turnstile> SF(A)_v = (\<diamond>\<box>\<not>Enabled(<A>_v) \<or> \<box>\<diamond><A>_v)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   664
  apply (unfold SF_def dmd_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   665
  apply fastforce
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   666
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   667
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   668
(* theorems to "box" fairness conditions *)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   669
lemma BoxWFI: "\<turnstile> WF(A)_v \<longrightarrow> \<box>WF(A)_v"
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   670
  by (auto simp: WF_alt [try_rewrite] more_temp_simps3 intro!: BoxOr [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   671
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   672
lemma WF_Box: "\<turnstile> (\<box>WF(A)_v) = WF(A)_v"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   673
  by (fastforce intro!: BoxWFI [temp_use] dest!: STL2 [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   674
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   675
lemma BoxSFI: "\<turnstile> SF(A)_v \<longrightarrow> \<box>SF(A)_v"
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   676
  by (auto simp: SF_alt [try_rewrite] more_temp_simps3 intro!: BoxOr [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   677
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   678
lemma SF_Box: "\<turnstile> (\<box>SF(A)_v) = SF(A)_v"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   679
  by (fastforce intro!: BoxSFI [temp_use] dest!: STL2 [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   680
26305
651371f29e00 avoid rebinding of existing facts;
wenzelm
parents: 21624
diff changeset
   681
lemmas more_temp_simps = more_temp_simps3 WF_Box [temp_rewrite] SF_Box [temp_rewrite]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   682
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   683
lemma SFImplWF: "\<turnstile> SF(A)_v \<longrightarrow> WF(A)_v"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   684
  apply (unfold SF_def WF_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   685
  apply (fastforce dest!: DBImplBD [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   686
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   687
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   688
(* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *)
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   689
ML \<open>
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   690
fun box_fair_tac ctxt =
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   691
  SELECT_GOAL (REPEAT (dresolve_tac ctxt [@{thm BoxWFI}, @{thm BoxSFI}] 1))
60592
c9bd1d902f04 isabelle update_cartouches;
wenzelm
parents: 60591
diff changeset
   692
\<close>
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   693
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   694
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   695
(* ------------------------------ leads-to ------------------------------ *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   696
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   697
section "\<leadsto>"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   698
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   699
lemma leadsto_init: "\<turnstile> (Init F) \<and> (F \<leadsto> G) \<longrightarrow> \<diamond>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   700
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   701
  apply (auto dest!: STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   702
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   703
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   704
(* \<turnstile> F & (F \<leadsto> G) \<longrightarrow> \<diamond>G *)
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   705
lemmas leadsto_init_temp = leadsto_init [where 'a = behavior, unfolded Init_simps]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   706
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   707
lemma streett_leadsto: "\<turnstile> (\<box>\<diamond>Init F \<longrightarrow> \<box>\<diamond>G) = (\<diamond>(F \<leadsto> G))"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   708
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   709
  apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   710
    apply (simp add: more_temp_simps)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   711
    apply (fastforce elim!: DmdImplE [temp_use] STL4E [temp_use])
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   712
   apply (fastforce intro!: InitDmd [temp_use] elim!: STL4E [temp_use])
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   713
  apply (subgoal_tac "sigma \<Turnstile> \<box>\<diamond>\<diamond>G")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   714
   apply (simp add: more_temp_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   715
  apply (drule BoxDmdDmdBox [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   716
   apply assumption
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   717
  apply (fastforce elim!: DmdImplE [temp_use] STL4E [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   718
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   719
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   720
lemma leadsto_infinite: "\<turnstile> \<box>\<diamond>F \<and> (F \<leadsto> G) \<longrightarrow> \<box>\<diamond>G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   721
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   722
  apply (erule InitDmd [temp_use, THEN streett_leadsto [temp_unlift, THEN iffD2, THEN mp]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   723
  apply (simp add: dmdInitD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   724
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   725
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   726
(* In particular, strong fairness is a Streett condition. The following
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   727
   rules are sometimes easier to use than WF2 or SF2 below.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   728
*)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   729
lemma leadsto_SF: "\<turnstile> (Enabled(<A>_v) \<leadsto> <A>_v) \<longrightarrow> SF(A)_v"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   730
  apply (unfold SF_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   731
  apply (clarsimp elim!: leadsto_infinite [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   732
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   733
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   734
lemma leadsto_WF: "\<turnstile> (Enabled(<A>_v) \<leadsto> <A>_v) \<longrightarrow> WF(A)_v"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   735
  by (clarsimp intro!: SFImplWF [temp_use] leadsto_SF [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   736
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   737
(* introduce an invariant into the proof of a leadsto assertion.
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   738
   \<box>I \<longrightarrow> ((P \<leadsto> Q)  =  (P /\ I \<leadsto> Q))
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   739
*)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   740
lemma INV_leadsto: "\<turnstile> \<box>I \<and> (P \<and> I \<leadsto> Q) \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   741
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   742
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   743
  apply (erule STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   744
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   745
  apply (auto simp: Init_simps dest!: STL2_gen [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   746
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   747
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   748
lemma leadsto_classical: "\<turnstile> (Init F \<and> \<box>\<not>G \<leadsto> G) \<longrightarrow> (F \<leadsto> G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   749
  apply (unfold leadsto_def dmd_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   750
  apply (force simp: Init_simps elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   751
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   752
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   753
lemma leadsto_false: "\<turnstile> (F \<leadsto> #False) = (\<box>\<not>F)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   754
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   755
  apply (simp add: boxNotInitD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   756
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   757
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   758
lemma leadsto_exists: "\<turnstile> ((\<exists>x. F x) \<leadsto> G) = (\<forall>x. (F x \<leadsto> G))"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   759
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   760
  apply (auto simp: allT [try_rewrite] Init_simps elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   761
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   762
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   763
(* basic leadsto properties, cf. Unity *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   764
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   765
lemma ImplLeadsto_gen: "\<turnstile> \<box>(Init F \<longrightarrow> Init G) \<longrightarrow> (F \<leadsto> G)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   766
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   767
  apply (auto intro!: InitDmd_gen [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   768
    elim!: STL4E_gen [temp_use] simp: Init_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   769
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   770
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   771
lemmas ImplLeadsto =
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
   772
  ImplLeadsto_gen [where 'a = behavior and 'b = behavior, unfolded Init_simps]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   773
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   774
lemma ImplLeadsto_simple: "\<And>F G. \<turnstile> F \<longrightarrow> G \<Longrightarrow> \<turnstile> F \<leadsto> G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   775
  by (auto simp: Init_def intro!: ImplLeadsto_gen [temp_use] necT [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   776
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   777
lemma EnsuresLeadsto:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   778
  assumes "\<turnstile> A \<and> $P \<longrightarrow> Q`"
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   779
  shows "\<turnstile> \<box>A \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   780
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   781
  apply (clarsimp elim!: INV_leadsto [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   782
  apply (erule STL4E_gen)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   783
  apply (auto simp: Init_defs intro!: PrimeDmd [temp_use] assms [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   784
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   785
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   786
lemma EnsuresLeadsto2: "\<turnstile> \<box>($P \<longrightarrow> Q`) \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   787
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   788
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   789
  apply (erule STL4E_gen)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   790
  apply (auto simp: Init_simps intro!: PrimeDmd [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   791
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   792
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   793
lemma ensures:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   794
  assumes 1: "\<turnstile> $P \<and> N \<longrightarrow> P` \<or> Q`"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   795
    and 2: "\<turnstile> ($P \<and> N) \<and> A \<longrightarrow> Q`"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   796
  shows "\<turnstile> \<box>N \<and> \<box>(\<box>P \<longrightarrow> \<diamond>A) \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   797
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   798
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   799
  apply (erule STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   800
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   801
  apply clarsimp
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   802
  apply (subgoal_tac "sigmaa \<Turnstile> \<box>($P \<longrightarrow> P` \<or> Q`) ")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   803
   apply (drule unless [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   804
   apply (clarsimp dest!: INV1 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   805
  apply (rule 2 [THEN DmdImpl, temp_use, THEN DmdPrime [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   806
   apply (force intro!: BoxDmd_simple [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   807
     simp: split_box_conj [try_rewrite] box_stp_act [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   808
  apply (force elim: STL4E [temp_use] dest: 1 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   809
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   810
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   811
lemma ensures_simple:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   812
  "\<lbrakk> \<turnstile> $P \<and> N \<longrightarrow> P` \<or> Q`;
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   813
      \<turnstile> ($P \<and> N) \<and> A \<longrightarrow> Q`
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   814
   \<rbrakk> \<Longrightarrow> \<turnstile> \<box>N \<and> \<box>\<diamond>A \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   815
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   816
  apply (erule (2) ensures [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   817
  apply (force elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   818
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   819
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   820
lemma EnsuresInfinite:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   821
    "\<lbrakk> sigma \<Turnstile> \<box>\<diamond>P; sigma \<Turnstile> \<box>A; \<turnstile> A \<and> $P \<longrightarrow> Q` \<rbrakk> \<Longrightarrow> sigma \<Turnstile> \<box>\<diamond>Q"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   822
  apply (erule leadsto_infinite [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   823
  apply (erule EnsuresLeadsto [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   824
  apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   825
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   826
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   827
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   828
(*** Gronning's lattice rules (taken from TLP) ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   829
section "Lattice rules"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   830
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   831
lemma LatticeReflexivity: "\<turnstile> F \<leadsto> F"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   832
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   833
  apply (rule necT InitDmd_gen)+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   834
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   835
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   836
lemma LatticeTransitivity: "\<turnstile> (G \<leadsto> H) \<and> (F \<leadsto> G) \<longrightarrow> (F \<leadsto> H)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   837
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   838
  apply clarsimp
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   839
  apply (erule dup_boxE) (* \<box>\<box>(Init G \<longrightarrow> H) *)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   840
  apply merge_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   841
  apply (clarsimp elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   842
  apply (rule dup_dmdD)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   843
  apply (subgoal_tac "sigmaa \<Turnstile> \<diamond>Init G")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   844
   apply (erule DmdImpl2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   845
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   846
  apply (simp add: dmdInitD)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   847
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   848
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   849
lemma LatticeDisjunctionElim1: "\<turnstile> (F \<or> G \<leadsto> H) \<longrightarrow> (F \<leadsto> H)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   850
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   851
  apply (auto simp: Init_simps elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   852
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   853
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   854
lemma LatticeDisjunctionElim2: "\<turnstile> (F \<or> G \<leadsto> H) \<longrightarrow> (G \<leadsto> H)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   855
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   856
  apply (auto simp: Init_simps elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   857
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   858
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   859
lemma LatticeDisjunctionIntro: "\<turnstile> (F \<leadsto> H) \<and> (G \<leadsto> H) \<longrightarrow> (F \<or> G \<leadsto> H)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   860
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   861
  apply clarsimp
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   862
  apply merge_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   863
  apply (auto simp: Init_simps elim!: STL4E [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   864
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   865
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   866
lemma LatticeDisjunction: "\<turnstile> (F \<or> G \<leadsto> H) = ((F \<leadsto> H) \<and> (G \<leadsto> H))"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   867
  by (auto intro: LatticeDisjunctionIntro [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   868
    LatticeDisjunctionElim1 [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   869
    LatticeDisjunctionElim2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   870
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   871
lemma LatticeDiamond: "\<turnstile> (A \<leadsto> B \<or> C) \<and> (B \<leadsto> D) \<and> (C \<leadsto> D) \<longrightarrow> (A \<leadsto> D)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   872
  apply clarsimp
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   873
  apply (subgoal_tac "sigma \<Turnstile> (B \<or> C) \<leadsto> D")
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   874
  apply (erule_tac G = "LIFT (B \<or> C)" in LatticeTransitivity [temp_use])
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 42814
diff changeset
   875
   apply (fastforce intro!: LatticeDisjunctionIntro [temp_use])+
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   876
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   877
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   878
lemma LatticeTriangle: "\<turnstile> (A \<leadsto> D \<or> B) \<and> (B \<leadsto> D) \<longrightarrow> (A \<leadsto> D)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   879
  apply clarsimp
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   880
  apply (subgoal_tac "sigma \<Turnstile> (D \<or> B) \<leadsto> D")
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   881
   apply (erule_tac G = "LIFT (D \<or> B)" in LatticeTransitivity [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   882
  apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   883
  apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   884
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   885
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   886
lemma LatticeTriangle2: "\<turnstile> (A \<leadsto> B \<or> D) \<and> (B \<leadsto> D) \<longrightarrow> (A \<leadsto> D)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   887
  apply clarsimp
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   888
  apply (subgoal_tac "sigma \<Turnstile> B \<or> D \<leadsto> D")
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   889
   apply (erule_tac G = "LIFT (B \<or> D)" in LatticeTransitivity [temp_use])
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   890
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   891
  apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   892
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   893
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   894
(*** Lamport's fairness rules ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   895
section "Fairness rules"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   896
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   897
lemma WF1:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   898
  "\<lbrakk> \<turnstile> $P \<and> N  \<longrightarrow> P` \<or> Q`;
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   899
      \<turnstile> ($P \<and> N) \<and> <A>_v \<longrightarrow> Q`;
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   900
      \<turnstile> $P \<and> N \<longrightarrow> $(Enabled(<A>_v)) \<rbrakk>
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   901
  \<Longrightarrow> \<turnstile> \<box>N \<and> WF(A)_v \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   902
  apply (clarsimp dest!: BoxWFI [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   903
  apply (erule (2) ensures [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   904
  apply (erule (1) STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   905
  apply (clarsimp simp: WF_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   906
  apply (rule STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   907
  apply (clarsimp elim!: mp intro!: InitDmd [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   908
  apply (erule STL4 [temp_use, THEN box_stp_actD [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   909
  apply (simp add: split_box_conj box_stp_actI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   910
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   911
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   912
(* Sometimes easier to use; designed for action B rather than state predicate Q *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   913
lemma WF_leadsto:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   914
  assumes 1: "\<turnstile> N \<and> $P \<longrightarrow> $Enabled (<A>_v)"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   915
    and 2: "\<turnstile> N \<and> <A>_v \<longrightarrow> B"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   916
    and 3: "\<turnstile> \<box>(N \<and> [\<not>A]_v) \<longrightarrow> stable P"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   917
  shows "\<turnstile> \<box>N \<and> WF(A)_v \<longrightarrow> (P \<leadsto> B)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   918
  apply (unfold leadsto_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   919
  apply (clarsimp dest!: BoxWFI [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   920
  apply (erule (1) STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   921
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   922
  apply (rule 2 [THEN DmdImpl, temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   923
  apply (rule BoxDmd_simple [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   924
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   925
  apply (rule classical)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   926
  apply (rule STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   927
  apply (clarsimp simp: WF_def elim!: mp intro!: InitDmd [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   928
  apply (rule 1 [THEN STL4, temp_use, THEN box_stp_actD])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   929
  apply (simp (no_asm_simp) add: split_box_conj [try_rewrite] box_stp_act [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   930
  apply (erule INV1 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   931
  apply (rule 3 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   932
  apply (simp add: split_box_conj [try_rewrite] NotDmd [temp_use] not_angle [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   933
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   934
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   935
lemma SF1:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   936
  "\<lbrakk> \<turnstile> $P \<and> N  \<longrightarrow> P` \<or> Q`;
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   937
      \<turnstile> ($P \<and> N) \<and> <A>_v \<longrightarrow> Q`;
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   938
      \<turnstile> \<box>P \<and> \<box>N \<and> \<box>F \<longrightarrow> \<diamond>Enabled(<A>_v) \<rbrakk>
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   939
  \<Longrightarrow> \<turnstile> \<box>N \<and> SF(A)_v \<and> \<box>F \<longrightarrow> (P \<leadsto> Q)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   940
  apply (clarsimp dest!: BoxSFI [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   941
  apply (erule (2) ensures [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   942
  apply (erule_tac F = F in dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   943
  apply merge_temp_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   944
  apply (erule STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   945
  apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   946
  apply (clarsimp simp: SF_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   947
  apply (rule STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   948
  apply (erule mp)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   949
  apply (erule STL4 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   950
  apply (simp add: split_box_conj [try_rewrite] STL3 [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   951
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   952
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   953
lemma WF2:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   954
  assumes 1: "\<turnstile> N \<and> <B>_f \<longrightarrow> <M>_g"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   955
    and 2: "\<turnstile> $P \<and> P` \<and> <N \<and> A>_f \<longrightarrow> B"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   956
    and 3: "\<turnstile> P \<and> Enabled(<M>_g) \<longrightarrow> Enabled(<A>_f)"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   957
    and 4: "\<turnstile> \<box>(N \<and> [\<not>B]_f) \<and> WF(A)_f \<and> \<box>F \<and> \<diamond>\<box>Enabled(<M>_g) \<longrightarrow> \<diamond>\<box>P"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   958
  shows "\<turnstile> \<box>N \<and> WF(A)_f \<and> \<box>F \<longrightarrow> WF(M)_g"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   959
  apply (clarsimp dest!: BoxWFI [temp_use] BoxDmdBox [temp_use, THEN iffD2]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   960
    simp: WF_def [where A = M])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   961
  apply (erule_tac F = F in dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   962
  apply merge_temp_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   963
  apply (erule STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   964
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   965
  apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   966
  apply (rule classical)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   967
  apply (subgoal_tac "sigmaa \<Turnstile> \<diamond> (($P \<and> P` \<and> N) \<and> <A>_f)")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   968
   apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   969
  apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   970
  apply (simp add: NotDmd [temp_use] not_angle [try_rewrite])
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   971
  apply merge_act_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   972
  apply (frule 4 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   973
     apply assumption+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   974
  apply (drule STL6 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   975
   apply assumption
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   976
  apply (erule_tac V = "sigmaa \<Turnstile> \<diamond>\<box>P" in thin_rl)
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
   977
  apply (erule_tac V = "sigmaa \<Turnstile> \<box>F" in thin_rl)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   978
  apply (drule BoxWFI [temp_use])
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   979
  apply (erule_tac F = "ACT N \<and> [\<not>B]_f" in dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
   980
  apply merge_temp_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   981
  apply (erule DmdImpldup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   982
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   983
  apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   984
    WF_Box [try_rewrite] box_stp_act [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   985
   apply (force elim!: TLA2E [where P = P, temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   986
  apply (rule STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   987
  apply (force simp: WF_def split_box_conj [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   988
    elim!: mp intro!: InitDmd [temp_use] 3 [THEN STL4, temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   989
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   990
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   991
lemma SF2:
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   992
  assumes 1: "\<turnstile> N \<and> <B>_f \<longrightarrow> <M>_g"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   993
    and 2: "\<turnstile> $P \<and> P` \<and> <N \<and> A>_f \<longrightarrow> B"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   994
    and 3: "\<turnstile> P \<and> Enabled(<M>_g) \<longrightarrow> Enabled(<A>_f)"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   995
    and 4: "\<turnstile> \<box>(N \<and> [\<not>B]_f) \<and> SF(A)_f \<and> \<box>F \<and> \<box>\<diamond>Enabled(<M>_g) \<longrightarrow> \<diamond>\<box>P"
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
   996
  shows "\<turnstile> \<box>N \<and> SF(A)_f \<and> \<box>F \<longrightarrow> SF(M)_g"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   997
  apply (clarsimp dest!: BoxSFI [temp_use] simp: 2 [try_rewrite] SF_def [where A = M])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   998
  apply (erule_tac F = F in dup_boxE)
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
   999
  apply (erule_tac F = "TEMP \<diamond>Enabled (<M>_g) " in dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
  1000
  apply merge_temp_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1001
  apply (erule STL4Edup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1002
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1003
  apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1004
  apply (rule classical)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1005
  apply (subgoal_tac "sigmaa \<Turnstile> \<diamond> (($P \<and> P` \<and> N) \<and> <A>_f)")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1006
   apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1007
  apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1008
  apply (simp add: NotDmd [temp_use] not_angle [try_rewrite])
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
  1009
  apply merge_act_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1010
  apply (frule 4 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1011
     apply assumption+
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1012
  apply (erule_tac V = "sigmaa \<Turnstile> \<box>F" in thin_rl)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1013
  apply (drule BoxSFI [temp_use])
60587
0318b43ee95c more symbols;
wenzelm
parents: 59780
diff changeset
  1014
  apply (erule_tac F = "TEMP \<diamond>Enabled (<M>_g)" in dup_boxE)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1015
  apply (erule_tac F = "ACT N \<and> [\<not>B]_f" in dup_boxE)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
  1016
  apply merge_temp_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1017
  apply (erule DmdImpldup)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1018
   apply assumption
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1019
  apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1020
    SF_Box [try_rewrite] box_stp_act [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1021
   apply (force elim!: TLA2E [where P = P, temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1022
  apply (rule STL2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1023
  apply (force simp: SF_def split_box_conj [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1024
    elim!: mp InfImpl [temp_use] intro!: 3 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1025
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1026
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1027
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1028
(***           Liveness proofs by well-founded orderings                   ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1029
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1030
section "Well-founded orderings"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1031
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1032
lemma wf_leadsto:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1033
  assumes 1: "wf r"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1034
    and 2: "\<And>x. sigma \<Turnstile> F x \<leadsto> (G \<or> (\<exists>y. #((y,x)\<in>r) \<and> F y))    "
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1035
  shows "sigma \<Turnstile> F x \<leadsto> G"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1036
  apply (rule 1 [THEN wf_induct])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1037
  apply (rule LatticeTriangle [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1038
   apply (rule 2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1039
  apply (auto simp: leadsto_exists [try_rewrite])
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1040
  apply (case_tac "(y,x) \<in> r")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1041
   apply force
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1042
  apply (force simp: leadsto_def Init_simps intro!: necT [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1043
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1044
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1045
(* If r is well-founded, state function v cannot decrease forever *)
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1046
lemma wf_not_box_decrease: "\<And>r. wf r \<Longrightarrow> \<turnstile> \<box>[ (v`, $v) \<in> #r ]_v \<longrightarrow> \<diamond>\<box>[#False]_v"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1047
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1048
  apply (rule ccontr)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1049
  apply (subgoal_tac "sigma \<Turnstile> (\<exists>x. v=#x) \<leadsto> #False")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1050
   apply (drule leadsto_false [temp_use, THEN iffD1, THEN STL2_gen [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1051
   apply (force simp: Init_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1052
  apply (clarsimp simp: leadsto_exists [try_rewrite] not_square [try_rewrite] more_temp_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1053
  apply (erule wf_leadsto)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1054
  apply (rule ensures_simple [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1055
   apply (auto simp: square_def angle_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1056
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1057
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1058
(* "wf r  \<Longrightarrow>  \<turnstile> \<diamond>\<box>[ (v`, $v) : #r ]_v \<longrightarrow> \<diamond>\<box>[#False]_v" *)
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1059
lemmas wf_not_dmd_box_decrease =
45605
a89b4bc311a5 eliminated obsolete "standard";
wenzelm
parents: 44890
diff changeset
  1060
  wf_not_box_decrease [THEN DmdImpl, unfolded more_temp_simps]
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1061
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1062
(* If there are infinitely many steps where v decreases, then there
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1063
   have to be infinitely many non-stuttering steps where v doesn't decrease.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1064
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1065
lemma wf_box_dmd_decrease:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1066
  assumes 1: "wf r"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1067
  shows "\<turnstile> \<box>\<diamond>((v`, $v) \<in> #r) \<longrightarrow> \<box>\<diamond><(v`, $v) \<notin> #r>_v"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1068
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1069
  apply (rule ccontr)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1070
  apply (simp add: not_angle [try_rewrite] more_temp_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1071
  apply (drule 1 [THEN wf_not_dmd_box_decrease [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1072
  apply (drule BoxDmdDmdBox [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1073
   apply assumption
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1074
  apply (subgoal_tac "sigma \<Turnstile> \<box>\<diamond> ((#False) ::action)")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1075
   apply force
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1076
  apply (erule STL4E)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1077
  apply (rule DmdImpl)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1078
  apply (force intro: 1 [THEN wf_irrefl, temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1079
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1080
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1081
(* In particular, for natural numbers, if n decreases infinitely often
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1082
   then it has to increase infinitely often.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1083
*)
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1084
lemma nat_box_dmd_decrease: "\<And>n::nat stfun. \<turnstile> \<box>\<diamond>(n` < $n) \<longrightarrow> \<box>\<diamond>($n < n`)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1085
  apply clarsimp
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1086
  apply (subgoal_tac "sigma \<Turnstile> \<box>\<diamond><\<not> ((n`,$n) \<in> #less_than)>_n")
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1087
   apply (erule thin_rl)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1088
   apply (erule STL4E)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1089
   apply (rule DmdImpl)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1090
   apply (clarsimp simp: angle_def [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1091
  apply (rule wf_box_dmd_decrease [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1092
   apply (auto elim!: STL4E [temp_use] DmdImplE [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1093
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1094
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1095
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1096
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1097
(***           Flexible quantification over state variables                ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1098
(* ------------------------------------------------------------------------- *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1099
section "Flexible quantification"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1100
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1101
lemma aallI:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1102
  assumes 1: "basevars vs"
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1103
    and 2: "(\<And>x. basevars (x,vs) \<Longrightarrow> sigma \<Turnstile> F x)"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1104
  shows "sigma \<Turnstile> (\<forall>\<forall>x. F x)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1105
  by (auto simp: aall_def elim!: eexE [temp_use] intro!: 1 dest!: 2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1106
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1107
lemma aallE: "\<turnstile> (\<forall>\<forall>x. F x) \<longrightarrow> F x"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1108
  apply (unfold aall_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1109
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1110
  apply (erule contrapos_np)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1111
  apply (force intro!: eexI [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1112
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1113
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1114
(* monotonicity of quantification *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1115
lemma eex_mono:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1116
  assumes 1: "sigma \<Turnstile> \<exists>\<exists>x. F x"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1117
    and 2: "\<And>x. sigma \<Turnstile> F x \<longrightarrow> G x"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1118
  shows "sigma \<Turnstile> \<exists>\<exists>x. G x"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1119
  apply (rule unit_base [THEN 1 [THEN eexE]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1120
  apply (rule eexI [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1121
  apply (erule 2 [unfolded intensional_rews, THEN mp])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1122
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1123
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1124
lemma aall_mono:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1125
  assumes 1: "sigma \<Turnstile> \<forall>\<forall>x. F(x)"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1126
    and 2: "\<And>x. sigma \<Turnstile> F(x) \<longrightarrow> G(x)"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1127
  shows "sigma \<Turnstile> \<forall>\<forall>x. G(x)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1128
  apply (rule unit_base [THEN aallI])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1129
  apply (rule 2 [unfolded intensional_rews, THEN mp])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1130
  apply (rule 1 [THEN aallE [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1131
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1132
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1133
(* Derived history introduction rule *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1134
lemma historyI:
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1135
  assumes 1: "sigma \<Turnstile> Init I"
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1136
    and 2: "sigma \<Turnstile> \<box>N"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1137
    and 3: "basevars vs"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1138
    and 4: "\<And>h. basevars(h,vs) \<Longrightarrow> \<turnstile> I \<and> h = ha \<longrightarrow> HI h"
60588
750c533459b1 more symbols;
wenzelm
parents: 60587
diff changeset
  1139
    and 5: "\<And>h s t. \<lbrakk> basevars(h,vs); N (s,t); h t = hb (h s) (s,t) \<rbrakk> \<Longrightarrow> HN h (s,t)"
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1140
  shows "sigma \<Turnstile> \<exists>\<exists>h. Init (HI h) \<and> \<box>(HN h)"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1141
  apply (rule history [temp_use, THEN eexE])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1142
  apply (rule 3)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1143
  apply (rule eexI [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1144
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1145
  apply (rule conjI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1146
   prefer 2
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1147
   apply (insert 2)
42787
dd3ab25eb9d1 proper method_setup;
wenzelm
parents: 42773
diff changeset
  1148
   apply merge_box
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1149
   apply (force elim!: STL4E [temp_use] 5 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1150
  apply (insert 1)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1151
  apply (force simp: Init_defs elim!: 4 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1152
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1153
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1154
(* ----------------------------------------------------------------------
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1155
   example of a history variable: existence of a clock
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1156
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1157
60591
e0b77517f9a9 more symbols;
wenzelm
parents: 60588
diff changeset
  1158
lemma "\<turnstile> \<exists>\<exists>h. Init(h = #True) \<and> \<box>(h` = (\<not>$h))"
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1159
  apply (rule tempI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1160
  apply (rule historyI)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1161
  apply (force simp: Init_defs intro!: unit_base [temp_use] necT [temp_use])+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1162
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1163
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
  1164
end