src/ZF/UNITY/WFair.ML
author paulson
Wed, 08 Aug 2001 14:33:10 +0200
changeset 11479 697dcaaf478f
child 12195 ed2893765a08
permissions -rw-r--r--
new ZF/UNITY theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/WFair.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
Weak Fairness versions of transient, ensures, leadsTo.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming", 1994
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     9
*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    10
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    11
(*** transient ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
Goalw [transient_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
"F:transient(A) ==> F:program & A:condition";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
qed "transientD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
Goalw [stable_def, constrains_def, transient_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
    "[| F : stable(A); F : transient(A) |] ==> A = 0";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
qed "stable_transient_empty";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
Goalw [transient_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
    "[| F : transient(A); B<=A |] ==> F : transient(B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
by (res_inst_tac [("x", "act")] bexI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
by (auto_tac (claset(), 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
   simpset() addsimps [condition_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
qed "transient_strengthen";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
Goalw [transient_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
"[| act:Acts(F); A <= domain(act); act``A <= state-A; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
\       F:program; A:condition |] ==> F : transient(A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
qed "transientI";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
val major::prems = 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
Goalw [transient_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
    "[| F:transient(A); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
\  !!act. [| act:Acts(F);  A <= domain(act); act``A <= state-A |] ==> P |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
\    ==> P";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
by (rtac (major RS CollectE) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
by (blast_tac (claset() addIs prems) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
qed "transientE";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
Goalw [transient_def] "transient(state) = 0";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
by (ALLGOALS(Clarify_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
by (dtac ActsD 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
by (asm_full_simp_tac (simpset() addsimps [Diff_cancel]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
by (blast_tac (claset() addSDs [state_subset_not_empty]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
qed "transient_state";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
Goalw [transient_def] "transient(0) = program";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
by (subgoal_tac "Id:Acts(x)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
by (Asm_simp_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
by (res_inst_tac [("x", "Id")] bexI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
by (ALLGOALS(Blast_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
qed "transient_empty";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
Addsimps [transient_empty, transient_state];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
(*** ensures ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
Goalw [ensures_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
    "[| F : (A-B) co (A Un B); F : transient(A-B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
\          ==> F : A ensures B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
qed "ensuresI";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
(** From Misra's notes, Progress chapter, exercise number 4 **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
Goal "[| F:A co A Un B; F: transient(A) |] ==> F: A ensures B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
by (dres_inst_tac [("B", "A-B")] constrains_weaken_L 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
by (dres_inst_tac [("B", "A-B")] transient_strengthen 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
by (auto_tac (claset(), simpset() addsimps [ensures_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
qed "ensuresI2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
Goalw [ensures_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
    "F : A ensures B ==> F : (A-B) co (A Un B) & F : transient (A-B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
qed "ensuresD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
Goalw [ensures_def, constrains_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
 "F : A ensures B ==> F:program & A:condition & B:condition";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
qed "ensuresD2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    94
Goalw [ensures_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    95
    "[| F : A ensures A'; A'<=B';  B':condition |] ==> F : A ensures B'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    96
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    97
by (blast_tac (claset()  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    98
          addIs [transient_strengthen, constrains_weaken]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    99
          addDs [constrainsD2]) 1);   
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   100
qed "ensures_weaken_R";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   101
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   102
(*The L-version (precondition strengthening) fails, but we have this*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   103
Goalw [ensures_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   104
    "[| F : stable(C);  F : A ensures B |]   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   105
\   ==> F : (C Int A) ensures (C Int B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   106
by (asm_full_simp_tac (simpset() addsimps [ensures_def,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   107
                                  Int_Un_distrib2,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   108
                                  Diff_Int_distrib RS sym]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   109
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   110
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   111
        addIs [transient_strengthen, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   112
               stable_constrains_Int, constrains_weaken] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   113
        addDs [constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   114
qed "stable_ensures_Int"; 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   115
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   116
Goal "[| F : stable(A);  F : transient(C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   117
\                A <= B Un C; B:condition|] ==> F : A ensures B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   118
by (asm_full_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   119
               addsimps [ensures_def, stable_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   120
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   121
by (blast_tac (claset()  addIs [transient_strengthen, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   122
                                constrains_weaken] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   123
                         addDs [constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   124
qed "stable_transient_ensures";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   125
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   126
Goal "(A ensures B) = (A unless B) Int transient (A-B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   127
by (simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   128
       addsimps [ensures_def, unless_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   129
qed "ensures_eq";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   130
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   131
(*** leadsTo ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   132
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   133
val leads_lhs_subset = leads.dom_subset RS subsetD RS SigmaD1;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   134
val leads_rhs_subset = leads.dom_subset RS subsetD RS SigmaD2;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   135
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   136
Goalw [leadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   137
 "F: A leadsTo B ==> F:program & A:condition & B:condition";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   138
by (blast_tac (claset() addDs [leads_lhs_subset,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   139
                               leads_rhs_subset]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   140
qed "leadsToD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   141
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   142
Goalw [leadsTo_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   143
  "F : A ensures B ==> F : A leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   144
by (blast_tac (claset() addDs [ensuresD2]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   145
                        addIs [leads.Basis]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   146
qed "leadsTo_Basis";                       
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   147
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   148
AddIs [leadsTo_Basis];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   149
Addsimps [leadsTo_Basis];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   150
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   151
Goalw [leadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   152
     "[| F : A leadsTo B;  F : B leadsTo C |] ==> F : A leadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   153
by (blast_tac (claset() addIs [leads.Trans]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   154
qed "leadsTo_Trans";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   155
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   156
(* To be move to State.thy *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   157
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   158
Goalw [condition_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   159
    "A:condition ==> state<=A <-> A=state";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   160
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   161
qed "state_upper";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   162
Addsimps [state_upper];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   163
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   164
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   165
Goalw [transient_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   166
    "F : transient(A) ==> F : A leadsTo (state - A )";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   167
by (rtac (ensuresI RS leadsTo_Basis) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   168
by (ALLGOALS(Clarify_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   169
by (blast_tac (claset() addIs [transientI]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   170
by (rtac constrains_weaken 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   171
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   172
qed "transient_imp_leadsTo";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   173
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   174
(*Useful with cancellation, disjunction*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   175
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   176
Goal "F : A leadsTo (A' Un A') ==> F : A leadsTo A'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   177
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   178
qed "leadsTo_Un_duplicate";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   179
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   180
Goal "F : A leadsTo (A' Un C Un C) ==> F : A leadsTo (A' Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   181
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   182
qed "leadsTo_Un_duplicate2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   183
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   184
(*The Union introduction rule as we should have liked to state it*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   185
Goalw [leadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   186
    "[|  ALL A:S. F : A leadsTo B; F:program; B:condition |]\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   187
\ ==> F : Union(S) leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   188
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   189
by (blast_tac (claset() addIs [leads.Union] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   190
                        addDs [leads_lhs_subset]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   191
bind_thm ("leadsTo_Union", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   192
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   193
Goalw [leadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   194
 "[| ALL A:S. F: (A Int C) leadsTo B;  F:program; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   195
 \ ==> F : (Union(S) Int C) leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   196
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   197
by (simp_tac (simpset()  addsimps [Int_Union_Union]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   198
by (blast_tac (claset() addIs [leads.Union] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   199
                        addDs [leads_lhs_subset, leads_rhs_subset]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   200
bind_thm ("leadsTo_Union_Int", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   201
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   202
Goalw [leadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   203
"[| ALL i:I. F : (A(i)) leadsTo B; F:program; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   204
\   ==> F:(UN i:I. A(i)) leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   205
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   206
by (simp_tac (simpset()  addsimps [Int_Union_Union]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   207
by (blast_tac (claset() addIs [leads.Union] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   208
                        addDs [leads_lhs_subset, leads_rhs_subset]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   209
bind_thm ("leadsTo_UN",  ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   210
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   211
(*Binary union introduction rule*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   212
Goal "[| F: A leadsTo C; F: B leadsTo C |] ==> F : (A Un B) leadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   213
by (stac Un_eq_Union 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   214
by (blast_tac (claset() addIs [leadsTo_Union] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   215
                        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   216
qed "leadsTo_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   217
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   218
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   219
Goal "[| ALL x:A. F:{x} leadsTo B; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   220
\        F:program; B:condition |] ==> F : A leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   221
by (res_inst_tac [("b", "A")] (UN_singleton RS subst) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   222
by (blast_tac (claset() addIs [leadsTo_UN]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   223
bind_thm("single_leadsTo_I", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   224
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   225
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   226
(*The INDUCTION rule as we should have liked to state it*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   227
val major::prems = Goalw [leadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   228
  "[| F: za leadsTo zb; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   229
\     !!A B. F : A ensures B ==> P(A, B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   230
\     !!A B C. [| F: A leadsTo B; P(A, B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   231
\                 F: B leadsTo C; P(B, C) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   232
\              ==> P(A, C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   233
\     !!B S. [| ALL A:S. F:A leadsTo B & P(A, B); B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   234
\        ==> P(Union(S), B) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   235
\  |] ==> P(za, zb)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   236
by (cut_facts_tac [major] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   237
by (rtac (major RS CollectD2 RS leads.induct) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   238
by (auto_tac (claset() addIs prems, simpset()));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   239
qed "leadsTo_induct";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   240
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   241
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   242
"[| A<=B; F:program; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   243
\   ==> F : A ensures B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   244
by (rewrite_goals_tac [ensures_def, constrains_def, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   245
                       transient_def, condition_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   246
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   247
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   248
by (res_inst_tac [("x", "Id")] bexI 5);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   249
by (REPEAT(blast_tac (claset() addDs [Id_in_Acts]) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   250
qed "subset_imp_ensures";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   251
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   252
bind_thm ("subset_imp_leadsTo", subset_imp_ensures RS leadsTo_Basis);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   253
bind_thm ("leadsTo_refl", subset_refl RS subset_imp_leadsTo);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   254
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   255
bind_thm ("empty_leadsTo", empty_subsetI RS subset_imp_leadsTo);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   256
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   257
Addsimps [empty_leadsTo];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   258
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   259
Goalw [condition_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   260
  "[| F:program; A:condition |] ==> F: A leadsTo state";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   261
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   262
qed "leadsTo_state";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   263
Addsimps [leadsTo_state];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   264
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   265
(* A nicer induction rule; without ensures *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   266
val [major,impl,basis,trans,unionp] = Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   267
  "[| F: za leadsTo zb; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   268
\     !!A B. [| A<=B; B:condition |] ==> P(A, B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   269
\     !!A B. [| F:A co A Un B; F:transient(A) |] ==> P(A, B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   270
\     !!A B C. [| F: A leadsTo B; P(A, B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   271
\                 F: B leadsTo C; P(B, C) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   272
\              ==> P(A, C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   273
\     !!B S. [| ALL A:S. F:A leadsTo B & P(A, B); B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   274
\        ==> P(Union(S), B) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   275
\  |] ==> P(za, zb)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   276
by (cut_facts_tac [major] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   277
by (etac leadsTo_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   278
by (auto_tac (claset() addIs [trans,unionp], simpset()));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   279
by (rewrite_goal_tac [ensures_def] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   280
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   281
by (forward_tac [constrainsD2] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   282
by (dres_inst_tac [("B'", "(A-B) Un B")] constrains_weaken_R 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   283
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   284
by (forward_tac [ensuresI2 RS leadsTo_Basis] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   285
by (dtac basis 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   286
by (subgoal_tac "A Int B <= B " 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   287
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   288
by (dtac impl 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   289
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   290
by (res_inst_tac [("a", "Union({A - B, A Int B})"), ("P", "%x. P(x, ?u)")] subst 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   291
by (rtac unionp 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   292
by (auto_tac (claset() addIs [subset_imp_leadsTo], simpset()));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   293
qed "leadsTo_induct2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   294
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   295
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   296
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   297
(** Variant induction rule: on the preconditions for B **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   298
(*Lemma is the weak version: can't see how to do it in one step*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   299
val major::prems = Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   300
  "[| F : za leadsTo zb;  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   301
\     P(zb); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   302
\     !!A B. [| F : A ensures B;  P(B) |] ==> P(A); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   303
\     !!S. [| ALL A:S. P(A) |] ==> P(Union(S)) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   304
\  |] ==> P(za)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   305
(*by induction on this formula*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   306
by (subgoal_tac "P(zb) --> P(za)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   307
(*now solve first subgoal: this formula is sufficient*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   308
by (blast_tac (claset() addIs leadsTo_refl::prems) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   309
by (rtac (major RS leadsTo_induct) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   310
by (REPEAT (blast_tac (claset() addIs prems) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   311
qed "lemma";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   312
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   313
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   314
val [major, cond, ensuresp, unionp] = Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   315
  "[| F : za leadsTo zb;  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   316
\     P(zb); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   317
\     !!A B. [| F : A ensures B;  F : B leadsTo zb;  P(B) |] ==> P(A); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   318
\     !!S. ALL A:S. F : A leadsTo zb & P(A) ==> P(Union(S)) \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   319
\  |] ==> P(za)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   320
by (cut_facts_tac [major] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   321
by (subgoal_tac "(F : za leadsTo zb) & P(za)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   322
by (etac conjunct2 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   323
by (rtac (major RS lemma) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   324
by (blast_tac (claset() addDs [leadsToD]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   325
                        addIs [leadsTo_Union,unionp]) 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   326
by (blast_tac (claset() addIs [leadsTo_Trans,ensuresp]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   327
by (blast_tac (claset() addIs [conjI,leadsTo_refl,cond] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   328
                        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   329
qed "leadsTo_induct_pre";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   330
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   331
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   332
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   333
"[| F : A leadsTo A'; A'<=B'; B':condition |]\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   334
\  ==> F : A leadsTo B'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   335
by (blast_tac (claset()  addIs [subset_imp_leadsTo, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   336
                                leadsTo_Trans]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   337
                         addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   338
qed "leadsTo_weaken_R";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   339
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   340
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   341
Goal "[| F : A leadsTo A'; B<=A |] ==> F : B leadsTo A'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   342
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   343
        addIs [leadsTo_Trans, subset_imp_leadsTo]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   344
        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   345
qed_spec_mp "leadsTo_weaken_L";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   346
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   347
(*Distributes over binary unions*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   348
Goal "F:(A Un B) leadsTo C  <->  (F:A leadsTo C & F : B leadsTo C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   349
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   350
qed "leadsTo_Un_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   351
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   352
Goal "[| F:program; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   353
\==> F : (UN i:I. A(i)) leadsTo B  <->  (ALL i : I. F : (A(i)) leadsTo B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   354
by (blast_tac (claset() addIs [leadsTo_UN, leadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   355
qed "leadsTo_UN_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   356
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   357
Goal "[| F:program; B:condition |]   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   358
\==> F : (Union(S)) leadsTo B <->  (ALL A:S. F : A leadsTo B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   359
by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   360
qed "leadsTo_Union_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   361
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   362
Goal 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   363
"[| F : A leadsTo A'; B<=A; A'<=B'; B':condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   364
\   ==> F : B leadsTo B'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   365
by (subgoal_tac "B:condition & A':condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   366
by (force_tac (claset() addSDs [leadsToD],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   367
               simpset() addsimps [condition_def]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   368
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   369
     addIs [leadsTo_weaken_R, leadsTo_weaken_L, leadsTo_Trans]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   370
qed "leadsTo_weaken";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   371
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   372
(*Set difference: maybe combine with leadsTo_weaken_L?*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   373
Goal "[| F : (A-B) leadsTo C; F : B leadsTo C|]   ==> F : A leadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   374
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   375
                        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   376
qed "leadsTo_Diff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   377
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   378
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   379
Goal "[| ALL i:I. F : (A(i)) leadsTo (A'(i)); F:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   380
\   ==> F:(UN i:I. A(i)) leadsTo (UN i:I. A'(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   381
by (rtac leadsTo_Union 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   382
by (ALLGOALS(Clarify_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   383
by (REPEAT(blast_tac (claset()
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   384
      addIs [leadsTo_weaken_R] addDs [leadsToD]) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   385
qed "leadsTo_UN_UN";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   386
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   387
(*Binary union version*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   388
Goal "[| F : A leadsTo A'; F : B leadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   389
\     ==> F : (A Un B) leadsTo (A' Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   390
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken_R]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   391
                        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   392
qed "leadsTo_Un_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   393
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   394
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   395
(** The cancellation law **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   396
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   397
Goal "[| F : A leadsTo (A' Un B); F : B leadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   398
\     ==> F : A leadsTo (A' Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   399
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   400
      addIs [leadsTo_Trans, leadsTo_Un_Un, leadsTo_refl]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   401
      addDs  [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   402
qed "leadsTo_cancel2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   403
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   404
Goal "[| F : A leadsTo (A' Un B); F : (B-A') leadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   405
\     ==> F : A leadsTo (A' Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   406
by (rtac leadsTo_cancel2 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   407
by (assume_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   408
by (blast_tac (claset() addIs [leadsTo_weaken_R]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   409
                        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   410
qed "leadsTo_cancel_Diff2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   411
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   412
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   413
Goal "[| F : A leadsTo (B Un A'); F : B leadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   414
\   ==> F : A leadsTo (B' Un A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   415
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   416
by (blast_tac (claset() addSIs [leadsTo_cancel2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   417
qed "leadsTo_cancel1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   418
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   419
Goal "[| F : A leadsTo (B Un A'); F : (B-A') leadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   420
\   ==> F : A leadsTo (B' Un A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   421
by (rtac leadsTo_cancel1 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   422
by (assume_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   423
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   424
        addIs [leadsTo_weaken_R]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   425
        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   426
qed "leadsTo_cancel_Diff1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   427
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   428
(** The impossibility law **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   429
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   430
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   431
   "F : A leadsTo 0 ==> A=0";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   432
by (etac leadsTo_induct_pre 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   433
by (rewrite_goals_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   434
     [ensures_def, constrains_def, transient_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   435
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   436
by (auto_tac (claset() addSDs [Acts_type],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   437
              simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   438
               [actionSet_def, condition_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   439
by (blast_tac (claset() addSDs [bspec]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   440
qed "leadsTo_empty";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   441
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   442
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   443
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   444
(** PSP: Progress-Safety-Progress **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   445
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   446
(*Special case of PSP: Misra's "stable conjunction"*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   447
Goalw [stable_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   448
   "[| F : A leadsTo A'; F : stable(B) |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   449
\  F:(A Int B) leadsTo (A' Int B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   450
by (etac leadsTo_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   451
by (rtac leadsTo_Union_Int 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   452
by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   453
by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   454
by (rtac leadsTo_Basis 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   455
by (asm_full_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   456
         addsimps [ensures_def, Diff_Int_distrib RS sym, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   457
                   Diff_Int_distrib2 RS sym, Int_Un_distrib RS sym]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   458
by (REPEAT(blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   459
               addIs [transient_strengthen,constrains_Int]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   460
               addDs [constrainsD2]) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   461
qed "psp_stable";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   462
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   463
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   464
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   465
   "[| F : A leadsTo A'; F : stable(B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   466
\ ==> F : (B Int A) leadsTo (B Int A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   467
by (asm_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   468
             addsimps psp_stable::Int_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   469
qed "psp_stable2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   470
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   471
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   472
Goalw [ensures_def, constrains_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   473
   "[| F : A ensures A'; F : B co B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   474
\   ==> F : (A Int B') ensures ((A' Int B) Un (B' - B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   475
(*speeds up the proof*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   476
by (Clarify_tac 1);  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   477
by (blast_tac (claset() addIs [transient_strengthen]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   478
qed "psp_ensures";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   480
Goal "[| F : A leadsTo A'; F : B co B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   481
\     ==> F : (A Int B') leadsTo ((A' Int B) Un (B' - B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   482
by (subgoal_tac "F:program & A:condition & \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   483
                \ A':condition & B:condition & B':condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   484
by (blast_tac (claset() addDs [leadsToD, constrainsD2]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   485
by (etac leadsTo_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   486
by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   487
(*Transitivity case has a delicate argument involving "cancellation"*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   488
by (rtac leadsTo_Un_duplicate2 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   489
by (etac leadsTo_cancel_Diff1 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   490
by (asm_full_simp_tac (simpset() addsimps [Int_Diff, Diff_triv]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   491
by (blast_tac (claset() addIs [leadsTo_weaken_L] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   492
                        addDs [constrains_imp_subset]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   493
(*Basis case*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   494
by (blast_tac (claset() addIs [psp_ensures]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   495
qed "psp";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   496
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   497
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   498
Goal "[| F : A leadsTo A'; F : B co B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   499
\   ==> F : (B' Int A) leadsTo ((B Int A') Un (B' - B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   500
by (asm_simp_tac (simpset() addsimps psp::Int_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   501
qed "psp2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   502
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   503
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   504
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   505
Goalw [unless_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   506
   "[| F : A leadsTo A';  F : B unless B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   507
\   ==> F : (A Int B) leadsTo ((A' Int B) Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   508
by (subgoal_tac "F:program & A:condition & A':condition &\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   509
                \            B:condition & B':condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   510
by (blast_tac (claset() addDs [leadsToD, constrainsD2]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   511
by (dtac psp 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   512
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   513
by (blast_tac (claset() addIs [leadsTo_weaken]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   514
qed "psp_unless";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   515
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   516
(*** Proving the induction rules ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   517
(** The most general rule: r is any wf relation; f is any variant function **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   518
Goal "[| wf(r); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   519
\        m:I; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   520
\        field(r)<=I; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   521
\        F:program; B:condition;\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   522
\        ALL m:I. F : (A Int f-``{m}) leadsTo                     \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   523
\                   ((A Int f-``(converse(r)``{m})) Un B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   524
\     ==> F : (A Int f-``{m}) leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   525
by (eres_inst_tac [("a","m")] wf_induct2 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   526
by (ALLGOALS(Asm_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   527
by (subgoal_tac "F : (A Int (f-``(converse(r)``{x}))) leadsTo B" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   528
by (stac vimage_eq_UN 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   529
by (asm_simp_tac (simpset() addsimps [Int_UN_distrib]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   530
by (blast_tac (claset() addIs [leadsTo_cancel1, leadsTo_Un_duplicate]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   531
by (case_tac "converse(r)``{x}=0" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   532
by (auto_tac (claset() addSEs [not_emptyE] addSIs [leadsTo_UN], simpset()));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   533
qed "lemma1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   534
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   535
(** Meta or object quantifier ? **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   536
Goal "[| wf(r); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   537
\        field(r)<=I; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   538
\        A<=f-``I;\ 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   539
\        F:program; A:condition; B:condition; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   540
\        ALL m:I. F : (A Int f-``{m}) leadsTo                     \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   541
\                   ((A Int f-``(converse(r)``{m})) Un B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   542
\     ==> F : A leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   543
by (res_inst_tac [("b", "A")] subst 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   544
by (res_inst_tac [("I", "I")] leadsTo_UN 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   545
by (REPEAT (assume_tac 2));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   546
by (Clarify_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   547
by (eres_inst_tac [("I", "I")] lemma1 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   548
by (REPEAT (assume_tac 2));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   549
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   550
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   551
by (thin_tac "field(r)<=I" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   552
by (dres_inst_tac [("c", "x")] subsetD 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   553
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   554
by (res_inst_tac [("b", "x")] UN_I 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   555
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   556
qed "leadsTo_wf_induct";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   557
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   558
Goalw [field_def] "field(less_than(nat)) = nat";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   559
by (simp_tac (simpset() addsimps [less_than_equals]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   560
by (rtac equalityI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   561
by (force_tac (claset() addSEs [rangeE], simpset()) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   562
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   563
by (thin_tac "x~:range(?y)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   564
by (etac nat_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   565
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [domain_def])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   566
by (res_inst_tac [("x", "<succ(xa),succ(succ(xa))>")] ReplaceI 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   567
by (res_inst_tac [("x", "<0,1>")] ReplaceI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   568
by (REPEAT(force_tac (claset() addIs [splitI], simpset()) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   569
qed "nat_less_than_field";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   570
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   571
(*Alternative proof is via the lemma F : (A Int f-`(lessThan m)) leadsTo B*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   572
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   573
 "[| A<=f-``nat;\ 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   574
\    F:program; A:condition; B:condition; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   575
\    ALL m:nat. F:(A Int f-``{m}) leadsTo ((A Int f-``lessThan(m,nat)) Un B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   576
\     ==> F : A leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   577
by (res_inst_tac [("A1", "nat")]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   578
        (wf_less_than RS leadsTo_wf_induct) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   579
by (Clarify_tac 6);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   580
by (ALLGOALS(asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   581
          (simpset() addsimps [nat_less_than_field])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   582
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   583
by (ALLGOALS(asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   584
    (simpset() addsimps [rewrite_rule [vimage_def] Image_inverse_less_than])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   585
qed "lessThan_induct";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   586
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   587
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   588
(*** wlt ****)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   589
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   590
(*Misra's property W3*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   591
Goalw [wlt_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   592
"[| F:program; B:condition |] ==> F:wlt(F, B) leadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   593
by (blast_tac (claset() addSIs [leadsTo_Union]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   594
qed "wlt_leadsTo";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   595
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   596
Goalw [wlt_def] "F : A leadsTo B ==> A <= wlt(F, B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   597
by (blast_tac (claset() addSIs [leadsTo_Union]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   598
                        addDs  [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   599
qed "leadsTo_subset";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   600
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   601
(*Misra's property W2*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   602
Goal "[| F:program; B:condition |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   603
\ F : A leadsTo B <-> (A <= wlt(F,B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   604
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   605
   addSIs [leadsTo_subset, wlt_leadsTo RS leadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   606
qed "leadsTo_eq_subset_wlt";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   607
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   608
(*Misra's property W4*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   609
Goal "[| F:program; B:condition |] ==> B <= wlt(F,B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   610
by (asm_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   611
         addsimps [leadsTo_eq_subset_wlt RS iff_sym,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   612
                   subset_imp_leadsTo]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   613
qed "wlt_increasing";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   614
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   615
(*Used in the Trans case below*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   616
Goalw [constrains_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   617
   "[| B <= A2; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   618
\      F : (A1 - B) co (A1 Un B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   619
\      F : (A2 - C) co (A2 Un C) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   620
\   ==> F : (A1 Un A2 - C) co (A1 Un A2 Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   621
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   622
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   623
qed "lemma1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   624
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   625
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   626
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   627
(*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   628
(* slightly different from the HOL one: B here is bounded *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   629
Goal "F : A leadsTo A' \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   630
\     ==> EX B:condition. A<=B & F:B leadsTo A' & F : (B-A') co (B Un A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   631
by (forward_tac [leadsToD] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   632
by (etac leadsTo_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   633
(*Basis*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   634
by (blast_tac (claset() addDs [ensuresD, constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   635
(*Trans*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   636
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   637
by (res_inst_tac [("x", "Ba Un Bb")] bexI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   638
by (blast_tac (claset() addIs [lemma1,leadsTo_Un_Un, leadsTo_cancel1,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   639
                               leadsTo_Un_duplicate]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   640
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   641
(*Union*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   642
by (clarify_tac (claset() addSDs [ball_conj_distrib RS iffD1]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   643
by (subgoal_tac "EX y. y:Pi(S, %A. {Ba:condition. A<=Ba & \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   644
                          \         F:Ba leadsTo B & F:Ba - B co Ba Un B})" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   645
by (rtac AC_ball_Pi 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   646
by (Clarify_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   647
by (rotate_tac 3 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   648
by (blast_tac (claset() addSDs [bspec]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   649
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   650
by (res_inst_tac [("x", "UN A:S. y`A")] bexI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   651
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   652
by (res_inst_tac [("I1", "S")] (constrains_UN RS constrains_weaken) 3);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   653
by (rtac leadsTo_Union 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   654
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   655
by (asm_full_simp_tac (simpset() addsimps [condition_def]) 7);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   656
by (asm_full_simp_tac (simpset() addsimps [condition_def]) 6);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   657
by (REPEAT(blast_tac (claset() addDs [apply_type]) 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   658
qed "leadsTo_123";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   659
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   660
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   661
(*Misra's property W5*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   662
Goal "[| F:program; B:condition |] ==>F : (wlt(F, B) - B) co (wlt(F,B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   663
by (cut_inst_tac [("F","F")] (wlt_leadsTo RS leadsTo_123) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   664
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   665
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   666
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   667
by (subgoal_tac "Ba = wlt(F,B)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   668
by (blast_tac (claset() addDs [leadsTo_eq_subset_wlt RS iffD1]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   669
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   670
by (asm_full_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   671
         addsimps [wlt_increasing RS (subset_Un_iff2 RS iffD1)]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   672
qed "wlt_constrains_wlt";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   673
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   674
Goalw [wlt_def, condition_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   675
   "wlt(F,B):condition";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   676
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   677
qed "wlt_in_condition";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   678
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   679
(*** Completion: Binary and General Finite versions ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   680
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   681
Goal "[| W = wlt(F, (B' Un C));     \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   682
\      F : A leadsTo (A' Un C);  F : A' co (A' Un C);   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   683
\      F : B leadsTo (B' Un C);  F : B' co (B' Un C) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   684
\   ==> F : (A Int B) leadsTo ((A' Int B') Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   685
by (subgoal_tac "W:condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   686
by (blast_tac (claset() addIs [wlt_in_condition]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   687
by (subgoal_tac "F : (W-C) co (W Un B' Un C)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   688
by (blast_tac (claset() addIs [[asm_rl, wlt_constrains_wlt] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   689
                               MRS constrains_Un RS constrains_weaken]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   690
                        addDs [leadsToD, constrainsD2]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   691
by (subgoal_tac "F : (W-C) co W" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   692
by (subgoals_tac ["F:program", "(B' Un C):condition"] 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   693
by (rotate_tac ~2 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   694
by (asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   695
    (simpset() addsimps 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   696
                [wlt_increasing RS (subset_Un_iff2 RS iffD1), Un_assoc]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   697
by (REPEAT(blast_tac (claset()  addDs [leadsToD, constrainsD]) 2));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   698
by (subgoal_tac "F : (A Int W - C) leadsTo (A' Int W Un C)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   699
by (blast_tac (claset() addIs [wlt_leadsTo, psp RS leadsTo_weaken]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   700
                        addDs [leadsToD, constrainsD2]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   701
(** LEVEL 6 **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   702
by (subgoal_tac "F : (A' Int W Un C) leadsTo (A' Int B' Un C)" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   703
by (subgoal_tac "A' Int W Un C:condition & A' Int B' Un C:condition" 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   704
by (rtac leadsTo_Un_duplicate2 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   705
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   706
        addIs  [leadsTo_Un_Un, wlt_leadsTo RS 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   707
                psp2 RS leadsTo_weaken,leadsTo_refl] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   708
        addDs [leadsToD, constrainsD]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   709
by (thin_tac "W=wlt(F, B' Un C)" 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   710
by (blast_tac (claset() addDs [leadsToD, constrainsD2]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   711
by (dtac leadsTo_Diff 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   712
by (blast_tac (claset() addIs [subset_imp_leadsTo] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   713
                        addDs [leadsToD, constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   714
by (subgoal_tac "A Int B <= A Int W" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   715
by (blast_tac (claset() addSDs [leadsTo_subset]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   716
                        addSIs [subset_refl RS Int_mono]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   717
(** To speed the proof **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   718
by (subgoal_tac "A Int B :condition & A \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   719
               \ Int W :condition & A' Int B' Un C:condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   720
by (blast_tac (claset() addIs [leadsTo_Trans, subset_imp_leadsTo] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   721
                        addDs [leadsToD, constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   722
by (blast_tac (claset() addDs [leadsToD, constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   723
bind_thm("completion", refl RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   724
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   725
Goal "[| I:Fin(X); F:program; C:condition |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   726
\(ALL i:I. F : (A(i)) leadsTo (A'(i) Un C)) -->  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   727
\                  (ALL i:I. F : (A'(i)) co (A'(i) Un C)) --> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   728
\                  F : (INT i:I. A(i)) leadsTo ((INT i:I. A'(i)) Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   729
by (etac Fin_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   730
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   731
by (case_tac "y=0" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   732
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   733
by (etac not_emptyE 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   734
by (subgoal_tac "Inter(cons(A(x), RepFun(y, A)))= A(x) Int Inter(RepFun(y,A)) & \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   735
               \ Inter(cons(A'(x), RepFun(y, A')))= A'(x) Int Inter(RepFun(y,A'))" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   736
by (Blast_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   737
by (Asm_simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   738
by (rtac completion 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   739
by (subgoal_tac "Inter(RepFun(y, A')) Un C = (INT x:RepFun(y, A'). x Un C)" 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   740
by (Blast_tac 5);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   741
by (Asm_simp_tac 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   742
by (rtac constrains_INT 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   743
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   744
qed "lemma";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   745
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   746
val prems = Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   747
     "[| I:Fin(X);  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   748
\        !!i. i:I ==> F : A(i) leadsTo (A'(i) Un C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   749
\        !!i. i:I ==> F : A'(i) co (A'(i) Un C); F:program; C:condition |]   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   750
\     ==> F : (INT i:I. A(i)) leadsTo ((INT i:I. A'(i)) Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   751
by (blast_tac (claset() addIs (lemma RS mp RS mp)::prems) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   752
qed "finite_completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   753
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   754
Goalw [stable_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   755
     "[| F : A leadsTo A';  F : stable(A');   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   756
\        F : B leadsTo B';  F : stable(B') |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   757
\   ==> F : (A Int B) leadsTo (A' Int B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   758
by (res_inst_tac [("C1", "0")] (completion RS leadsTo_weaken_R) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   759
by (REPEAT(blast_tac (claset() addDs [leadsToD, constrainsD2]) 5));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   760
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   761
qed "stable_completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   762
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   763
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   764
val prems = Goalw [stable_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   765
     "[| I:Fin(X); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   766
\        ALL i:I. F : A(i) leadsTo A'(i); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   767
\        ALL i:I.  F: stable(A'(i));  F:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   768
\     ==> F : (INT i:I. A(i)) leadsTo (INT i:I. A'(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   769
by (subgoal_tac "(INT i:I. A'(i)):condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   770
by (blast_tac (claset() addDs [leadsToD, constrainsD2]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   771
by (res_inst_tac [("C1", "0")] (finite_completion RS leadsTo_weaken_R) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   772
by (assume_tac 7);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   773
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   774
by (ALLGOALS (Blast_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   775
qed "finite_stable_completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   776