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