src/ZF/UNITY/Guar.ML
author wenzelm
Fri, 28 Sep 2001 19:18:46 +0200
changeset 11627 abf9cda4a4d2
parent 11479 697dcaaf478f
child 12152 46f128d8133c
permissions -rw-r--r--
updated;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      ZF/UNITY/Guar.ML
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Sidi O Ehmety, Computer Laboratory
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   2001  University of Cambridge
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     5
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     6
Guarantees, etc.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
From Chandy and Sanders, "Reasoning About Program Composition"
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     9
Revised by Sidi Ehmety on January 2001
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
Proofs ported from HOL.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
"F~:program ==> F Join G = programify(G)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
by (rtac program_equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
by (auto_tac (claset(),
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
    simpset() addsimps [Join_def, programify_def, SKIP_def, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
                        Acts_def, Init_def, AllowedActs_def, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
                        RawInit_eq, RawActs_eq, RawAllowedActs_eq,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
                        Int_absorb, Int_assoc, Un_absorb]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
by (forward_tac [Id_in_RawActs] 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
by (forward_tac [Id_in_RawAllowedActs] 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
by (dtac   RawInit_type 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
by (dtac   RawActs_type 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
by (dtac   RawAllowedActs_type 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
by (auto_tac (claset(), simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
      addsimps [condition_def, actionSet_def, cons_absorb]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
qed "not_program_Join1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
"G~:program ==> F Join G = programify(F)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
by (stac Join_commute 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
by (blast_tac (claset() addIs [not_program_Join1]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
qed "not_program_Join2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
Goal "F~:program ==> F ok G";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
by (auto_tac (claset(),
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
    simpset() addsimps [ok_def, programify_def, SKIP_def, mk_program_def,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
                        Acts_def, Init_def, AllowedActs_def, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
                        RawInit_def, RawActs_def, RawAllowedActs_def,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
                        Int_absorb, Int_assoc, Un_absorb]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
by (auto_tac (claset(), simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
      addsimps [condition_def, actionSet_def, program_def, cons_absorb]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
qed "not_program_ok1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
Goal "G~:program ==> F ok G";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
by (rtac ok_sym  1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
by (blast_tac (claset() addIs [not_program_ok1]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
qed "not_program_ok2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
Goal "OK(cons(i, I), F) <-> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
\ (i:I & OK(I, F)) | (i~:I & OK(I, F) & F(i) ok JOIN(I,F))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
by (asm_full_simp_tac (simpset() addsimps [OK_iff_ok]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
(** Auto_tac proves the goal in one step, but take more time **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
by (ALLGOALS(Clarify_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
by (REPEAT(blast_tac (claset() addIs [ok_sym]) 10));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
qed "OK_cons_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
(*** existential properties ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
Goalw [ex_prop_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
 "GG:Fin(program) ==> (ex_prop(X) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
\ --> GG Int X~=0 --> OK(GG, (%G. G)) -->(JN G:GG. G):X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
by (etac Fin_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
by (ALLGOALS(asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
         (simpset() addsimps [OK_cons_iff])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
(* Auto_tac proves the goal in one step *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
by (Fast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
qed_spec_mp "ex1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
Goalw [ex_prop_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
     "X<=program ==> (ALL GG. GG:Fin(program) & GG Int X ~= 0\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
\  --> OK(GG,(%G. G)) -->(JN G:GG. G):X) --> ex_prop(X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
by (dres_inst_tac [("x", "{F,G}")] spec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [OK_iff_ok])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
by (auto_tac (claset() addIs [ok_sym], simpset()));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
qed "ex2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
(*Chandy & Sanders take this as a definition*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
Goal "X<=program ==> ex_prop(X) <-> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
\ (ALL GG. GG:Fin(program) & GG Int X ~= 0 &\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
\ OK(GG,( %G. G)) --> (JN G:GG. G):X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    94
by (blast_tac (claset() addIs [ex1, ex2 RS mp]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    95
qed "ex_prop_finite";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    96
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    97
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    98
(*Their "equivalent definition" given at the end of section 3*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    99
Goalw [ex_prop2_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   100
 "ex_prop(X) <-> ex_prop2(X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   101
by (Asm_full_simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   102
by (rewrite_goals_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   103
          [ex_prop_def, component_of_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   104
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   105
by (stac Join_commute 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   106
by (dtac  ok_sym 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   107
by (case_tac "G:program" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   108
by (dres_inst_tac [("x", "G")] bspec 5);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   109
by (dres_inst_tac [("x", "F")] bspec 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   110
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   111
by (force_tac (claset(), 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   112
           simpset() addsimps [not_program_Join1]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   113
by (REPEAT(Force_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   114
qed "ex_prop_equiv";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   115
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   116
(*** universal properties ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   117
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   118
Goalw [uv_prop_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   119
     "GG:Fin(program) ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   120
\ (uv_prop(X)-->  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   121
\  GG <= X & OK(GG, (%G. G)) --> (JN G:GG. G):X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   122
by (etac Fin_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   123
by (auto_tac (claset(), simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   124
           [OK_cons_iff]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   125
qed_spec_mp "uv1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   126
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   127
Goalw [uv_prop_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   128
"X<=program  ==> (ALL GG. GG:Fin(program) & GG <= X & OK(GG,(%G. G)) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   129
\ --> (JN G:GG. G):X)  --> uv_prop(X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   130
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   131
by (Clarify_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   132
by (dres_inst_tac [("x", "{x,xa}")] spec 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   133
by (dres_inst_tac [("x", "0")] spec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   134
by (auto_tac (claset() addDs [ok_sym], 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   135
    simpset() addsimps [OK_iff_ok]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   136
qed "uv2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   137
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   138
(*Chandy & Sanders take this as a definition*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   139
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   140
"X<=program ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   141
\ uv_prop(X) <-> (ALL GG. GG:Fin(program) & GG <= X &\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   142
\    OK(GG, %G. G) --> (JN G:GG. G): X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   143
by (REPEAT(blast_tac (claset() addIs [uv1,uv2 RS mp]) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   144
qed "uv_prop_finite";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   145
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   146
(*** guarantees ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   147
(* The premise G:program is needed later *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   148
val major::prems = Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   149
     "[| (!!G. [| F ok G; F Join G:X; G:program |] ==> F Join G : Y); F:program |]  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   150
\   ==> F : X guarantees Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   151
by (cut_facts_tac prems 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   152
by (simp_tac (simpset() addsimps [guar_def, component_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   153
by (blast_tac (claset() addIs [major]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   154
qed "guaranteesI";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   155
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   156
Goalw [guar_def, component_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   157
     "[| F : X guarantees Y;  F ok G;  F Join G:X; G:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   158
\     ==> F Join G : Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   159
by (Asm_full_simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   160
qed "guaranteesD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   161
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   162
(*This version of guaranteesD matches more easily in the conclusion
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   163
  The major premise can no longer be  F<=H since we need to reason about G*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   164
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   165
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   166
     "[| F : X guarantees Y;  F Join G = H;  H : X;  F ok G; G:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   167
\     ==> H : Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   168
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   169
qed "component_guaranteesD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   170
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   171
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   172
     "[| F: X guarantees X'; Y <= X; X' <= Y' |] ==> F: Y guarantees Y'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   173
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   174
qed "guarantees_weaken";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   175
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   176
Goalw [guar_def] "X <= Y \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   177
\  ==> X guarantees Y = program";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   178
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   179
qed "subset_imp_guarantees_program";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   180
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   181
(*Equivalent to subset_imp_guarantees_UNIV but more intuitive*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   182
Goalw [guar_def] "[| X <= Y; F:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   183
\  ==> F : X guarantees Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   184
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   185
qed "subset_imp_guarantees";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   186
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   187
(*Remark at end of section 4.1 *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   188
Goalw [guar_def, component_of_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   189
"Y<=program ==>ex_prop(Y) --> (Y = (program guarantees Y))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   190
by (simp_tac (simpset() addsimps [ex_prop_equiv]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   191
(* Simplification tactics with ex_prop2_def loops *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   192
by (rewrite_goals_tac [ex_prop2_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   193
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   194
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   195
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   196
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   197
by (dres_inst_tac [("x", "x")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   198
by (dres_inst_tac [("x", "x")] bspec 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   199
by (dtac iff_sym 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   200
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   201
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   202
by (etac iffE 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   203
by (ALLGOALS(full_simp_tac (simpset() addsimps [component_of_def])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   204
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   205
by (REPEAT(Force_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   206
qed "ex_prop_imp";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   207
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   208
Goalw [guar_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   209
  "(Y = program guarantees Y) ==> ex_prop(Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   210
by (asm_simp_tac (simpset() addsimps [ex_prop_equiv]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   211
by (rewrite_goals_tac [ex_prop2_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   212
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   213
by (ALLGOALS(full_simp_tac (simpset() addsimps [component_of_def])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   214
by (dtac sym 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   215
by (ALLGOALS(etac equalityE));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   216
by (REPEAT(Force_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   217
qed "guarantees_imp";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   218
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   219
Goal "Y<=program ==>(ex_prop(Y)) <-> (Y = program guarantees Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   220
by (rtac iffI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   221
by (rtac (ex_prop_imp RS mp) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   222
by (rtac guarantees_imp 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   223
by (ALLGOALS(Asm_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   224
qed "ex_prop_equiv2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   225
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   226
(** Distributive laws.  Re-orient to perform miniscoping **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   227
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   228
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   229
     "i:I ==>(UN i:I. X(i)) guarantees Y = (INT i:I. X(i) guarantees Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   230
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   231
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   232
by (Force_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   233
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   234
qed "guarantees_UN_left";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   235
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   236
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   237
     "(X Un Y) guarantees Z = (X guarantees Z) Int (Y guarantees Z)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   238
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   239
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   240
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   241
qed "guarantees_Un_left";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   242
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   243
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   244
     "i:I ==> X guarantees (INT i:I. Y(i)) = (INT i:I. X guarantees Y(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   245
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   246
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   247
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   248
qed "guarantees_INT_right";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   249
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   250
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   251
     "Z guarantees (X Int Y) = (Z guarantees X) Int (Z guarantees Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   252
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   253
qed "guarantees_Int_right";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   254
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   255
Goal "[| F : Z guarantees X;  F : Z guarantees Y |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   256
\    ==> F : Z guarantees (X Int Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   257
by (asm_simp_tac (simpset() addsimps [guarantees_Int_right]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   258
qed "guarantees_Int_right_I";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   259
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   260
Goal "i:I==> (F : X guarantees (INT i:I. Y(i))) <-> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   261
\     (ALL i:I. F : X guarantees Y(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   262
by (asm_simp_tac (simpset() addsimps [guarantees_INT_right, INT_iff]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   263
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   264
qed "guarantees_INT_right_iff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   265
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   266
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   267
Goalw [guar_def] "(X guarantees Y) = (program guarantees ((program-X) Un Y))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   268
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   269
qed "shunting";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   270
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   271
Goalw [guar_def] "(X guarantees Y) = (program - Y) guarantees (program -X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   272
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   273
qed "contrapositive";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   274
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   275
(** The following two can be expressed using intersection and subset, which
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   276
    is more faithful to the text but looks cryptic.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   277
**)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   278
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   279
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   280
    "[| F : V guarantees X;  F : (X Int Y) guarantees Z |]\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   281
\    ==> F : (V Int Y) guarantees Z";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   282
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   283
qed "combining1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   284
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   285
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   286
    "[| F : V guarantees (X Un Y);  F : Y guarantees Z |]\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   287
\    ==> F : V guarantees (X Un Z)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   288
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   289
qed "combining2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   290
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   291
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   292
(** The following two follow Chandy-Sanders, but the use of object-quantifiers
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   293
    does not suit Isabelle... **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   294
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   295
(*Premise should be (!!i. i: I ==> F: X guarantees Y i) *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   296
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   297
     "[| ALL i:I. F : X guarantees Y(i); i:I |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   298
\   ==> F : X guarantees (INT i:I. Y(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   299
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   300
qed "all_guarantees";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   301
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   302
(*Premises should be [| F: X guarantees Y i; i: I |] *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   303
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   304
     "EX i:I. F : X guarantees Y(i) ==> F : X guarantees (UN i:I. Y(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   305
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   306
qed "ex_guarantees";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   307
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   308
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   309
(*** Additional guarantees laws, by lcp ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   310
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   311
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   312
    "[| F: U guarantees V;  G: X guarantees Y; F ok G |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   313
\    ==> F Join G: (U Int X) guarantees (V Int Y)"; 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   314
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   315
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   316
by (asm_full_simp_tac (simpset() addsimps [Join_assoc]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   317
by (subgoal_tac "F Join G Join x = G Join (F Join x)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   318
by (asm_full_simp_tac (simpset() addsimps [ok_commute]) 1); 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   319
by (asm_simp_tac (simpset() addsimps Join_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   320
qed "guarantees_Join_Int";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   321
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   322
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   323
    "[| F: U guarantees V;  G: X guarantees Y; F ok G |]  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   324
\    ==> F Join G: (U Un X) guarantees (V Un Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   325
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   326
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   327
by (asm_full_simp_tac (simpset() addsimps [Join_assoc]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   328
by (subgoal_tac "F Join G Join x = G Join (F Join x)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   329
by (rotate_tac 4 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   330
by (dres_inst_tac [("x", "F Join x")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   331
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   332
by (force_tac (claset(), simpset() addsimps [ok_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   333
by (asm_simp_tac (simpset() addsimps Join_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   334
qed "guarantees_Join_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   335
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   336
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   337
     "[| ALL i:I. F(i) : X(i) guarantees Y(i);  OK(I,F); i:I |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   338
\     ==> (JN i:I. F(i)) : (INT i:I. X(i)) guarantees (INT i:I. Y(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   339
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   340
by (Blast_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   341
by (dres_inst_tac [("x", "xa")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   342
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [INT_iff])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   343
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   344
by (rotate_tac ~1 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   345
by (dres_inst_tac [("x", "(JN x:(I-{xa}). F(x)) Join G")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   346
by (auto_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   347
    (claset() addIs [OK_imp_ok],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   348
     simpset() addsimps [Join_assoc RS sym, JN_Join_diff, JN_absorb]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   349
qed "guarantees_JN_INT";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   350
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   351
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   352
    "[| ALL i:I. F(i) : X(i) guarantees Y(i);  OK(I,F) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   353
\    ==> JOIN(I,F) : (UN i:I. X(i)) guarantees (UN i:I. Y(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   354
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   355
by (dres_inst_tac [("x", "xa")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   356
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   357
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   358
by (rotate_tac ~1 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   359
by (dres_inst_tac [("x", "JOIN(I-{xa}, F) Join x")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   360
by (auto_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   361
    (claset() addIs [OK_imp_ok],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   362
     simpset() addsimps [Join_assoc RS sym, JN_Join_diff, JN_absorb]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   363
qed "guarantees_JN_UN";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   364
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   365
(*** guarantees laws for breaking down the program, by lcp ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   366
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   367
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   368
     "[| F: X guarantees Y;  F ok G |] ==> F Join G: X guarantees Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   369
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   370
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   371
by (asm_full_simp_tac (simpset() addsimps [Join_assoc]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   372
qed "guarantees_Join_I1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   373
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   374
Goal "[| G: X guarantees Y;  F ok G |] ==> F Join G: X guarantees Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   375
by (asm_full_simp_tac (simpset() addsimps [inst "G" "G" Join_commute, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   376
                                           inst "G" "G" ok_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   377
by (blast_tac (claset() addIs [guarantees_Join_I1]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   378
qed "guarantees_Join_I2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   379
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   380
Goalw [guar_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   381
     "[| i:I; F(i): X guarantees Y;  OK(I,F) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   382
\     ==> (JN i:I. F(i)) : X guarantees Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   383
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   384
by (dres_inst_tac [("x", "JOIN(I-{i},F) Join G")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   385
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   386
by (auto_tac (claset() addIs [OK_imp_ok],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   387
              simpset() addsimps [JN_Join_diff, Join_assoc RS sym]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   388
qed "guarantees_JN_I";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   389
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   390
(*** well-definedness ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   391
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   392
Goalw [welldef_def] "F Join G: welldef ==> programify(F): welldef";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   393
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   394
qed "Join_welldef_D1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   395
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   396
Goalw [welldef_def] "F Join G: welldef ==> programify(G): welldef";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   397
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   398
qed "Join_welldef_D2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   399
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   400
(*** refinement ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   401
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   402
Goalw [refines_def] "F refines F wrt X";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   403
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   404
qed "refines_refl";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   405
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   406
(* Added by Sidi Ehmety from Chandy & Sander, section 6 *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   407
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   408
Goalw [guar_def, component_of_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   409
"(F:X guarantees Y) <-> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   410
\  F:program & (ALL H:program. H:X --> (F component_of H --> H:Y))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   411
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   412
by (REPEAT(Force_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   413
qed "guarantees_equiv";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   414
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   415
Goalw [wg_def] "!!X. [| F:(X guarantees Y); X <= program |] ==> X <= wg(F,Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   416
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   417
qed "wg_weakest";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   418
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   419
Goalw [wg_def, guar_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   420
"F:program ==> F:wg(F,Y) guarantees Y";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   421
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   422
qed "wg_guarantees";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   423
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   424
Goalw [wg_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   425
  "[| F:program; H:program |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   426
\ (H: wg(F,X)) <-> H:program & (F component_of H --> H:X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   427
by (simp_tac (simpset() addsimps [guarantees_equiv]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   428
by (rtac iffI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   429
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   430
by (res_inst_tac [("x", "{H}")] bexI 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   431
by (res_inst_tac [("x", "{H}")] bexI 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   432
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   433
qed "wg_equiv";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   434
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   435
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   436
"[| F component_of H; F:program; H:program |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   437
\ H:wg(F,X) <-> H:X";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   438
by (asm_full_simp_tac (simpset() addsimps [wg_equiv]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   439
qed "component_of_wg";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   440
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   441
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   442
"ALL FF:Fin(program). FF Int X ~= 0 --> OK(FF, %F. F) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   443
\  --> (ALL F:FF. ((JN F:FF. F): wg(F,X)) <-> ((JN F:FF. F):X))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   444
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   445
by (subgoal_tac "F component_of (JN F:FF. F)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   446
by (dres_inst_tac [("X", "X")] component_of_wg 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   447
by (force_tac (claset() addSDs [Fin.dom_subset RS subsetD RS PowD],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   448
               simpset()) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   449
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [component_of_def])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   450
by (res_inst_tac [("x", "JN F:(FF-{F}). F")] exI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   451
by (auto_tac (claset() addIs [JN_Join_diff] addDs [ok_sym], 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   452
              simpset() addsimps [OK_iff_ok]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   453
qed "wg_finite";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   454
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   455
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   456
(* "!!FF. [| FF:Fin(program); FF Int X ~=0; OK(FF, %F. F); G:FF |] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   457
   ==> JOIN(FF, %F. F):wg(G, X) <-> JOIN(FF, %F. F):X"  *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   458
val wg_finite2 = wg_finite RS bspec RS mp RS mp RS bspec;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   459
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   460
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   461
Goal "[| ex_prop(X); F:program |] ==> (F:X) <-> (ALL H:program. H : wg(F,X))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   462
by (asm_full_simp_tac (simpset() addsimps [ex_prop_equiv, wg_equiv]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   463
by (rewrite_goals_tac [ex_prop2_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   464
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   465
qed "wg_ex_prop";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   466
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   467
(** From Charpentier and Chandy "Theorems About Composition" **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   468
(* Proposition 2 *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   469
Goalw [wx_def] "wx(X)<=X";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   470
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   471
qed "wx_subset";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   472
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   473
Goalw [wx_def] "wx(X)<=program";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   474
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   475
qed "wx_into_program";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   476
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   477
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   478
"ex_prop(wx(X))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   479
by (full_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   480
        addsimps [ex_prop_def, wx_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   481
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   482
by (ALLGOALS(res_inst_tac [("x", "xb")] bexI));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   483
by (REPEAT(Force_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   484
qed "wx_ex_prop";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   485
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   486
Goalw [wx_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   487
"ALL Z. Z<=program --> Z<= X --> ex_prop(Z) --> Z <= wx(X)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   488
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   489
qed "wx_weakest";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   490
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   491
(* Proposition 6 *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   492
Goalw [ex_prop_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   493
 "ex_prop({F:program. ALL G:program. F ok G --> F Join G:X})";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   494
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   495
by (dres_inst_tac [("x", "G Join Ga")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   496
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   497
by (force_tac (claset(), simpset() addsimps [ok_Join_iff1, Join_assoc]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   498
by (dres_inst_tac [("x", "F Join Ga")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   499
by (Simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   500
by (full_simp_tac (simpset() addsimps [ok_Join_iff1]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   501
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   502
by (asm_simp_tac (simpset() addsimps [ok_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   503
by (subgoal_tac "F Join G = G Join F" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   504
by (asm_simp_tac (simpset() addsimps [Join_assoc]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   505
by (simp_tac (simpset() addsimps [Join_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   506
qed "wx'_ex_prop";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   507
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   508
(* Equivalence with the other definition of wx *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   509
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   510
Goalw [wx_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   511
 "wx(X) = {F:program. ALL G:program. F ok G --> (F Join G):X}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   512
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   513
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   514
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   515
by (full_simp_tac (simpset() addsimps [ex_prop_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   516
by (dres_inst_tac [("x", "x")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   517
by (dres_inst_tac [("x", "G")] bspec 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   518
by (forw_inst_tac [("c", "x Join G")] subsetD 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   519
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   520
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   521
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   522
by (res_inst_tac [("B", "{F:program. ALL G:program. F ok G --> F Join G:X}")] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   523
                   UnionI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   524
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   525
by (rtac wx'_ex_prop 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   526
by (rotate_tac 2 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   527
by (dres_inst_tac [("x", "SKIP")] bspec 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   528
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   529
qed "wx_equiv";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   530
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   531
(* Propositions 7 to 11 are all about this second definition of wx. And
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   532
   by equivalence between the two definition, they are the same as the ones proved *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   533
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   534
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   535
(* Proposition 12 *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   536
(* Main result of the paper *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   537
Goalw [guar_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   538
   "(X guarantees Y) = wx((program-X) Un Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   539
by (simp_tac (simpset() addsimps [wx_equiv]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   540
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   541
qed "guarantees_wx_eq";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   542
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   543
(* {* Corollary, but this result has already been proved elsewhere *}
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   544
 "ex_prop(X guarantees Y)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   545
  by (simp_tac (simpset() addsimps [guar_wx_iff, wx_ex_prop]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   546
  qed "guarantees_ex_prop";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   547
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   548
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   549
(* Rules given in section 7 of Chandy and Sander's
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   550
    Reasoning About Program composition paper *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   551
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   552
Goal "[| Init(F) <= A; F:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   553
\   ==> F:(stable(A)) guarantees (Always(A))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   554
by (rtac guaranteesI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   555
by (assume_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   556
by (simp_tac (simpset() addsimps [Join_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   557
by (rtac stable_Join_Always1 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   558
by (ALLGOALS(asm_full_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   559
     addsimps [invariant_def, Join_stable])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   560
by (auto_tac (claset(), simpset() addsimps [programify_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   561
qed "stable_guarantees_Always";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   562
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   563
(* To be moved to WFair.ML *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   564
Goal "[| F:A co A Un B; F:transient(A) |] ==> F:A leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   565
by (dres_inst_tac [("B", "A-B")] constrains_weaken_L 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   566
by (dres_inst_tac [("B", "A-B")] transient_strengthen 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   567
by (rtac (ensuresI RS leadsTo_Basis) 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   568
by (ALLGOALS(Blast_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   569
qed "leadsTo_Basis'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   570
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   571
Goal "[| F:transient(A); B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   572
\  ==> F: (A co A Un B) guarantees (A leadsTo (B-A))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   573
by (rtac guaranteesI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   574
by (rtac leadsTo_Basis' 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   575
by (dtac constrains_weaken_R 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   576
by (assume_tac 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   577
by (blast_tac (claset() addIs [Join_transient_I1]) 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   578
by (REPEAT(blast_tac (claset() addDs [transientD]) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   579
qed "constrains_guarantees_leadsTo";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   580
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   581