src/ZF/UNITY/SubstAx.ML
author wenzelm
Wed, 15 Aug 2001 22:20:30 +0200
changeset 11501 3b6415035d1a
parent 11479 697dcaaf478f
child 12195 ed2893765a08
permissions -rw-r--r--
support for absolute namespace entry paths;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11479
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      ZF/UNITY/SubstAx.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
LeadsTo relation, restricted to the set of reachable states.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     7
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
     8
*)
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
(*Resembles the previous definition of LeadsTo*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    12
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    13
Goalw [LeadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    14
     "A LeadsTo B = \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    15
\ {F:program. F : (reachable(F) Int A) leadsTo (reachable(F) Int B) & \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    16
\    A:condition & B:condition}";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    17
by (blast_tac (claset() addDs [psp_stable2, leadsToD, constrainsD2] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    18
                        addIs [leadsTo_weaken]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    19
qed "LeadsTo_eq_leadsTo";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    20
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    21
Goalw [LeadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    22
"F: A LeadsTo B ==> F:program & A:condition & B:condition";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    23
by (Blast_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    24
qed "LeadsToD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    25
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    26
(*** Specialized laws for handling invariants ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    27
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    28
(** Conjoining an Always property **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    29
Goal "[| F : Always(INV); A:condition |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    30
\  (F : (INV Int A) LeadsTo A') <-> (F : A LeadsTo A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    31
by (asm_full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    32
    (simpset() addsimps [LeadsTo_def, Always_eq_includes_reachable,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    33
              Int_absorb2, Int_assoc RS sym, leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    34
qed "Always_LeadsTo_pre";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    35
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    36
Goal "[| F : Always(INV); A':condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    37
  \ ==> (F : A LeadsTo (INV Int A')) <-> (F : A LeadsTo A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    38
by (asm_full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    39
    (simpset() addsimps [LeadsTo_eq_leadsTo, Always_eq_includes_reachable, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    40
          Int_absorb2, Int_assoc RS sym,leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    41
qed "Always_LeadsTo_post";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    42
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    43
(* Like 'Always_LeadsTo_pre RS iffD1', but with premises in the good order *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    44
Goal "[| F:Always(C); F : (C Int A) LeadsTo A'; A:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    45
\ ==> F: A LeadsTo A'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    46
by (blast_tac (claset() addIs [Always_LeadsTo_pre RS iffD1]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    47
qed "Always_LeadsToI";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    48
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    49
(* Like 'Always_LeadsTo_post RS iffD2', but with premises in the good order *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    50
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    51
"[| F : Always(C);  F : A LeadsTo A'; A':condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    52
\  ==> F : A LeadsTo (C Int A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    53
by (blast_tac (claset() addIs [Always_LeadsTo_post RS iffD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    54
qed "Always_LeadsToD";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    55
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    56
(*** Introduction rules: Basis, Trans, Union ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    57
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    58
Goal "F : A leadsTo B ==> F : A LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    59
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    60
by (blast_tac (claset() addIs [leadsTo_weaken_L]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    61
                        addDs [leadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    62
qed "leadsTo_imp_LeadsTo";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    63
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    64
Goal "[| F : A LeadsTo B;  F : B LeadsTo C |] ==> F : A LeadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    65
by (full_simp_tac (simpset() addsimps [LeadsTo_eq_leadsTo]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    66
by (blast_tac (claset() addIs [leadsTo_Trans]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    67
qed "LeadsTo_Trans";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    68
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    69
Goalw [LeadsTo_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    70
     "[| ALL A:S. F : A LeadsTo B; F:program; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    71
\ ==> F : Union(S) LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    72
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    73
by (stac Int_Union_Union2 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    74
by (blast_tac (claset() addIs [leadsTo_UN]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    75
bind_thm("LeadsTo_Union", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    76
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    77
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    78
(*** Derived rules ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    79
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    80
Goalw [LeadsTo_def] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    81
"[| F:program; A:condition |] ==>F : A LeadsTo state";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    82
by (blast_tac (claset() addIs [leadsTo_state]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    83
qed "LeadsTo_state";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    84
Addsimps [LeadsTo_state];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    85
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    86
(*Useful with cancellation, disjunction*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    87
Goal "F : A LeadsTo (A' Un A') ==> F : A LeadsTo A'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    88
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    89
qed "LeadsTo_Un_duplicate";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    90
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    91
Goal "F : A LeadsTo (A' Un C Un C) ==> F : A LeadsTo (A' Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    92
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    93
qed "LeadsTo_Un_duplicate2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    94
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    95
Goal "[| ALL i:I. F : A(i) LeadsTo B; F:program; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    96
\  ==> F : (UN i:I. A(i)) LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    97
by (simp_tac (simpset() addsimps [Int_Union_Union]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    98
by (blast_tac (claset() addIs [LeadsTo_Union]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
    99
bind_thm("LeadsTo_UN", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   100
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   101
(*Binary union introduction rule*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   102
Goal "[| F : A LeadsTo C; F : B LeadsTo C |] ==> F : (A Un B) LeadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   103
by (stac Un_eq_Union 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   104
by (blast_tac (claset() addIs [LeadsTo_Union] 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   105
                        addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   106
qed "LeadsTo_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   107
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   108
(*Lets us look at the starting state*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   109
Goal "[| ALL s:A. F : {s} LeadsTo B; F:program; B:condition |]\
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   110
\  ==> F : A LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   111
by (stac (UN_singleton RS sym) 1 THEN rtac LeadsTo_UN 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   112
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   113
bind_thm("single_LeadsTo_I", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   114
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   115
Goal "[| A <= B; B:condition; F:program |] ==> F : A LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   116
by (subgoal_tac "A:condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   117
by (force_tac (claset(), 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   118
         simpset() addsimps [condition_def]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   119
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   120
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   121
qed "subset_imp_LeadsTo";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   122
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   123
bind_thm ("empty_LeadsTo", empty_subsetI RS subset_imp_LeadsTo);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   124
Addsimps [empty_LeadsTo];
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   125
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   126
Goal "[| F : A LeadsTo A';  A' <= B'; B':condition |] ==> F : A LeadsTo B'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   127
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   128
by (blast_tac (claset() addIs [leadsTo_weaken_R]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   129
qed_spec_mp "LeadsTo_weaken_R";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   130
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   131
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   132
Goal "[| F : A LeadsTo A';  B <= A |]  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   133
\     ==> F : B LeadsTo A'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   134
by (subgoal_tac "B:condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   135
by (force_tac (claset() addSDs [LeadsToD],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   136
               simpset() addsimps [condition_def]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   137
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   138
by (blast_tac (claset() addIs [leadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   139
qed_spec_mp "LeadsTo_weaken_L";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   140
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   141
Goal "[| F : A LeadsTo A';   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   142
\        B  <= A;   A' <= B'; B':condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   143
\     ==> F : B LeadsTo B'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   144
by (blast_tac (claset() addIs [LeadsTo_weaken_R, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   145
                    LeadsTo_weaken_L, LeadsTo_Trans]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   146
qed "LeadsTo_weaken";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   147
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   148
Goal "[| F : Always(C);  F : A LeadsTo A';   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   149
\        C Int B <= A;   C Int A' <= B'; B:condition; B':condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   150
\     ==> F : B LeadsTo B'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   151
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   152
      addDs [AlwaysD2, LeadsToD, Always_LeadsToI]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   153
      addIs [LeadsTo_weaken, Always_LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   154
qed "Always_LeadsTo_weaken";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   155
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   156
(** Two theorems for "proof lattices" **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   157
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   158
Goal "F : A LeadsTo B ==> F:(A Un B) LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   159
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   160
         addIs [LeadsTo_Un, subset_imp_LeadsTo]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   161
         addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   162
qed "LeadsTo_Un_post";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   163
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   164
Goal "[| F : A LeadsTo B;  F : B LeadsTo C |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   165
\     ==> F : (A Un B) LeadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   166
by (blast_tac (claset() addIs [LeadsTo_Un, subset_imp_LeadsTo, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   167
                               LeadsTo_weaken_L, LeadsTo_Trans]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   168
                        addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   169
qed "LeadsTo_Trans_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   170
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   171
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   172
(** Distributive laws **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   173
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   174
Goal "(F : (A Un B) LeadsTo C)  <-> (F : A LeadsTo C & F : B LeadsTo C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   175
by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   176
qed "LeadsTo_Un_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   177
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   178
Goal "[| F:program; B:condition |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   179
\ (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
   180
by (blast_tac (claset() addIs [LeadsTo_UN, LeadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   181
qed "LeadsTo_UN_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   182
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   183
Goal "[| F:program; B:condition |] ==> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   184
\ (F : Union(S) LeadsTo B)  <->  (ALL A : S. F : A LeadsTo B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   185
by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_L]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   186
qed "LeadsTo_Union_distrib";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   187
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   188
(** More rules using the premise "Always INV" **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   189
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   190
Goal "F : A Ensures B ==> F : A LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   191
by (asm_full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   192
    (simpset() addsimps [Ensures_def, LeadsTo_def, leadsTo_Basis]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   193
qed "LeadsTo_Basis";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   194
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   195
Goal "[| F : (A-B) Co (A Un B);  F : transient (A-B) |]   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   196
\     ==> F : A Ensures B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   197
by (asm_full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   198
    (simpset() addsimps [Ensures_def, Constrains_eq_constrains]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   199
by (blast_tac (claset() addIs [ensuresI, constrains_weaken, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   200
                               transient_strengthen]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   201
                        addDs [constrainsD2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   202
qed "EnsuresI";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   203
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   204
Goal "[| F : Always(INV);      \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   205
\        F : (INV Int (A-A')) Co (A Un A'); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   206
\        F : transient (INV Int (A-A')) |]   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   207
\ ==> F : A LeadsTo A'";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   208
by (rtac Always_LeadsToI 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   209
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   210
by (blast_tac (claset() addDs [ConstrainsD]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   211
by (blast_tac (claset() addIs [EnsuresI, LeadsTo_Basis,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   212
                               Always_ConstrainsD RS Constrains_weaken, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   213
                               transient_strengthen]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   214
                        addDs [AlwaysD2, ConstrainsD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   215
qed "Always_LeadsTo_Basis";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   216
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   217
(*Set difference: maybe combine with leadsTo_weaken_L??
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   218
  This is the most useful form of the "disjunction" rule*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   219
Goal "[| F : (A-B) LeadsTo C;  F : (A Int B) LeadsTo C; \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   220
\ A:condition; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   221
\     ==> F : A LeadsTo C";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   222
by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   223
                addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   224
qed "LeadsTo_Diff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   225
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   226
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   227
Goal "[| ALL i:I. F: A(i) LeadsTo A'(i); F:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   228
\     ==> F : (UN i:I. A(i)) LeadsTo (UN i:I. A'(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   229
by (rtac LeadsTo_Union 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   230
by (ALLGOALS(Clarify_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   231
by (blast_tac (claset() addDs [LeadsToD]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   232
by (blast_tac (claset()  addIs [LeadsTo_weaken_R]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   233
                         addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   234
bind_thm ("LeadsTo_UN_UN", ballI RS result());
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   235
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   236
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   237
(*Version with no index set*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   238
Goal "[| ALL i. F: A(i) LeadsTo A'(i); F:program |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   239
\     ==> F : (UN i:I. A(i)) LeadsTo (UN i:I. A'(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   240
by (blast_tac (claset() addIs [LeadsTo_UN_UN]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   241
qed "all_LeadsTo_UN_UN";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   242
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   243
bind_thm ("LeadsTo_UN_UN_noindex", allI RS all_LeadsTo_UN_UN);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   244
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   245
(*Binary union version*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   246
Goal "[| F : A LeadsTo A'; F : B LeadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   247
\           ==> F : (A Un B) LeadsTo (A' Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   248
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   249
        addIs [LeadsTo_Un, LeadsTo_weaken_R]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   250
        addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   251
qed "LeadsTo_Un_Un";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   252
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   253
(** The cancellation law **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   254
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   255
Goal "[| F : A LeadsTo (A' Un B); F : B LeadsTo B' |]    \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   256
\     ==> F : A LeadsTo (A' Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   257
by (blast_tac (claset() addIs [LeadsTo_Un_Un, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   258
                               subset_imp_LeadsTo, LeadsTo_Trans]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   259
                    addDs [LeadsToD]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   260
qed "LeadsTo_cancel2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   261
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   262
Goal "A Un (B - A) = A Un B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   263
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   264
qed "Un_Diff";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   265
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   266
Goal "[| F : A LeadsTo (A' Un B); F : (B-A') LeadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   267
\     ==> F : A LeadsTo (A' Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   268
by (rtac LeadsTo_cancel2 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   269
by (assume_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   270
by (asm_simp_tac (simpset() addsimps [Un_Diff]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   271
qed "LeadsTo_cancel_Diff2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   272
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   273
Goal "[| F : A LeadsTo (B Un A'); F : B LeadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   274
\     ==> F : A LeadsTo (B' Un A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   275
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   276
by (blast_tac (claset() addSIs [LeadsTo_cancel2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   277
qed "LeadsTo_cancel1";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   278
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   279
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   280
Goal "(B - A) Un A = B Un A";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   281
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   282
qed "Diff_Un2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   283
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   284
Goal "[| F : A LeadsTo (B Un A'); F : (B-A') LeadsTo B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   285
\     ==> F : A LeadsTo (B' Un A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   286
by (rtac LeadsTo_cancel1 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   287
by (assume_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   288
by (asm_simp_tac (simpset() addsimps [Diff_Un2]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   289
qed "LeadsTo_cancel_Diff1";
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 impossibility law **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   293
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   294
(*The set "A" may be non-empty, but it contains no reachable states*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   295
Goal "F : A LeadsTo 0 ==> F : Always (state -A)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   296
by (full_simp_tac (simpset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   297
           addsimps [LeadsTo_def,Always_eq_includes_reachable]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   298
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   299
by (forward_tac [reachableD] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   300
by (auto_tac (claset() addSDs [leadsTo_empty],
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   301
              simpset() addsimps [condition_def]));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   302
qed "LeadsTo_empty";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   303
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   304
(** PSP: Progress-Safety-Progress **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   305
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   306
(*Special case of PSP: Misra's "stable conjunction"*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   307
Goal "[| F : A LeadsTo A';  F : Stable(B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   308
\     ==> F : (A Int B) LeadsTo (A' Int B)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   309
by (forward_tac [StableD2] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   310
by (rotate_tac ~1 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   311
by (asm_full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   312
    (simpset() addsimps [LeadsTo_eq_leadsTo, Stable_eq_stable]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   313
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   314
by (dtac psp_stable 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   315
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   316
by (asm_full_simp_tac (simpset() addsimps (Int_absorb::Int_ac)) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   317
qed "PSP_Stable";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   318
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   319
Goal "[| F : A LeadsTo A'; F : Stable(B) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   320
\     ==> F : (B Int A) LeadsTo (B Int A')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   321
by (asm_simp_tac (simpset() addsimps PSP_Stable::Int_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   322
qed "PSP_Stable2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   323
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   324
Goal "[| F : A LeadsTo A'; F : B Co B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   325
\     ==> F : (A Int B') LeadsTo ((A' Int B) Un (B' - B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   326
by (full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   327
    (simpset() addsimps [LeadsTo_def, Constrains_eq_constrains]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   328
by (blast_tac (claset() addDs [psp] addIs [leadsTo_weaken]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   329
qed "PSP";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   330
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   331
Goal "[| F : A LeadsTo A'; F : B Co B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   332
\     ==> F : (B' Int A) LeadsTo ((B Int A') Un (B' - B))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   333
by (asm_simp_tac (simpset() addsimps PSP::Int_ac) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   334
qed "PSP2";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   335
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   336
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   337
Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   338
"[| F : A LeadsTo A'; F : B Unless B' |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   339
\     ==> F : (A Int B) LeadsTo ((A' Int B) Un B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   340
by (forward_tac [LeadsToD] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   341
by (forward_tac [UnlessD] 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   342
by (rewrite_goals_tac [Unless_def]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   343
by (dtac PSP 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   344
by (assume_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   345
by (blast_tac (claset() 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   346
        addIs [LeadsTo_Diff, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   347
               LeadsTo_weaken, subset_imp_LeadsTo]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   348
qed "PSP_Unless";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   349
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   350
(*** Induction rules ***)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   351
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   352
(** Meta or object quantifier ????? **)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   353
Goal "[| wf(r);     \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   354
\        ALL m:I. F : (A Int f-``{m}) LeadsTo                     \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   355
\                           ((A Int f-``(converse(r) `` {m})) Un B); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   356
\        field(r)<=I; A<=f-``I; F:program; A:condition; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   357
\     ==> F : A LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   358
by (full_simp_tac (simpset() addsimps [LeadsTo_eq_leadsTo]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   359
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   360
by (eres_inst_tac [("I", "I"), ("f", "f")] leadsTo_wf_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   361
by (ALLGOALS(Clarify_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   362
by (dres_inst_tac [("x", "m")] bspec 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   363
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   364
by (res_inst_tac [("A'", 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   365
           "reachable(F) Int (A Int f -``(converse(r)``{m}) Un B)")]  
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   366
        leadsTo_weaken_R 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   367
by (asm_simp_tac (simpset() addsimps [Int_assoc]) 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   368
by (asm_simp_tac (simpset() addsimps [Int_assoc]) 5);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   369
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   370
qed "LeadsTo_wf_induct";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   371
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   372
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   373
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   374
Goal "[| 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
   375
\     A<=f-``nat; F:program; A:condition; B:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   376
\     ==> F : A LeadsTo B";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   377
by (res_inst_tac [("A1", "nat"), ("I", "nat")] (wf_less_than RS LeadsTo_wf_induct) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   378
by (ALLGOALS(asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   379
          (simpset() addsimps [nat_less_than_field])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   380
by (Clarify_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   381
by (ALLGOALS(asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   382
    (simpset() addsimps [rewrite_rule [vimage_def] Image_inverse_less_than])));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   383
qed "LessThan_induct";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   384
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   385
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   386
(****** 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   387
 To be ported ??? I am not sure.
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   388
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   389
  integ_0_le_induct
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   390
  LessThan_bounded_induct
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   391
  GreaterThan_bounded_induct
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   392
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
(*** Completion: Binary and General Finite versions ***)
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 C);  F : A' Co (A' Un C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   398
\        F : B LeadsTo (B' Un C);  F : B' Co (B' Un C) |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   399
\     ==> F : (A Int B) LeadsTo ((A' Int B') Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   400
by (full_simp_tac
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   401
    (simpset() addsimps [LeadsTo_eq_leadsTo, Constrains_eq_constrains, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   402
                         Int_Un_distrib2 RS sym]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   403
by Safe_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   404
by (REPEAT(Blast_tac 2));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   405
by (blast_tac (claset() addIs [completion, leadsTo_weaken]) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   406
qed "Completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   407
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   408
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   409
Goal "[| I:Fin(X);F:program; C:condition |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   410
\     ==> (ALL i:I. F : (A(i)) LeadsTo (A'(i) Un C)) -->  \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   411
\         (ALL i:I. F : (A'(i)) Co (A'(i) Un C)) --> \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   412
\         F : (INT i:I. A(i)) LeadsTo ((INT i:I. A'(i)) Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   413
by (etac Fin_induct 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   414
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   415
by (case_tac "y=0" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   416
by Auto_tac;
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   417
by (etac not_emptyE 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   418
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
   419
               \ Inter(cons(A'(x), RepFun(y, A')))= A'(x) Int Inter(RepFun(y,A'))" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   420
by (Blast_tac 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   421
by (Asm_simp_tac 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   422
by (rtac Completion 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   423
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
   424
by (Blast_tac 5);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   425
by (Asm_simp_tac 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   426
by (rtac Constrains_INT 4);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   427
by (REPEAT(Blast_tac 1));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   428
val lemma = result();
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   429
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   430
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   431
val prems = Goal
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   432
     "[| I:Fin(X); !!i. i:I ==> F : A(i) LeadsTo (A'(i) Un C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   433
\        !!i. i:I ==> F : A'(i) Co (A'(i) Un C); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   434
\        F:program; C:condition |]   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   435
\     ==> F : (INT i:I. A(i)) LeadsTo ((INT i:I. A'(i)) Un C)";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   436
by (blast_tac (claset() addIs (lemma RS mp RS mp)::prems) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   437
qed "Finite_completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   438
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   439
Goalw [Stable_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   440
     "[| F : A LeadsTo A';  F : Stable(A');   \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   441
\        F : B LeadsTo B';  F : Stable(B') |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   442
\   ==> F : (A Int B) LeadsTo (A' Int B')";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   443
by (res_inst_tac [("C1", "0")] (Completion RS LeadsTo_weaken_R) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   444
by (REPEAT(blast_tac (claset() addDs [LeadsToD,ConstrainsD]) 5));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   445
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   446
qed "Stable_completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   447
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   448
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   449
val prems = Goalw [Stable_def]
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   450
     "[| I:Fin(X); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   451
\        ALL i:I. F : A(i) LeadsTo A'(i); \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   452
\        ALL i:I.  F: Stable(A'(i));   F:program  |] \
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   453
\     ==> F : (INT i:I. A(i)) LeadsTo (INT i:I. A'(i))";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   454
by (subgoal_tac "(INT i:I. A'(i)):condition" 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   455
by (blast_tac (claset() addDs  [LeadsToD,ConstrainsD]) 2);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   456
by (res_inst_tac [("C1", "0")] (Finite_completion RS LeadsTo_weaken_R) 1);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   457
by (assume_tac 7);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   458
by (ALLGOALS(Asm_full_simp_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   459
by (ALLGOALS (Blast_tac));
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   460
qed "Finite_stable_completion";
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   461
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   462
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   463
(*proves "ensures/leadsTo" properties when the program is specified*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   464
fun ensures_tac sact = 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   465
    SELECT_GOAL
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   466
      (EVERY [REPEAT (Always_Int_tac 1),
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   467
              etac Always_LeadsTo_Basis 1 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   468
                  ORELSE   (*subgoal may involve LeadsTo, leadsTo or ensures*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   469
                  REPEAT (ares_tac [LeadsTo_Basis, leadsTo_Basis,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   470
                                    EnsuresI, ensuresI] 1),
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   471
              (*now there are two subgoals: co & transient*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   472
              simp_tac (simpset() addsimps !program_defs_ref) 2,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   473
              res_inst_tac [("act", sact)] transientI 2,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   474
                 (*simplify the command's domain*)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   475
              simp_tac (simpset() addsimps [domain_def]) 3, 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   476
              (* proving the domain part *)
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   477
             Clarify_tac 3, dtac swap 3, Force_tac 4,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   478
             rtac ReplaceI 3, Force_tac 3, Force_tac 4,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   479
             Asm_full_simp_tac 3, rtac conjI 3, Simp_tac 4,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   480
             REPEAT (rtac update_type2 3),
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   481
             constrains_tac 1,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   482
             ALLGOALS Clarify_tac,
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   483
             ALLGOALS (asm_full_simp_tac 
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   484
            (simpset() addsimps [condition_def])),
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   485
            ALLGOALS Clarify_tac]);
697dcaaf478f new ZF/UNITY theory
paulson
parents:
diff changeset
   486