src/ZF/UNITY/Constrains.thy
author aspinall
Fri, 30 Sep 2005 18:18:34 +0200
changeset 17740 fc385ce6187d
parent 16183 052d9aba392d
child 21588 cd0dc678a205
permissions -rw-r--r--
Add icon for interface.
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{*Weak Safety Properties*}
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 Constrains
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
     9
imports UNITY
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    10
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    11
begin
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
consts traces :: "[i, i] => i"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
  (* Initial states and program => (final state, reversed trace to it)... 
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    14
      the domain may also be state*list(state) *)
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
inductive 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
  domains 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
     "traces(init, acts)" <=
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
         "(init Un (UN act:acts. field(act)))*list(UN act:acts. field(act))"
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    19
  intros 
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
         (*Initial trace is empty*)
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    21
    Init: "s: init ==> <s,[]> : traces(init,acts)"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    23
    Acts: "[| act:acts;  <s,evs> : traces(init,acts);  <s,s'>: act |]
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
           ==> <s', Cons(s,evs)> : traces(init, acts)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
  
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    26
  type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    29
consts reachable :: "i=>i"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
inductive
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
  domains
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
  "reachable(F)" <= "Init(F) Un (UN act:Acts(F). field(act))"
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    33
  intros 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    34
    Init: "s:Init(F) ==> s:reachable(F)"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    36
    Acts: "[| act: Acts(F);  s:reachable(F);  <s,s'>: act |]
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
           ==> s':reachable(F)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    39
  type_intros UnI1 UnI2 fieldI2 UN_I
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
consts
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
  Constrains :: "[i,i] => i"  (infixl "Co"     60)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
  op_Unless  :: "[i, i] => i"  (infixl "Unless" 60)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
defs
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    47
  Constrains_def:
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    48
    "A Co B == {F:program. F:(reachable(F) Int A) co B}"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    50
  Unless_def:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
    "A Unless B == (A-B) Co (A Un B)"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
constdefs
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
  Stable     :: "i => i"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
    "Stable(A) == A Co A"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
  (*Always is the weak form of "invariant"*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
  Always :: "i => i"
12195
ed2893765a08 *** empty log message ***
ehmety
parents: 11479
diff changeset
    58
    "Always(A) == initially(A) Int Stable(A)"
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    60
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    61
(*** traces and reachable ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    62
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    63
lemma reachable_type: "reachable(F) <= state"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    64
apply (cut_tac F = F in Init_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    65
apply (cut_tac F = F in Acts_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    66
apply (cut_tac F = F in reachable.dom_subset, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    67
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    68
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    69
lemma st_set_reachable: "st_set(reachable(F))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    70
apply (unfold st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    71
apply (rule reachable_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    72
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    73
declare st_set_reachable [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    74
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    75
lemma reachable_Int_state: "reachable(F) Int state = reachable(F)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    76
by (cut_tac reachable_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    77
declare reachable_Int_state [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    78
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    79
lemma state_Int_reachable: "state Int reachable(F) = reachable(F)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    80
by (cut_tac reachable_type, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    81
declare state_Int_reachable [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    82
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    83
lemma reachable_equiv_traces: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    84
"F \<in> program ==> reachable(F)={s \<in> state. \<exists>evs. <s,evs>:traces(Init(F), Acts(F))}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    85
apply (rule equalityI, safe)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    86
apply (blast dest: reachable_type [THEN subsetD])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    87
apply (erule_tac [2] traces.induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    88
apply (erule reachable.induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    89
apply (blast intro: reachable.intros traces.intros)+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    90
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    91
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    92
lemma Init_into_reachable: "Init(F) <= reachable(F)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    93
by (blast intro: reachable.intros)
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
lemma stable_reachable: "[| F \<in> program; G \<in> program;  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    96
    Acts(G) <= Acts(F)  |] ==> G \<in> stable(reachable(F))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    97
apply (blast intro: stableI constrainsI st_setI
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    98
             reachable_type [THEN subsetD] reachable.intros)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
    99
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   100
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   101
declare stable_reachable [intro!]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   102
declare stable_reachable [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   103
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   104
(*The set of all reachable states is an invariant...*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   105
lemma invariant_reachable: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   106
   "F \<in> program ==> F \<in> invariant(reachable(F))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   107
apply (unfold invariant_def initially_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   108
apply (blast intro: reachable_type [THEN subsetD] reachable.intros)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   109
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   110
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   111
(*...in fact the strongest invariant!*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   112
lemma invariant_includes_reachable: "F \<in> invariant(A) ==> reachable(F) <= A"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   113
apply (cut_tac F = F in Acts_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   114
apply (cut_tac F = F in Init_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   115
apply (cut_tac F = F in reachable_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   116
apply (simp (no_asm_use) add: stable_def constrains_def invariant_def initially_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   117
apply (rule subsetI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   118
apply (erule reachable.induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   119
apply (blast intro: reachable.intros)+
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
(*** Co ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   123
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   124
lemma constrains_reachable_Int: "F \<in> B co B'==>F:(reachable(F) Int B) co (reachable(F) Int B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   125
apply (frule constrains_type [THEN subsetD])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   126
apply (frule stable_reachable [OF _ _ subset_refl])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   127
apply (simp_all add: stable_def constrains_Int)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   128
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   129
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   130
(*Resembles the previous definition of Constrains*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   131
lemma Constrains_eq_constrains: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   132
"A Co B = {F \<in> program. F:(reachable(F) Int A) co (reachable(F)  Int  B)}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   133
apply (unfold Constrains_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   134
apply (blast dest: constrains_reachable_Int constrains_type [THEN subsetD]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   135
             intro: constrains_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   136
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   137
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   138
lemmas Constrains_def2 = Constrains_eq_constrains [THEN eq_reflection]
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
lemma constrains_imp_Constrains: "F \<in> A co A' ==> F \<in> A Co A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   141
apply (unfold Constrains_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   142
apply (blast intro: constrains_weaken_L dest: constrainsD2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   143
done
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
lemma ConstrainsI: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   146
    "[|!!act s s'. [| act \<in> Acts(F); <s,s'>:act; s \<in> A |] ==> s':A'; 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   147
       F \<in> program|]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   148
     ==> F \<in> A Co A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   149
apply (auto simp add: Constrains_def constrains_def st_set_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   150
apply (blast dest: reachable_type [THEN subsetD])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   151
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   152
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   153
lemma Constrains_type: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   154
 "A Co B <= program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   155
apply (unfold Constrains_def, blast)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   156
done
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
lemma Constrains_empty: "F \<in> 0 Co B <-> F \<in> program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   159
by (auto dest: Constrains_type [THEN subsetD]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   160
            intro: constrains_imp_Constrains)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   161
declare Constrains_empty [iff]
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
lemma Constrains_state: "F \<in> A Co state <-> F \<in> program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   164
apply (unfold Constrains_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   165
apply (auto dest: Constrains_type [THEN subsetD] intro: constrains_imp_Constrains)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   166
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   167
declare Constrains_state [iff]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   168
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   169
lemma Constrains_weaken_R: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   170
        "[| F \<in> A Co A'; A'<=B' |] ==> F \<in> A Co B'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   171
apply (unfold Constrains_def2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   172
apply (blast intro: constrains_weaken_R)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   173
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   174
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   175
lemma Constrains_weaken_L: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   176
    "[| F \<in> A Co A'; B<=A |] ==> F \<in> B Co A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   177
apply (unfold Constrains_def2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   178
apply (blast intro: constrains_weaken_L st_set_subset)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   179
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   180
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   181
lemma Constrains_weaken: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   182
   "[| F \<in> A Co A'; B<=A; A'<=B' |] ==> F \<in> B Co B'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   183
apply (unfold Constrains_def2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   184
apply (blast intro: constrains_weaken st_set_subset)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   185
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   186
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   187
(** Union **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   188
lemma Constrains_Un: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   189
    "[| F \<in> A Co A'; F \<in> B Co B' |] ==> F \<in> (A Un B) Co (A' Un B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   190
apply (unfold Constrains_def2, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   191
apply (simp add: Int_Un_distrib)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   192
apply (blast intro: constrains_Un)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   193
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   194
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   195
lemma Constrains_UN: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   196
    "[|(!!i. i \<in> I==>F \<in> A(i) Co A'(i)); F \<in> program|] 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   197
     ==> F:(\<Union>i \<in> I. A(i)) Co (\<Union>i \<in> I. A'(i))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   198
by (auto intro: constrains_UN simp del: UN_simps 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   199
         simp add: Constrains_def2 Int_UN_distrib)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   200
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
(** Intersection **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   203
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   204
lemma Constrains_Int: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   205
    "[| F \<in> A Co A'; F \<in> B Co B'|]==> F:(A Int B) Co (A' Int B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   206
apply (unfold Constrains_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   207
apply (subgoal_tac "reachable (F) Int (A Int B) = (reachable (F) Int A) Int (reachable (F) Int B) ")
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   208
apply (auto intro: constrains_Int)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   209
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   210
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   211
lemma Constrains_INT: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   212
    "[| (!!i. i \<in> I ==>F \<in> A(i) Co A'(i)); F \<in> program  |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   213
     ==> F:(\<Inter>i \<in> I. A(i)) Co (\<Inter>i \<in> I. A'(i))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   214
apply (simp (no_asm_simp) del: INT_simps add: Constrains_def INT_extend_simps)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   215
apply (rule constrains_INT)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   216
apply (auto simp add: Constrains_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   217
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   218
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   219
lemma Constrains_imp_subset: "F \<in> A Co A' ==> reachable(F) Int A <= A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   220
apply (unfold Constrains_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   221
apply (blast dest: constrains_imp_subset)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   222
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   223
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   224
lemma Constrains_trans: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   225
 "[| F \<in> A Co B; F \<in> B Co C |] ==> F \<in> A Co C"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   226
apply (unfold Constrains_def2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   227
apply (blast intro: constrains_trans constrains_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   228
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   229
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   230
lemma Constrains_cancel: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   231
"[| F \<in> A Co (A' Un B); F \<in> B Co B' |] ==> F \<in> A Co (A' Un B')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   232
apply (unfold Constrains_def2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   233
apply (simp (no_asm_use) add: Int_Un_distrib)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   234
apply (blast intro: constrains_cancel)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   235
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   236
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   237
(*** Stable ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   238
(* Useful because there's no Stable_weaken.  [Tanja Vos] *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   239
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   240
lemma stable_imp_Stable: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   241
"F \<in> stable(A) ==> F \<in> Stable(A)"
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
apply (unfold stable_def Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   244
apply (erule constrains_imp_Constrains)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   245
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   246
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   247
lemma Stable_eq: "[| F \<in> Stable(A); A = B |] ==> F \<in> Stable(B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   248
by blast
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
lemma Stable_eq_stable: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   251
"F \<in> Stable(A) <->  (F \<in> stable(reachable(F) Int A))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   252
apply (auto dest: constrainsD2 simp add: Stable_def stable_def Constrains_def2)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   253
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   254
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   255
lemma StableI: "F \<in> A Co A ==> F \<in> Stable(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   256
by (unfold Stable_def, assumption)
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 StableD: "F \<in> Stable(A) ==> F \<in> A Co A"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   259
by (unfold Stable_def, assumption)
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 Stable_Un: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   262
    "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable(A Un A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   263
apply (unfold Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   264
apply (blast intro: Constrains_Un)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   265
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   266
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   267
lemma Stable_Int: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   268
    "[| F \<in> Stable(A); F \<in> Stable(A') |] ==> F \<in> Stable (A Int A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   269
apply (unfold Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   270
apply (blast intro: Constrains_Int)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   271
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   272
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   273
lemma Stable_Constrains_Un: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   274
    "[| F \<in> Stable(C); F \<in> A Co (C Un A') |]    
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   275
     ==> F \<in> (C Un A) Co (C Un A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   276
apply (unfold Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   277
apply (blast intro: Constrains_Un [THEN Constrains_weaken_R])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   278
done
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
lemma Stable_Constrains_Int: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   281
    "[| F \<in> Stable(C); F \<in> (C Int A) Co A' |]    
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   282
     ==> F \<in> (C Int A) Co (C Int A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   283
apply (unfold Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   284
apply (blast intro: Constrains_Int [THEN Constrains_weaken])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   285
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   286
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   287
lemma Stable_UN: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   288
    "[| (!!i. i \<in> I ==> F \<in> Stable(A(i))); F \<in> program |]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   289
     ==> F \<in> Stable (\<Union>i \<in> I. A(i))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   290
apply (simp add: Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   291
apply (blast intro: Constrains_UN)
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 Stable_INT: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   295
    "[|(!!i. i \<in> I ==> F \<in> Stable(A(i))); F \<in> program |]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   296
     ==> F \<in> Stable (\<Inter>i \<in> I. A(i))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   297
apply (simp add: Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   298
apply (blast intro: Constrains_INT)
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 Stable_reachable: "F \<in> program ==>F \<in> Stable (reachable(F))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   302
apply (simp (no_asm_simp) add: Stable_eq_stable Int_absorb)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   303
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   304
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   305
lemma Stable_type: "Stable(A) <= program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   306
apply (unfold Stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   307
apply (rule Constrains_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   308
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   309
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   310
(*** The Elimination Theorem.  The "free" m has become universally quantified!
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   311
     Should the premise be !!m instead of \<forall>m ?  Would make it harder to use
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   312
     in forward proof. ***)
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
lemma Elimination: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   315
    "[| \<forall>m \<in> M. F \<in> ({s \<in> A. x(s) = m}) Co (B(m)); F \<in> program |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   316
     ==> F \<in> ({s \<in> A. x(s):M}) Co (\<Union>m \<in> M. B(m))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   317
apply (unfold Constrains_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   318
apply (rule_tac A1 = "reachable (F) Int A" 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   319
	in UNITY.elimination [THEN constrains_weaken_L])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   320
apply (auto intro: constrains_weaken_L)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   321
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   322
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   323
(* As above, but for the special case of A=state *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   324
lemma Elimination2: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   325
 "[| \<forall>m \<in> M. F \<in> {s \<in> state. x(s) = m} Co B(m); F \<in> program |]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   326
     ==> F \<in> {s \<in> state. x(s):M} Co (\<Union>m \<in> M. B(m))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   327
apply (blast intro: Elimination)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   328
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   329
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   330
(** Unless **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   331
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   332
lemma Unless_type: "A Unless B <=program"
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
apply (unfold Unless_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   335
apply (rule Constrains_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   336
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   337
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   338
(*** Specialized laws for handling Always ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   339
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   340
(** Natural deduction rules for "Always A" **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   341
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   342
lemma AlwaysI: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   343
"[| Init(F)<=A;  F \<in> Stable(A) |] ==> F \<in> Always(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   344
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   345
apply (unfold Always_def initially_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   346
apply (frule Stable_type [THEN subsetD], auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   347
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   348
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   349
lemma AlwaysD: "F \<in> Always(A) ==> Init(F)<=A & F \<in> Stable(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   350
by (simp add: Always_def initially_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   351
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   352
lemmas AlwaysE = AlwaysD [THEN conjE, standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   353
lemmas Always_imp_Stable = AlwaysD [THEN conjunct2, standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   354
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   355
(*The set of all reachable states is Always*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   356
lemma Always_includes_reachable: "F \<in> Always(A) ==> reachable(F) <= A"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   357
apply (simp (no_asm_use) add: Stable_def Constrains_def constrains_def Always_def initially_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   358
apply (rule subsetI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   359
apply (erule reachable.induct)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   360
apply (blast intro: reachable.intros)+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   361
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   362
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   363
lemma invariant_imp_Always: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   364
     "F \<in> invariant(A) ==> F \<in> Always(A)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   365
apply (unfold Always_def invariant_def Stable_def stable_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   366
apply (blast intro: constrains_imp_Constrains)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   367
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   368
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   369
lemmas Always_reachable = invariant_reachable [THEN invariant_imp_Always, standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   370
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   371
lemma Always_eq_invariant_reachable: "Always(A) = {F \<in> program. F \<in> invariant(reachable(F) Int A)}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   372
apply (simp (no_asm) add: Always_def invariant_def Stable_def Constrains_def2 stable_def initially_def)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   373
apply (rule equalityI, auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   374
apply (blast intro: reachable.intros reachable_type)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   375
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   376
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   377
(*the RHS is the traditional definition of the "always" operator*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   378
lemma Always_eq_includes_reachable: "Always(A) = {F \<in> program. reachable(F) <= A}"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   379
apply (rule equalityI, safe)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   380
apply (auto dest: invariant_includes_reachable 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   381
   simp add: subset_Int_iff invariant_reachable Always_eq_invariant_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   382
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   383
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   384
lemma Always_type: "Always(A) <= program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   385
by (unfold Always_def initially_def, auto)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   386
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   387
lemma Always_state_eq: "Always(state) = program"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   388
apply (rule equalityI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   389
apply (auto dest: Always_type [THEN subsetD] reachable_type [THEN subsetD]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   390
            simp add: Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   391
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   392
declare Always_state_eq [simp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   393
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   394
lemma state_AlwaysI: "F \<in> program ==> F \<in> Always(state)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   395
by (auto dest: reachable_type [THEN subsetD]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   396
            simp add: Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   397
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   398
lemma Always_eq_UN_invariant: "st_set(A) ==> Always(A) = (\<Union>I \<in> Pow(A). invariant(I))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   399
apply (simp (no_asm) add: Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   400
apply (rule equalityI, auto) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   401
apply (blast intro: invariantI rev_subsetD [OF _ Init_into_reachable] 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   402
		    rev_subsetD [OF _ invariant_includes_reachable]  
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   403
             dest: invariant_type [THEN subsetD])+
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   404
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   405
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   406
lemma Always_weaken: "[| F \<in> Always(A); A <= B |] ==> F \<in> Always(B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   407
by (auto simp add: Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   408
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   409
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   410
(*** "Co" rules involving Always ***)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   411
lemmas Int_absorb2 = subset_Int_iff [unfolded iff_def, THEN conjunct1, THEN mp]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   412
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   413
lemma Always_Constrains_pre: "F \<in> Always(I) ==> (F:(I Int A) Co A') <-> (F \<in> A Co A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   414
apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_def Int_assoc [symmetric])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   415
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   416
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   417
lemma Always_Constrains_post: "F \<in> Always(I) ==> (F \<in> A Co (I Int A')) <->(F \<in> A Co A')"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   418
apply (simp (no_asm_simp) add: Always_includes_reachable [THEN Int_absorb2] Constrains_eq_constrains Int_assoc [symmetric])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   419
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   420
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   421
lemma Always_ConstrainsI: "[| F \<in> Always(I);  F \<in> (I Int A) Co A' |] ==> F \<in> A Co A'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   422
by (blast intro: Always_Constrains_pre [THEN iffD1])
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   423
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   424
(* [| F \<in> Always(I);  F \<in> A Co A' |] ==> F \<in> A Co (I Int A') *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   425
lemmas Always_ConstrainsD = Always_Constrains_post [THEN iffD2, standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   426
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   427
(*The analogous proof of Always_LeadsTo_weaken doesn't terminate*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   428
lemma Always_Constrains_weaken: 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   429
"[|F \<in> Always(C); F \<in> A Co A'; C Int B<=A; C Int A'<=B'|]==>F \<in> B Co B'"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   430
apply (rule Always_ConstrainsI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   431
apply (drule_tac [2] Always_ConstrainsD, simp_all) 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   432
apply (blast intro: Constrains_weaken)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   433
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   434
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   435
(** Conjoining Always properties **)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   436
lemma Always_Int_distrib: "Always(A Int B) = Always(A) Int Always(B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   437
by (auto simp add: Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   438
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   439
(* the premise i \<in> I is need since \<Inter>is formally not defined for I=0 *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   440
lemma Always_INT_distrib: "i \<in> I==>Always(\<Inter>i \<in> I. A(i)) = (\<Inter>i \<in> I. Always(A(i)))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   441
apply (rule equalityI)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   442
apply (auto simp add: Inter_iff Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   443
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   444
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   445
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   446
lemma Always_Int_I: "[| F \<in> Always(A);  F \<in> Always(B) |] ==> F \<in> Always(A Int B)"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   447
apply (simp (no_asm_simp) add: Always_Int_distrib)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   448
done
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   449
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   450
(*Allows a kind of "implication introduction"*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   451
lemma Always_Diff_Un_eq: "[| F \<in> Always(A) |] ==> (F \<in> Always(C-A Un B)) <-> (F \<in> Always(B))"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   452
by (auto simp add: Always_eq_includes_reachable)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   453
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   454
(*Delete the nearest invariance assumption (which will be the second one
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   455
  used by Always_Int_I) *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   456
lemmas Always_thin = thin_rl [of "F \<in> Always(A)", standard]
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   457
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   458
ML
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   459
{*
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   460
val reachable_type = thm "reachable_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   461
val st_set_reachable = thm "st_set_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   462
val reachable_Int_state = thm "reachable_Int_state";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   463
val state_Int_reachable = thm "state_Int_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   464
val reachable_equiv_traces = thm "reachable_equiv_traces";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   465
val Init_into_reachable = thm "Init_into_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   466
val stable_reachable = thm "stable_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   467
val invariant_reachable = thm "invariant_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   468
val invariant_includes_reachable = thm "invariant_includes_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   469
val constrains_reachable_Int = thm "constrains_reachable_Int";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   470
val Constrains_eq_constrains = thm "Constrains_eq_constrains";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   471
val Constrains_def2 = thm "Constrains_def2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   472
val constrains_imp_Constrains = thm "constrains_imp_Constrains";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   473
val ConstrainsI = thm "ConstrainsI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   474
val Constrains_type = thm "Constrains_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   475
val Constrains_empty = thm "Constrains_empty";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   476
val Constrains_state = thm "Constrains_state";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   477
val Constrains_weaken_R = thm "Constrains_weaken_R";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   478
val Constrains_weaken_L = thm "Constrains_weaken_L";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   479
val Constrains_weaken = thm "Constrains_weaken";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   480
val Constrains_Un = thm "Constrains_Un";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   481
val Constrains_UN = thm "Constrains_UN";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   482
val Constrains_Int = thm "Constrains_Int";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   483
val Constrains_INT = thm "Constrains_INT";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   484
val Constrains_imp_subset = thm "Constrains_imp_subset";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   485
val Constrains_trans = thm "Constrains_trans";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   486
val Constrains_cancel = thm "Constrains_cancel";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   487
val stable_imp_Stable = thm "stable_imp_Stable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   488
val Stable_eq = thm "Stable_eq";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   489
val Stable_eq_stable = thm "Stable_eq_stable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   490
val StableI = thm "StableI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   491
val StableD = thm "StableD";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   492
val Stable_Un = thm "Stable_Un";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   493
val Stable_Int = thm "Stable_Int";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   494
val Stable_Constrains_Un = thm "Stable_Constrains_Un";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   495
val Stable_Constrains_Int = thm "Stable_Constrains_Int";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   496
val Stable_UN = thm "Stable_UN";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   497
val Stable_INT = thm "Stable_INT";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   498
val Stable_reachable = thm "Stable_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   499
val Stable_type = thm "Stable_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   500
val Elimination = thm "Elimination";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   501
val Elimination2 = thm "Elimination2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   502
val Unless_type = thm "Unless_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   503
val AlwaysI = thm "AlwaysI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   504
val AlwaysD = thm "AlwaysD";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   505
val AlwaysE = thm "AlwaysE";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   506
val Always_imp_Stable = thm "Always_imp_Stable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   507
val Always_includes_reachable = thm "Always_includes_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   508
val invariant_imp_Always = thm "invariant_imp_Always";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   509
val Always_reachable = thm "Always_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   510
val Always_eq_invariant_reachable = thm "Always_eq_invariant_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   511
val Always_eq_includes_reachable = thm "Always_eq_includes_reachable";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   512
val Always_type = thm "Always_type";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   513
val Always_state_eq = thm "Always_state_eq";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   514
val state_AlwaysI = thm "state_AlwaysI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   515
val Always_eq_UN_invariant = thm "Always_eq_UN_invariant";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   516
val Always_weaken = thm "Always_weaken";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   517
val Int_absorb2 = thm "Int_absorb2";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   518
val Always_Constrains_pre = thm "Always_Constrains_pre";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   519
val Always_Constrains_post = thm "Always_Constrains_post";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   520
val Always_ConstrainsI = thm "Always_ConstrainsI";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   521
val Always_ConstrainsD = thm "Always_ConstrainsD";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   522
val Always_Constrains_weaken = thm "Always_Constrains_weaken";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   523
val Always_Int_distrib = thm "Always_Int_distrib";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   524
val Always_INT_distrib = thm "Always_INT_distrib";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   525
val Always_Int_I = thm "Always_Int_I";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   526
val Always_Diff_Un_eq = thm "Always_Diff_Un_eq";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   527
val Always_thin = thm "Always_thin";
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   528
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   529
(*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   530
val Always_Int_tac = dtac Always_Int_I THEN' assume_tac THEN' etac Always_thin;
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   531
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   532
(*Combines a list of invariance THEOREMS into one.*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   533
val Always_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS Always_Int_I);
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   534
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   535
(*To allow expansion of the program's definition when appropriate*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   536
val program_defs_ref = ref ([]: thm list);
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   537
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   538
(*proves "co" properties when the program is specified*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   539
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   540
fun gen_constrains_tac(cs,ss) i = 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   541
   SELECT_GOAL
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   542
      (EVERY [REPEAT (Always_Int_tac 1),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   543
              REPEAT (etac Always_ConstrainsI 1
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   544
                      ORELSE
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   545
                      resolve_tac [StableI, stableI,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   546
                                   constrains_imp_Constrains] 1),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   547
              rtac constrainsI 1,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   548
              (* Three subgoals *)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   549
              rewrite_goal_tac [st_set_def] 3,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   550
              REPEAT (Force_tac 2),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   551
              full_simp_tac (ss addsimps !program_defs_ref) 1,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   552
              ALLGOALS (clarify_tac cs),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   553
              REPEAT (FIRSTGOAL (etac disjE)),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   554
              ALLGOALS Clarify_tac,
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   555
              REPEAT (FIRSTGOAL (etac disjE)),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   556
              ALLGOALS (clarify_tac cs),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   557
              ALLGOALS (asm_full_simp_tac ss),
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   558
              ALLGOALS (clarify_tac cs)]) i;
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   559
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   560
fun constrains_tac st = gen_constrains_tac (claset(), simpset()) st;
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   561
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   562
(*For proving invariants*)
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   563
fun always_tac i = 
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   564
    rtac AlwaysI i THEN Force_tac i THEN constrains_tac i;
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   565
*}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   566
16183
052d9aba392d renamed "constrains" to "safety" to avoid keyword clash
paulson
parents: 15634
diff changeset
   567
method_setup safety = {*
15634
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   568
    Method.ctxt_args (fn ctxt =>
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   569
        Method.METHOD (fn facts =>
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   570
            gen_constrains_tac (local_clasimpset_of ctxt) 1)) *}
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   571
    "for proving safety properties"
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   572
bca33c49b083 conversion of UNITY to Isar scripts
paulson
parents: 14046
diff changeset
   573
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   574
end
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   575