src/HOL/TLA/Buffer/DBuffer.thy
author paulson
Wed, 04 Jul 2007 13:56:26 +0200
changeset 23563 42f2f90b51a6
parent 21624 6f79647cf536
child 26342 0f65fa163304
permissions -rw-r--r--
simplified a proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     1
(*
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     2
    File:        DBuffer.thy
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
     3
    ID:          $Id$
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     4
    Author:      Stephan Merz
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     5
    Copyright:   1997 University of Munich
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     6
*)
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
     7
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
     8
header {* Two FIFO buffers in a row, with interleaving assumption *}
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
     9
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    10
theory DBuffer
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    11
imports Buffer
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
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
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    15
  (* implementation variables *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    16
  inp  :: "nat stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    17
  mid  :: "nat stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    18
  out  :: "nat stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    19
  q1   :: "nat list stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    20
  q2   :: "nat list stfun"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    21
  qc   :: "nat list stfun"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    22
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    23
  DBInit :: stpred
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    24
  DBEnq :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    25
  DBDeq :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    26
  DBPass :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    27
  DBNext :: action
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    28
  DBuffer :: temporal
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    29
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    30
axioms
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    31
  DB_base:        "basevars (inp,mid,out,q1,q2)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    32
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    33
  (* the concatenation of the two buffers *)
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    34
  qc_def:         "PRED qc == PRED (q2 @ q1)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    35
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    36
  DBInit_def:     "DBInit   == PRED (BInit inp q1 mid  &  BInit mid q2 out)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    37
  DBEnq_def:      "DBEnq    == ACT  Enq inp q1 mid  &  unchanged (q2,out)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    38
  DBDeq_def:      "DBDeq    == ACT  Deq mid q2 out  &  unchanged (inp,q1)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    39
  DBPass_def:     "DBPass   == ACT  Deq inp q1 mid
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    40
                                 & (q2$ = $q2 @ [ mid$ ])
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    41
                                 & (out$ = $out)"
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    42
  DBNext_def:     "DBNext   == ACT  (DBEnq | DBDeq | DBPass)"
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
    43
  DBuffer_def:    "DBuffer  == TEMP Init DBInit
6255
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    44
                                 & [][DBNext]_(inp,mid,out,q1,q2)
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    45
                                 & WF(DBDeq)_(inp,mid,out,q1,q2)
db63752140c7 updated (Stephan Merz);
wenzelm
parents: 3807
diff changeset
    46
                                 & WF(DBPass)_(inp,mid,out,q1,q2)"
3807
82a99b090d9d A formalization of TLA in HOL -- by Stephan Merz;
wenzelm
parents:
diff changeset
    47
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    48
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    49
declare qc_def [simp]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    50
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    51
lemmas db_defs =
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    52
  BInit_def Enq_def Deq_def Next_def IBuffer_def Buffer_def
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    53
  DBInit_def DBEnq_def DBDeq_def DBPass_def DBNext_def DBuffer_def
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    54
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    55
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    56
(*** Proper initialization ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    57
lemma DBInit: "|- Init DBInit --> Init (BInit inp qc out)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    58
  by (auto simp: Init_def DBInit_def BInit_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    59
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    60
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    61
(*** Step simulation ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    62
lemma DB_step_simulation: "|- [DBNext]_(inp,mid,out,q1,q2) --> [Next inp qc out]_(inp,qc,out)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    63
  apply (rule square_simulation)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    64
   apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    65
  apply (tactic
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    66
    {* action_simp_tac (simpset () addsimps (thm "hd_append" :: thms "db_defs")) [] [] 1 *})
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    67
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    68
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    69
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    70
(*** Simulation of fairness ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    71
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    72
(* Compute enabledness predicates for DBDeq and DBPass actions *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    73
lemma DBDeq_visible: "|- <DBDeq>_(inp,mid,out,q1,q2) = DBDeq"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    74
  apply (unfold angle_def DBDeq_def Deq_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    75
  apply (safe, simp (asm_lr))+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    76
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    77
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    78
lemma DBDeq_enabled: 
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    79
    "|- Enabled (<DBDeq>_(inp,mid,out,q1,q2)) = (q2 ~= #[])"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    80
  apply (unfold DBDeq_visible [action_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    81
  apply (force intro!: DB_base [THEN base_enabled, temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    82
    elim!: enabledE simp: angle_def DBDeq_def Deq_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    83
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    84
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    85
lemma DBPass_visible: "|- <DBPass>_(inp,mid,out,q1,q2) = DBPass"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    86
  by (auto simp: angle_def DBPass_def Deq_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    87
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    88
lemma DBPass_enabled: 
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    89
    "|- Enabled (<DBPass>_(inp,mid,out,q1,q2)) = (q1 ~= #[])"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    90
  apply (unfold DBPass_visible [action_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    91
  apply (force intro!: DB_base [THEN base_enabled, temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    92
    elim!: enabledE simp: angle_def DBPass_def Deq_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    93
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    94
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    95
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    96
(* The plan for proving weak fairness at the higher level is to prove
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    97
   (0)  DBuffer => (Enabled (Deq inp qc out) ~> (Deq inp qc out))
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    98
   which is in turn reduced to the two leadsto conditions
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
    99
   (1)  DBuffer => (Enabled (Deq inp qc out) ~> q2 ~= [])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   100
   (2)  DBuffer => (q2 ~= [] ~> DBDeq)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   101
   and the fact that DBDeq implies <Deq inp qc out>_(inp,qc,out)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   102
   (and therefore DBDeq ~> <Deq inp qc out>_(inp,qc,out) trivially holds).
17309
c43ed29bd197 converted to Isar theory format;
wenzelm
parents: 6255
diff changeset
   103
21624
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   104
   Condition (1) is reduced to
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   105
   (1a) DBuffer => (qc ~= [] /\ q2 = [] ~> q2 ~= [])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   106
   by standard leadsto rules (leadsto_classical) and rule Deq_enabledE.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   107
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   108
   Both (1a) and (2) are proved from DBuffer's WF conditions by standard
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   109
   WF reasoning (Lamport's WF1 and WF_leadsto).
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   110
   The condition WF(Deq inp qc out) follows from (0) by rule leadsto_WF.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   111
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   112
   One could use Lamport's WF2 instead.
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   113
*)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   114
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   115
(* Condition (1a) *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   116
lemma DBFair_1a: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2)  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   117
         --> (qc ~= #[] & q2 = #[] ~> q2 ~= #[])"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   118
  apply (rule WF1)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   119
    apply (force simp: db_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   120
   apply (force simp: angle_def DBPass_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   121
  apply (force simp: DBPass_enabled [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   122
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   123
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   124
(* Condition (1) *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   125
lemma DBFair_1: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2)  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   126
         --> (Enabled (<Deq inp qc out>_(inp,qc,out)) ~> q2 ~= #[])"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   127
  apply clarsimp
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   128
  apply (rule leadsto_classical [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   129
  apply (rule DBFair_1a [temp_use, THEN LatticeTransitivity [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   130
  apply assumption+
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   131
  apply (rule ImplLeadsto_gen [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   132
  apply (force intro!: necT [temp_use] dest!: STL2_gen [temp_use] Deq_enabledE [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   133
    simp add: Init_defs)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   134
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   135
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   136
(* Condition (2) *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   137
lemma DBFair_2: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBDeq)_(inp,mid,out,q1,q2)  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   138
         --> (q2 ~= #[] ~> DBDeq)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   139
  apply (rule WF_leadsto)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   140
    apply (force simp: DBDeq_enabled [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   141
   apply (force simp: angle_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   142
  apply (force simp: db_defs elim!: Stable [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   143
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   144
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   145
(* High-level fairness *)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   146
lemma DBFair: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2)  
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   147
                                        & WF(DBDeq)_(inp,mid,out,q1,q2)   
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   148
         --> WF(Deq inp qc out)_(inp,qc,out)"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   149
  apply (auto simp del: qc_def intro!: leadsto_WF [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   150
    DBFair_1 [temp_use, THEN [2] LatticeTransitivity [temp_use]]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   151
    DBFair_2 [temp_use, THEN [2] LatticeTransitivity [temp_use]])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   152
  apply (auto intro!: ImplLeadsto_simple [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   153
    simp: angle_def DBDeq_def Deq_def hd_append [try_rewrite])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   154
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   155
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   156
(*** Main theorem ***)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   157
lemma DBuffer_impl_Buffer: "|- DBuffer --> Buffer inp out"
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   158
  apply (unfold DBuffer_def Buffer_def IBuffer_def)
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   159
  apply (force intro!: eexI [temp_use] DBInit [temp_use]
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   160
    DB_step_simulation [THEN STL4, temp_use] DBFair [temp_use])
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   161
  done
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   162
6f79647cf536 TLA: converted legacy ML scripts;
wenzelm
parents: 17309
diff changeset
   163
end