src/HOL/TLA/Inc/Inc.thy
author wenzelm
Sun, 16 Jan 2011 15:53:03 +0100
changeset 41589 bbd861837ebc
parent 26342 0f65fa163304
child 42769 053b4b487085
permissions -rw-r--r--
tuned headers;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41589
bbd861837ebc tuned headers;
wenzelm
parents: 26342
diff changeset
     1
(*  Title:      HOL/TLA/Inc/Inc.thy
bbd861837ebc tuned headers;
wenzelm
parents: 26342
diff changeset
     2
    Author:     Stephan Merz, University of Munich
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     3
*)
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     4
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
     5
header {* Lamport's "increment" example *}
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
     6
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
     7
theory Inc
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
     8
imports TLA
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
     9
begin
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    10
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    11
(* program counter as an enumeration type *)
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    12
datatype pcount = a | b | g
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
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    15
  (* program variables *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    16
  x :: "nat stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    17
  y :: "nat stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    18
  sem :: "nat stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    19
  pc1 :: "pcount stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    20
  pc2 :: "pcount stfun"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    21
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    22
  (* names of actions and predicates *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    23
  M1 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    24
  M2 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    25
  N1 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    26
  N2 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    27
  alpha1 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    28
  alpha2 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    29
  beta1 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    30
  beta2 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    31
  gamma1 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    32
  gamma2 :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    33
  InitPhi :: stpred
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    34
  InitPsi :: stpred
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    35
  PsiInv :: stpred
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    36
  PsiInv1 :: stpred
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    37
  PsiInv2 :: stpred
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    38
  PsiInv3 :: stpred
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    39
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    40
  (* temporal formulas *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    41
  Phi :: temporal
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    42
  Psi :: temporal
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    43
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    44
axioms
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    45
  (* the "base" variables, required to compute enabledness predicates *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    46
  Inc_base:      "basevars (x, y, sem, pc1, pc2)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    47
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    48
  (* definitions for high-level program *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    49
  InitPhi_def:   "InitPhi == PRED x = # 0 & y = # 0"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    50
  M1_def:        "M1      == ACT  x$ = Suc<$x> & y$ = $y"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    51
  M2_def:        "M2      == ACT  y$ = Suc<$y> & x$ = $x"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    52
  Phi_def:       "Phi     == TEMP Init InitPhi & [][M1 | M2]_(x,y)
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    53
                                 & WF(M1)_(x,y) & WF(M2)_(x,y)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    54
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    55
  (* definitions for low-level program *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    56
  InitPsi_def:   "InitPsi == PRED pc1 = #a & pc2 = #a
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    57
                                 & x = # 0 & y = # 0 & sem = # 1"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    58
  alpha1_def:    "alpha1  == ACT  $pc1 = #a & pc1$ = #b & $sem = Suc<sem$>
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    59
                                 & unchanged(x,y,pc2)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    60
  alpha2_def:    "alpha2  == ACT  $pc2 = #a & pc2$ = #b & $sem = Suc<sem$>
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    61
                                 & unchanged(x,y,pc1)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    62
  beta1_def:     "beta1   == ACT  $pc1 = #b & pc1$ = #g & x$ = Suc<$x>
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    63
                                 & unchanged(y,sem,pc2)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    64
  beta2_def:     "beta2   == ACT  $pc2 = #b & pc2$ = #g & y$ = Suc<$y>
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    65
                                 & unchanged(x,sem,pc1)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    66
  gamma1_def:    "gamma1  == ACT  $pc1 = #g & pc1$ = #a & sem$ = Suc<$sem>
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    67
                                 & unchanged(x,y,pc2)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    68
  gamma2_def:    "gamma2  == ACT  $pc2 = #g & pc2$ = #a & sem$ = Suc<$sem>
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    69
                                 & unchanged(x,y,pc1)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    70
  N1_def:        "N1      == ACT  (alpha1 | beta1 | gamma1)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    71
  N2_def:        "N2      == ACT  (alpha2 | beta2 | gamma2)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    72
  Psi_def:       "Psi     == TEMP Init InitPsi
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    73
                               & [][N1 | N2]_(x,y,sem,pc1,pc2)
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    74
                               & SF(N1)_(x,y,sem,pc1,pc2)
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    75
                               & SF(N2)_(x,y,sem,pc1,pc2)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    76
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    77
  PsiInv1_def:  "PsiInv1  == PRED sem = # 1 & pc1 = #a & pc2 = #a"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    78
  PsiInv2_def:  "PsiInv2  == PRED sem = # 0 & pc1 = #a & (pc2 = #b | pc2 = #g)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    79
  PsiInv3_def:  "PsiInv3  == PRED sem = # 0 & pc2 = #a & (pc1 = #b | pc1 = #g)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    80
  PsiInv_def:   "PsiInv   == PRED (PsiInv1 | PsiInv2 | PsiInv3)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
    81
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    82
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    83
lemmas PsiInv_defs = PsiInv_def PsiInv1_def PsiInv2_def PsiInv3_def
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    84
lemmas Psi_defs = Psi_def InitPsi_def N1_def N2_def alpha1_def alpha2_def
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    85
  beta1_def beta2_def gamma1_def gamma2_def
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    86
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    87
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    88
(*** Invariant proof for Psi: "manual" proof proves individual lemmas ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    89
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    90
lemma PsiInv_Init: "|- InitPsi --> PsiInv"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    91
  by (auto simp: InitPsi_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    92
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    93
lemma PsiInv_alpha1: "|- alpha1 & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    94
  by (auto simp: alpha1_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    95
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    96
lemma PsiInv_alpha2: "|- alpha2 & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    97
  by (auto simp: alpha2_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    98
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    99
lemma PsiInv_beta1: "|- beta1 & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   100
  by (auto simp: beta1_def PsiInv_defs)
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 PsiInv_beta2: "|- beta2 & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   103
  by (auto simp: beta2_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   104
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   105
lemma PsiInv_gamma1: "|- gamma1 & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   106
  by (auto simp: gamma1_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   107
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   108
lemma PsiInv_gamma2: "|- gamma2 & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   109
  by (auto simp: gamma2_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   110
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   111
lemma PsiInv_stutter: "|- unchanged (x,y,sem,pc1,pc2) & $PsiInv --> PsiInv$"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   112
  by (auto simp: PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   113
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   114
lemma PsiInv: "|- Psi --> []PsiInv"
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   115
  apply (tactic {* inv_tac (@{clasimpset} addsimps2 [@{thm Psi_def}]) 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   116
   apply (force simp: PsiInv_Init [try_rewrite] Init_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   117
  apply (auto intro: PsiInv_alpha1 [try_rewrite] PsiInv_alpha2 [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   118
    PsiInv_beta1 [try_rewrite] PsiInv_beta2 [try_rewrite] PsiInv_gamma1 [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   119
    PsiInv_gamma2 [try_rewrite] PsiInv_stutter [try_rewrite]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   120
    simp add: square_def N1_def N2_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   121
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   122
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   123
(* Automatic proof works too, but it make take a while on a slow machine.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   124
   More realistic examples require user guidance anyway.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   125
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   126
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   127
lemma "|- Psi --> []PsiInv"
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   128
  by (tactic {* auto_inv_tac (@{simpset} addsimps (@{thms PsiInv_defs} @ @{thms Psi_defs})) 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   129
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   130
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   131
(**** Step simulation ****)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   132
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   133
lemma Init_sim: "|- Psi --> Init InitPhi"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   134
  by (auto simp: InitPhi_def Psi_def InitPsi_def Init_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   135
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   136
lemma Step_sim: "|- Psi --> [][M1 | M2]_(x,y)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   137
  by (auto simp: square_def M1_def M2_def Psi_defs elim!: STL4E [temp_use])
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
(**** Proof of fairness ****)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   141
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   142
(*
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   143
   The goal is to prove Fair_M1 far below, which asserts
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   144
         |- Psi --> WF(M1)_(x,y)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   145
   (the other fairness condition is symmetrical).
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   146
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   147
   The strategy is to use WF2 (with beta1 as the helpful action). Proving its
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   148
   temporal premise needs two auxiliary lemmas:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   149
   1. Stuck_at_b: control can only proceed at pc1 = b by executing beta1
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   150
   2. N1_live: the first component will eventually reach b
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   151
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   152
   Lemma 1 is easy, lemma 2 relies on the invariant, the strong fairness
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   153
   of the semaphore, and needs auxiliary lemmas that ensure that the second
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   154
   component will eventually release the semaphore. Most of the proofs of
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   155
   the auxiliary lemmas are very similar.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   156
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   157
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   158
lemma Stuck_at_b: "|- [][(N1 | N2) & ~ beta1]_(x,y,sem,pc1,pc2) --> stable(pc1 = #b)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   159
  by (auto elim!: Stable squareE simp: Psi_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   160
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   161
lemma N1_enabled_at_g: "|- pc1 = #g --> Enabled (<N1>_(x,y,sem,pc1,pc2))"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   162
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   163
  apply (rule_tac F = gamma1 in enabled_mono)
24180
9f818139951b tuned ML setup;
wenzelm
parents: 21624
diff changeset
   164
   apply (tactic {* enabled_tac @{clasimpset} @{thm Inc_base} 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   165
  apply (force simp: gamma1_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   166
  apply (force simp: angle_def gamma1_def N1_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   167
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   168
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   169
lemma g1_leadsto_a1:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   170
  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N1)_(x,y,sem,pc1,pc2) & []#True  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   171
    --> (pc1 = #g ~> pc1 = #a)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   172
  apply (rule SF1)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   173
    apply (tactic
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   174
      {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   175
   apply (tactic
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   176
      {* action_simp_tac (@{simpset} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   177
  (* reduce |- []A --> <>Enabled B  to  |- A --> Enabled B *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   178
  apply (auto intro!: InitDmd_gen [temp_use] N1_enabled_at_g [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   179
    dest!: STL2_gen [temp_use] simp: Init_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   180
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   181
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   182
(* symmetrical for N2, and similar for beta2 *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   183
lemma N2_enabled_at_g: "|- pc2 = #g --> Enabled (<N2>_(x,y,sem,pc1,pc2))"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   184
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   185
  apply (rule_tac F = gamma2 in enabled_mono)
24180
9f818139951b tuned ML setup;
wenzelm
parents: 21624
diff changeset
   186
  apply (tactic {* enabled_tac @{clasimpset} @{thm Inc_base} 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   187
   apply (force simp: gamma2_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   188
  apply (force simp: angle_def gamma2_def N2_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   189
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   190
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   191
lemma g2_leadsto_a2:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   192
  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   193
    --> (pc2 = #g ~> pc2 = #a)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   194
  apply (rule SF1)
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   195
  apply (tactic {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *})
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   196
  apply (tactic {* action_simp_tac (@{simpset} addsimps @{thm angle_def} :: @{thms Psi_defs})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   197
    [] [] 1 *})
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   198
  apply (auto intro!: InitDmd_gen [temp_use] N2_enabled_at_g [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   199
    dest!: STL2_gen [temp_use] simp add: Init_def)
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
lemma N2_enabled_at_b: "|- pc2 = #b --> Enabled (<N2>_(x,y,sem,pc1,pc2))"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   203
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   204
  apply (rule_tac F = beta2 in enabled_mono)
24180
9f818139951b tuned ML setup;
wenzelm
parents: 21624
diff changeset
   205
   apply (tactic {* enabled_tac @{clasimpset} @{thm Inc_base} 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   206
   apply (force simp: beta2_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   207
  apply (force simp: angle_def beta2_def N2_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   208
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   209
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   210
lemma b2_leadsto_g2:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   211
  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   212
    --> (pc2 = #b ~> pc2 = #g)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   213
  apply (rule SF1)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   214
    apply (tactic
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   215
      {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   216
   apply (tactic
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   217
     {* action_simp_tac (@{simpset} addsimps @{thm angle_def} :: @{thms Psi_defs}) [] [] 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   218
  apply (auto intro!: InitDmd_gen [temp_use] N2_enabled_at_b [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   219
    dest!: STL2_gen [temp_use] simp: Init_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   220
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   221
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   222
(* Combine above lemmas: the second component will eventually reach pc2 = a *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   223
lemma N2_leadsto_a:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   224
  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   225
    --> (pc2 = #a | pc2 = #b | pc2 = #g ~> pc2 = #a)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   226
  apply (auto intro!: LatticeDisjunctionIntro [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   227
    apply (rule LatticeReflexivity [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   228
   apply (rule LatticeTransitivity [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   229
  apply (auto intro!: b2_leadsto_g2 [temp_use] g2_leadsto_a2 [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   230
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   231
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   232
(* Get rid of disjunction on the left-hand side of ~> above. *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   233
lemma N2_live:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   234
  "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2)  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   235
    --> <>(pc2 = #a)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   236
  apply (auto simp: Init_defs intro!: N2_leadsto_a [temp_use, THEN [2] leadsto_init [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   237
  apply (case_tac "pc2 (st1 sigma)")
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   238
    apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   239
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   240
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   241
(* Now prove that the first component will eventually reach pc1 = b from pc1 = a *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   242
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   243
lemma N1_enabled_at_both_a:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   244
  "|- pc2 = #a & (PsiInv & pc1 = #a) --> Enabled (<N1>_(x,y,sem,pc1,pc2))"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   245
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   246
  apply (rule_tac F = alpha1 in enabled_mono)
24180
9f818139951b tuned ML setup;
wenzelm
parents: 21624
diff changeset
   247
  apply (tactic {* enabled_tac @{clasimpset} @{thm Inc_base} 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   248
   apply (force simp: alpha1_def PsiInv_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   249
  apply (force simp: angle_def alpha1_def N1_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   250
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   251
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   252
lemma a1_leadsto_b1:
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   253
  "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))       
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   254
         & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)   
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   255
         --> (pc1 = #a ~> pc1 = #b)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   256
  apply (rule SF1)
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   257
  apply (tactic {* action_simp_tac (@{simpset} addsimps @{thms Psi_defs}) [] [@{thm squareE}] 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   258
  apply (tactic
26342
0f65fa163304 more antiquotations;
wenzelm
parents: 24180
diff changeset
   259
    {* action_simp_tac (@{simpset} addsimps (@{thm angle_def} :: @{thms Psi_defs})) [] [] 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   260
  apply (clarsimp intro!: N1_enabled_at_both_a [THEN DmdImpl [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   261
  apply (auto intro!: BoxDmd2_simple [temp_use] N2_live [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   262
    simp: split_box_conj more_temp_simps)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   263
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   264
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   265
(* Combine the leadsto properties for N1: it will arrive at pc1 = b *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   266
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   267
lemma N1_leadsto_b: "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))              
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   268
         & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)   
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   269
         --> (pc1 = #b | pc1 = #g | pc1 = #a ~> pc1 = #b)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   270
  apply (auto intro!: LatticeDisjunctionIntro [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   271
    apply (rule LatticeReflexivity [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   272
   apply (rule LatticeTransitivity [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   273
    apply (auto intro!: a1_leadsto_b1 [temp_use] g1_leadsto_a1 [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   274
      simp: split_box_conj)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   275
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   276
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   277
lemma N1_live: "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2))              
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   278
         & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2)   
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   279
         --> <>(pc1 = #b)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   280
  apply (auto simp: Init_defs intro!: N1_leadsto_b [temp_use, THEN [2] leadsto_init [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   281
  apply (case_tac "pc1 (st1 sigma)")
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   282
    apply auto
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   283
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   284
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   285
lemma N1_enabled_at_b: "|- pc1 = #b --> Enabled (<N1>_(x,y,sem,pc1,pc2))"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   286
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   287
  apply (rule_tac F = beta1 in enabled_mono)
24180
9f818139951b tuned ML setup;
wenzelm
parents: 21624
diff changeset
   288
   apply (tactic {* enabled_tac @{clasimpset} @{thm Inc_base} 1 *})
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   289
   apply (force simp: beta1_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   290
  apply (force simp: angle_def beta1_def N1_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   291
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   292
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   293
(* Now assemble the bits and pieces to prove that Psi is fair. *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   294
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   295
lemma Fair_M1_lemma: "|- []($PsiInv & [(N1 | N2)]_(x,y,sem,pc1,pc2))    
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   296
         & SF(N1)_(x,y,sem,pc1,pc2) & []SF(N2)_(x,y,sem,pc1,pc2)   
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   297
         --> SF(M1)_(x,y)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   298
  apply (rule_tac B = beta1 and P = "PRED pc1 = #b" in SF2)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   299
   (* action premises *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   300
     apply (force simp: angle_def M1_def beta1_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   301
  apply (force simp: angle_def Psi_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   302
  apply (force elim!: N1_enabled_at_b [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   303
    (* temporal premise: use previous lemmas and simple TL *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   304
  apply (force intro!: DmdStable [temp_use] N1_live [temp_use] Stuck_at_b [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   305
    elim: STL4E [temp_use] simp: square_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   306
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   307
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   308
lemma Fair_M1: "|- Psi --> WF(M1)_(x,y)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   309
  by (auto intro!: SFImplWF [temp_use] Fair_M1_lemma [temp_use] PsiInv [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   310
    simp: Psi_def split_box_conj [temp_use] more_temp_simps)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 11703
diff changeset
   311
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
   312
end