src/ZF/UNITY/Mutex.thy
author paulson
Mon, 28 Mar 2005 16:19:56 +0200
changeset 15634 bca33c49b083
parent 14046 6616e6c53d48
child 16183 052d9aba392d
permissions -rw-r--r--
conversion of UNITY to Isar scripts
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
     1
(*  ID:         $Id$
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
     6
header{*Mutual Exclusion*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
     7
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
     8
theory Mutex
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
     9
imports SubstAx
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    10
begin
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    11
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    12
text{*Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    13
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    14
Variables' types are introduced globally so that type verification reduces to
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    15
the usual ZF typechecking: an ill-tyed expressions reduce to the empty set.
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    16
*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    17
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
consts
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    19
  p :: i
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    20
  m :: i
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    21
  n :: i
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    22
  u :: i
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    23
  v :: i
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
translations
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    26
  "p" == "Var([0])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    27
  "m" == "Var([1])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    28
  "n" == "Var([0,0])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    29
  "u" == "Var([0,1])"
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    30
  "v" == "Var([1,0])"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
  
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    32
axioms --{** Type declarations  **}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    33
  p_type:  "type_of(p)=bool & default_val(p)=0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    34
  m_type:  "type_of(m)=int  & default_val(m)=#0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    35
  n_type:  "type_of(n)=int  & default_val(n)=#0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    36
  u_type:  "type_of(u)=bool & default_val(u)=0"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    37
  v_type:  "type_of(v)=bool & default_val(v)=0"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
  (** The program for process U **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
   U0 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
    "U0 == {<s,t>:state*state. t = s(u:=1, m:=#1) & s`m = #0}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
  U1 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
  "U1 == {<s,t>:state*state. t = s(p:= s`v, m:=#2) &  s`m = #1}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
  U2 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
    "U2 == {<s,t>:state*state. t = s(m:=#3) & s`p=0 & s`m = #2}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
  U3 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
    "U3 == {<s,t>:state*state. t=s(u:=0, m:=#4) & s`m = #3}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
  U4 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
    "U4 == {<s,t>:state*state. t = s(p:=1, m:=#0) & s`m = #4}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
   (** The program for process V **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
  V0 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
    "V0 == {<s,t>:state*state. t = s (v:=1, n:=#1) & s`n = #0}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
  V1 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
    "V1 == {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) & s`n = #1}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
  V2 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
    "V2 == {<s,t>:state*state. t  = s(n:=#3) & s`p=1 & s`n = #2}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
  V3 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
  "V3 == {<s,t>:state*state. t = s (v:=0, n:=#4) & s`n = #3}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
  V4 :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
    "V4 == {<s,t>:state*state. t  = s (p:=0, n:=#0) & s`n = #4}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
 Mutex :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
 "Mutex == mk_program({s:state. s`u=0 & s`v=0 & s`m = #0 & s`n = #0},
14046
6616e6c53d48 updating ZF-UNITY with Sidi's new material
paulson
parents: 12195
diff changeset
    76
              {U0, U1, U2, U3, U4, V0, V1, V2, V3, V4}, Pow(state*state))"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
  (** The correct invariants **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
  IU :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
    "IU == {s:state. (s`u = 1<->(#1 $<= s`m & s`m $<= #3))
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
	             & (s`m = #3 --> s`p=0)}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
  IV :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
    "IV == {s:state. (s`v = 1 <-> (#1 $<= s`n & s`n $<= #3))
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
	              & (s`n = #3 --> s`p=1)}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
  (** The faulty invariant (for U alone) **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
  bad_IU :: i
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
    "bad_IU == {s:state. (s`u = 1 <-> (#1 $<= s`m & s`m  $<= #3))&
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
                   (#3 $<= s`m & s`m $<= #4 --> s`p=0)}"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    94
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    95
(*  Title:      ZF/UNITY/Mutex.ML
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    96
    ID:         $Id \<in> Mutex.ML,v 1.4 2003/05/27 09:39:05 paulson Exp $
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    97
    Author:     Sidi O Ehmety, Computer Laboratory
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    98
    Copyright   2001  University of Cambridge
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    99
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   100
Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   101
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   102
Variables' types are introduced globally so that type verification
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   103
reduces to the usual ZF typechecking \<in> an ill-tyed expression will
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   104
reduce to the empty set.
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   105
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   106
*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   107
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   108
(** Variables' types **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   109
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   110
declare p_type [simp] u_type [simp] v_type [simp] m_type [simp] n_type [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   111
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   112
lemma u_value_type: "s \<in> state ==>s`u \<in> bool"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   113
apply (unfold state_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   114
apply (drule_tac a = u in apply_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   115
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   116
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   117
lemma v_value_type: "s \<in> state ==> s`v \<in> bool"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   118
apply (unfold state_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   119
apply (drule_tac a = v in apply_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   120
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   121
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   122
lemma p_value_type: "s \<in> state ==> s`p \<in> bool"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   123
apply (unfold state_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   124
apply (drule_tac a = p in apply_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   125
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   126
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   127
lemma m_value_type: "s \<in> state ==> s`m \<in> int"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   128
apply (unfold state_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   129
apply (drule_tac a = m in apply_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   130
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   131
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   132
lemma n_value_type: "s \<in> state ==>s`n \<in> int"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   133
apply (unfold state_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   134
apply (drule_tac a = n in apply_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   135
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   136
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   137
declare p_value_type [simp] u_value_type [simp] v_value_type [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   138
        m_value_type [simp] n_value_type [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   139
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   140
declare p_value_type [TC] u_value_type [TC] v_value_type [TC]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   141
        m_value_type [TC] n_value_type [TC]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   142
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   143
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   144
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   145
text{*Mutex is a program*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   146
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   147
lemma Mutex_in_program [simp,TC]: "Mutex \<in> program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   148
by (simp add: Mutex_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   149
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   150
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   151
method_setup constrains = {*
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   152
    Method.ctxt_args (fn ctxt =>
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   153
        Method.METHOD (fn facts => 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   154
            gen_constrains_tac (local_clasimpset_of ctxt) 1)) *}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   155
    "for proving safety properties"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   156
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   157
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   158
declare Mutex_def [THEN def_prg_Init, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   159
ML
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   160
{*
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   161
program_defs_ref := [thm"Mutex_def"]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   162
*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   163
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   164
declare  U0_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   165
declare  U1_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   166
declare  U2_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   167
declare  U3_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   168
declare  U4_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   169
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   170
declare  V0_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   171
declare  V1_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   172
declare  V2_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   173
declare  V3_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   174
declare  V4_def [THEN def_act_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   175
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   176
declare  U0_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   177
declare  U1_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   178
declare  U2_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   179
declare  U3_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   180
declare  U4_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   181
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   182
declare  V0_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   183
declare  V1_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   184
declare  V2_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   185
declare  V3_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   186
declare  V4_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   187
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   188
declare  IU_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   189
declare  IV_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   190
declare  bad_IU_def [THEN def_set_simp, simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   191
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   192
lemma IU: "Mutex \<in> Always(IU)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   193
apply (rule AlwaysI, force) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   194
apply (unfold Mutex_def, constrains, auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   195
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   196
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   197
lemma IV: "Mutex \<in> Always(IV)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   198
apply (rule AlwaysI, force) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   199
apply (unfold Mutex_def, constrains) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   200
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   201
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   202
(*The safety property: mutual exclusion*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   203
lemma mutual_exclusion: "Mutex \<in> Always({s \<in> state. ~(s`m = #3 & s`n = #3)})"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   204
apply (rule Always_weaken) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   205
apply (rule Always_Int_I [OF IU IV], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   206
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   207
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   208
(*The bad invariant FAILS in V1*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   209
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   210
lemma less_lemma: "[| x$<#1; #3 $<= x |] ==> P"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   211
apply (drule_tac j = "#1" and k = "#3" in zless_zle_trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   212
apply (drule_tac [2] j = x in zle_zless_trans, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   213
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   214
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   215
lemma "Mutex \<in> Always(bad_IU)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   216
apply (rule AlwaysI, force) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   217
apply (unfold Mutex_def, constrains, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   218
apply (subgoal_tac "#1 $<= #3")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   219
apply (drule_tac x = "#1" and y = "#3" in zle_trans, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   220
apply (simp (no_asm) add: not_zless_iff_zle [THEN iff_sym])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   221
apply auto
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   222
(*Resulting state: n=1, p=false, m=4, u=false.  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   223
  Execution of V1 (the command of process v guarded by n=1) sets p:=true,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   224
  violating the invariant!*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   225
oops
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   226
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   227
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   228
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   229
(*** Progress for U ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   230
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   231
lemma U_F0: "Mutex \<in> {s \<in> state. s`m=#2} Unless {s \<in> state. s`m=#3}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   232
by (unfold Unless_def Mutex_def, constrains)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   233
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   234
lemma U_F1:
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   235
     "Mutex \<in> {s \<in> state. s`m=#1} LeadsTo {s \<in> state. s`p = s`v & s`m = #2}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   236
by (unfold Mutex_def, ensures_tac U1)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   237
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   238
lemma U_F2: "Mutex \<in> {s \<in> state. s`p =0 & s`m = #2} LeadsTo {s \<in> state. s`m = #3}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   239
apply (cut_tac IU)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   240
apply (unfold Mutex_def, ensures_tac U2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   241
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   242
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   243
lemma U_F3: "Mutex \<in> {s \<in> state. s`m = #3} LeadsTo {s \<in> state. s`p=1}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   244
apply (rule_tac B = "{s \<in> state. s`m = #4}" in LeadsTo_Trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   245
 apply (unfold Mutex_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   246
 apply (ensures_tac U3)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   247
apply (ensures_tac U4)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   248
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   249
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   250
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   251
lemma U_lemma2: "Mutex \<in> {s \<in> state. s`m = #2} LeadsTo {s \<in> state. s`p=1}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   252
apply (rule LeadsTo_Diff [OF LeadsTo_weaken_L
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   253
                             Int_lower2 [THEN subset_imp_LeadsTo]])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   254
apply (rule LeadsTo_Trans [OF U_F2 U_F3], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   255
apply (auto dest!: p_value_type simp add: bool_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   256
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   257
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   258
lemma U_lemma1: "Mutex \<in> {s \<in> state. s`m = #1} LeadsTo {s \<in> state. s`p =1}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   259
by (rule LeadsTo_Trans [OF U_F1 [THEN LeadsTo_weaken_R] U_lemma2], blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   260
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   261
lemma eq_123: "i \<in> int ==> (#1 $<= i & i $<= #3) <-> (i=#1 | i=#2 | i=#3)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   262
apply auto
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   263
apply (auto simp add: neq_iff_zless)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   264
apply (drule_tac [4] j = "#3" and i = i in zle_zless_trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   265
apply (drule_tac [2] j = i and i = "#1" in zle_zless_trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   266
apply (drule_tac j = i and i = "#1" in zle_zless_trans, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   267
apply (rule zle_anti_sym)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   268
apply (simp_all (no_asm_simp) add: zless_add1_iff_zle [THEN iff_sym])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   269
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   270
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   271
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   272
lemma U_lemma123: "Mutex \<in> {s \<in> state. #1 $<= s`m & s`m $<= #3} LeadsTo {s \<in> state. s`p=1}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   273
by (simp add: eq_123 Collect_disj_eq LeadsTo_Un_distrib U_lemma1 U_lemma2 U_F3)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   274
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   275
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   276
(*Misra's F4*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   277
lemma u_Leadsto_p: "Mutex \<in> {s \<in> state. s`u = 1} LeadsTo {s \<in> state. s`p=1}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   278
by (rule Always_LeadsTo_weaken [OF IU U_lemma123], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   279
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   280
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   281
(*** Progress for V ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   282
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   283
lemma V_F0: "Mutex \<in> {s \<in> state. s`n=#2} Unless {s \<in> state. s`n=#3}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   284
by (unfold Unless_def Mutex_def, constrains)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   285
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   286
lemma V_F1: "Mutex \<in> {s \<in> state. s`n=#1} LeadsTo {s \<in> state. s`p = not(s`u) & s`n = #2}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   287
by (unfold Mutex_def, ensures_tac "V1")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   288
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   289
lemma V_F2: "Mutex \<in> {s \<in> state. s`p=1 & s`n = #2} LeadsTo {s \<in> state. s`n = #3}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   290
apply (cut_tac IV)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   291
apply (unfold Mutex_def, ensures_tac "V2")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   292
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   293
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   294
lemma V_F3: "Mutex \<in> {s \<in> state. s`n = #3} LeadsTo {s \<in> state. s`p=0}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   295
apply (rule_tac B = "{s \<in> state. s`n = #4}" in LeadsTo_Trans)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   296
 apply (unfold Mutex_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   297
 apply (ensures_tac V3)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   298
apply (ensures_tac V4)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   299
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   300
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   301
lemma V_lemma2: "Mutex \<in> {s \<in> state. s`n = #2} LeadsTo {s \<in> state. s`p=0}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   302
apply (rule LeadsTo_Diff [OF LeadsTo_weaken_L
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   303
                             Int_lower2 [THEN subset_imp_LeadsTo]])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   304
apply (rule LeadsTo_Trans [OF V_F2 V_F3], auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   305
apply (auto dest!: p_value_type simp add: bool_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   306
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   307
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   308
lemma V_lemma1: "Mutex \<in> {s \<in> state. s`n = #1} LeadsTo {s \<in> state. s`p = 0}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   309
by (rule LeadsTo_Trans [OF V_F1 [THEN LeadsTo_weaken_R] V_lemma2], blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   310
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   311
lemma V_lemma123: "Mutex \<in> {s \<in> state. #1 $<= s`n & s`n $<= #3} LeadsTo {s \<in> state. s`p = 0}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   312
by (simp add: eq_123 Collect_disj_eq LeadsTo_Un_distrib V_lemma1 V_lemma2 V_F3)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   313
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   314
(*Misra's F4*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   315
lemma v_Leadsto_not_p: "Mutex \<in> {s \<in> state. s`v = 1} LeadsTo {s \<in> state. s`p = 0}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   316
by (rule Always_LeadsTo_weaken [OF IV V_lemma123], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   317
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   318
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   319
(** Absence of starvation **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   320
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   321
(*Misra's F6*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   322
lemma m1_Leadsto_3: "Mutex \<in> {s \<in> state. s`m = #1} LeadsTo {s \<in> state. s`m = #3}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   323
apply (rule LeadsTo_cancel2 [THEN LeadsTo_Un_duplicate])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   324
apply (rule_tac [2] U_F2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   325
apply (simp add: Collect_conj_eq)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   326
apply (subst Un_commute)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   327
apply (rule LeadsTo_cancel2 [THEN LeadsTo_Un_duplicate])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   328
 apply (rule_tac [2] PSP_Unless [OF v_Leadsto_not_p U_F0])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   329
apply (rule U_F1 [THEN LeadsTo_weaken_R], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   330
apply (auto dest!: v_value_type simp add: bool_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   331
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   332
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   333
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   334
(*The same for V*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   335
lemma n1_Leadsto_3: "Mutex \<in> {s \<in> state. s`n = #1} LeadsTo {s \<in> state. s`n = #3}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   336
apply (rule LeadsTo_cancel2 [THEN LeadsTo_Un_duplicate])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   337
apply (rule_tac [2] V_F2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   338
apply (simp add: Collect_conj_eq)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   339
apply (subst Un_commute)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   340
apply (rule LeadsTo_cancel2 [THEN LeadsTo_Un_duplicate])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   341
 apply (rule_tac [2] PSP_Unless [OF u_Leadsto_p V_F0])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   342
apply (rule V_F1 [THEN LeadsTo_weaken_R], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   343
apply (auto dest!: u_value_type simp add: bool_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   344
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   345
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   346
end