src/ZF/UNITY/Mutex.ML
author kleing
Mon, 12 May 2003 15:07:11 +0200
changeset 14014 f3f16f9f2030
parent 12195 ed2893765a08
child 14046 6616e6c53d48
permissions -rw-r--r--
done
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      ZF/UNITY/Mutex.ML
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
     8
Variables' types are introduced globally so that type verification
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
     9
reduces to the usual ZF typechecking: an ill-tyed expression will
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    10
reduce to the empty set.
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    11
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
(** Variables' types **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
AddIffs  [p_type, u_type, v_type, m_type, n_type];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
Goal "!!s. s:state ==> s`u:bool";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    19
by (dres_inst_tac [("a", "u"), ("A", "var")] apply_type 1);
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    20
by Auto_tac;
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
qed "u_apply_type";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
Goal "!!s. s:state ==> s`v:bool";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    24
by (dres_inst_tac [("a", "v"), ("A", "var")] apply_type 1);
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    25
by Auto_tac;
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
qed "v_apply_type";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
Goal "!!s. s:state ==> s`p:bool";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    29
by (dres_inst_tac [("a", "p"), ("A", "var")] apply_type 1);
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    30
by Auto_tac;
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
qed "p_apply_type";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
Goal "!!s. s:state ==> s`m:int";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    34
by (dres_inst_tac [("a", "m"), ("A", "var")] apply_type 1);
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    35
by Auto_tac;
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
qed "m_apply_type";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
Goal "!!s. s:state ==> s`n:int";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    39
by (dres_inst_tac [("a", "n"), ("A", "var")] apply_type 1);
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    40
by Auto_tac;
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
qed "n_apply_type";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    42
Addsimps [p_apply_type, u_apply_type, v_apply_type,
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    43
          m_apply_type, n_apply_type];
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
(** Mutex is a program **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    47
Goalw [Mutex_def] "Mutex:program";
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
qed "Mutex_in_program";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
AddIffs [Mutex_in_program];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
Addsimps [Mutex_def RS def_prg_Init];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
program_defs_ref := [Mutex_def];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
Addsimps (map simp_of_act
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
          [U0_def, U1_def, U2_def, U3_def, U4_def, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
           V0_def, V1_def, V2_def, V3_def, V4_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
Addsimps (map simp_of_set [U0_def, U1_def, U2_def, U3_def, U4_def, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
           V0_def, V1_def, V2_def, V3_def, V4_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
Addsimps (map simp_of_set [IU_def, IV_def, bad_IU_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
(** In case one wants to be sure that the initial state is not empty **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
Goal "some(u:=0,v:=0, m:= #0,n:= #0):Init(Mutex)";
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    66
by (auto_tac (claset() addSIs [update_type2], simpset()));
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
qed "Mutex_Init_not_empty";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
Goal "Mutex : Always(IU)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
by (always_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
qed "IU";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
Goal "Mutex : Always(IV)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
by (always_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
qed "IV";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
(*The safety property: mutual exclusion*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
Goal "Mutex : Always({s:state. ~(s`m = #3 & s`n = #3)})";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
by (rtac ([IU, IV] MRS Always_Int_I RS Always_weaken) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
qed "mutual_exclusion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
(*The bad invariant FAILS in V1*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
Goal "[| x$<#1; #3 $<= x |] ==> P";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
by (dres_inst_tac [("j", "#1"), ("k", "#3")]  zless_zle_trans 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
by (dres_inst_tac [("j", "x")]  zle_zless_trans 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
qed "less_lemma";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
Goal "Mutex : Always(bad_IU)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    94
by (always_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    95
by (auto_tac (claset(), simpset() addsimps [not_zle_iff_zless]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    96
by (auto_tac (claset() addSDs [u_apply_type], simpset() addsimps [bool_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    97
by (subgoal_tac "#1 $<= #3" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    98
by (dres_inst_tac [("x", "#1"), ("y", "#3")] zle_trans 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    99
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   100
by (simp_tac (simpset() addsimps [not_zless_iff_zle RS iff_sym]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   101
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   102
(*Resulting state: n=1, p=false, m=4, u=false.  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   103
  Execution of V1 (the command of process v guarded by n=1) sets p:=true,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   104
  violating the invariant!*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   105
(*Check that subgoals remain: proof failed.*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   106
getgoal 1;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   107
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   108
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   109
(*** Progress for U ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   110
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   111
Goalw [Unless_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   112
"Mutex : {s:state. s`m=#2} Unless {s:state. s`m=#3}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   113
by (constrains_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   114
qed "U_F0";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   115
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   116
Goal "Mutex : {s:state. s`m=#1} LeadsTo {s:state. s`p = s`v & s`m = #2}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   117
by (ensures_tac "U1" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   118
qed "U_F1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   119
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   120
Goal "Mutex : {s:state. s`p =0 & s`m = #2} LeadsTo {s:state. s`m = #3}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   121
by (cut_facts_tac [IU] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   122
by (ensures_tac "U2" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   123
qed "U_F2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   124
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   125
Goal "Mutex : {s:state. s`m = #3} LeadsTo {s:state. s`p=1}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   126
by (res_inst_tac [("B", "{s:state. s`m = #4}")] LeadsTo_Trans 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   127
by (ensures_tac "U4" 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   128
by (ensures_tac "U3" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   129
qed "U_F3";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   130
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   131
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   132
Goal "Mutex : {s:state. s`m = #2} LeadsTo {s:state. s`p=1}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   133
by (rtac ([LeadsTo_weaken_L, Int_lower2 RS subset_imp_LeadsTo] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   134
          MRS LeadsTo_Diff) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   135
by (rtac ([U_F2, U_F3] MRS LeadsTo_Trans) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   136
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   137
by (auto_tac (claset() addSDs [p_apply_type], simpset() addsimps [bool_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   138
val U_lemma2 = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   139
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   140
Goal "Mutex : {s:state. s`m = #1} LeadsTo {s:state. s`p =1}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   141
by (rtac ([U_F1 RS LeadsTo_weaken_R, U_lemma2] MRS LeadsTo_Trans) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   142
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   143
val U_lemma1 = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   144
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   145
Goal "i:int ==> (#1 $<= i & i $<= #3) <-> (i=#1 | i=#2 | i=#3)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   146
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   147
by (auto_tac (claset(), simpset() addsimps [neq_iff_zless]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   148
by (dres_inst_tac [("j", "#3"), ("i", "i")] zle_zless_trans 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   149
by (dres_inst_tac [("j", "i"), ("i", "#1")] zle_zless_trans 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   150
by (dres_inst_tac [("j", "i"), ("i", "#1")] zle_zless_trans 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   151
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   152
by (rtac zle_anti_sym 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   153
by (ALLGOALS(asm_simp_tac (simpset()
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   154
      addsimps [zless_add1_iff_zle RS iff_sym])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   155
qed "eq_123";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   156
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   157
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   158
Goal "Mutex : {s:state. #1 $<= s`m & s`m $<= #3} LeadsTo {s:state. s`p=1}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   159
by (simp_tac (simpset() addsimps [m_apply_type RS eq_123, Collect_disj_eq,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   160
                                  LeadsTo_Un_distrib,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   161
                                  U_lemma1, U_lemma2, U_F3] ) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   162
val U_lemma123 = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   163
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   164
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   165
(*Misra's F4*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   166
Goal "Mutex : {s:state. s`u = 1} LeadsTo {s:state. s`p=1}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   167
by (rtac ([IU, U_lemma123] MRS Always_LeadsTo_weaken) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   168
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   169
qed "u_Leadsto_p";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   170
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   171
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   172
(*** Progress for V ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   173
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   174
Goalw [Unless_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   175
"Mutex : {s:state. s`n=#2} Unless {s:state. s`n=#3}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   176
by (constrains_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   177
qed "V_F0";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   178
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   179
Goal "Mutex : {s:state. s`n=#1} LeadsTo {s:state. s`p = not(s`u) & s`n = #2}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   180
by (ensures_tac "V1" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   181
by (auto_tac (claset() addIs [not_type], simpset()));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   182
qed "V_F1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   183
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   184
Goal "Mutex : {s:state. s`p=1 & s`n = #2} LeadsTo {s:state. s`n = #3}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   185
by (cut_facts_tac [IV] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   186
by (ensures_tac "V2" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   187
qed "V_F2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   188
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   189
Goal "Mutex : {s:state. s`n = #3} LeadsTo {s:state. s`p=0}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   190
by (res_inst_tac [("B", "{s:state. s`n = #4}")] LeadsTo_Trans 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   191
by (ensures_tac "V4" 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   192
by (ensures_tac "V3" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   193
qed "V_F3";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   194
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   195
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   196
Goal "Mutex : {s:state. s`n = #2} LeadsTo {s:state. s`p=0}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   197
by (rtac ([LeadsTo_weaken_L, Int_lower2 RS subset_imp_LeadsTo] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   198
          MRS LeadsTo_Diff) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   199
by (rtac ([V_F2, V_F3] MRS LeadsTo_Trans) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   200
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   201
by (auto_tac (claset() addSDs [p_apply_type], simpset() addsimps [bool_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   202
val V_lemma2 = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   203
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   204
Goal "Mutex : {s:state. s`n = #1} LeadsTo {s:state. s`p = 0}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   205
by (rtac ([V_F1 RS LeadsTo_weaken_R, V_lemma2] MRS LeadsTo_Trans) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   206
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   207
val V_lemma1 = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   208
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   209
Goal "Mutex : {s:state. #1 $<= s`n & s`n $<= #3} LeadsTo {s:state. s`p = 0}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   210
by (simp_tac (simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   211
     [n_apply_type RS eq_123, Collect_disj_eq, LeadsTo_Un_distrib,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   212
                  V_lemma1, V_lemma2, V_F3] ) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   213
val V_lemma123 = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   214
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   215
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   216
(*Misra's F4*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   217
Goal "Mutex : {s:state. s`v = 1} LeadsTo {s:state. s`p = 0}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   218
by (rtac ([IV, V_lemma123] MRS Always_LeadsTo_weaken) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   219
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   220
qed "v_Leadsto_not_p";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   221
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   222
(** Absence of starvation **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   223
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   224
(*Misra's F6*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   225
Goal "Mutex : {s:state. s`m = #1} LeadsTo {s:state. s`m = #3}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   226
by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   227
by (rtac U_F2 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   228
by (simp_tac (simpset() addsimps [Collect_conj_eq] ) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   229
by (stac Un_commute 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   230
by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   231
by (rtac ([v_Leadsto_not_p, U_F0] MRS PSP_Unless) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   232
by (rtac (U_F1 RS LeadsTo_weaken_R) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   233
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   234
by (auto_tac (claset() addSDs [v_apply_type], simpset() addsimps [bool_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   235
qed "m1_Leadsto_3";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   236
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   237
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   238
(*The same for V*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   239
Goal "Mutex : {s:state. s`n = #1} LeadsTo {s:state. s`n = #3}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   240
by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   241
by (rtac V_F2 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   242
by (simp_tac (simpset() addsimps [Collect_conj_eq] ) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   243
by (stac Un_commute 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   244
by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   245
by (rtac ([u_Leadsto_p, V_F0] MRS PSP_Unless) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   246
by (rtac (V_F1 RS LeadsTo_weaken_R) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   247
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   248
by (auto_tac (claset() addSDs [u_apply_type], simpset() addsimps [bool_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   249
qed "n1_Leadsto_3";