src/HOL/UNITY/ELT.ML
author paulson
Wed, 01 Dec 1999 11:20:24 +0100
changeset 8044 296b03b79505
child 8055 bb15396278fb
permissions -rw-r--r--
new generalized leads-to theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/ELT
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     5
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     6
leadsTo strengthened with a specification of the allowable sets transient parts
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     7
*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     8
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
     9
Goalw [givenBy_def] "(givenBy v) = {A. ALL x:A. ALL y. v x = v y --> y: A}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    10
by Safe_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    11
by (res_inst_tac [("x", "v `` ?u")] image_eqI 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    12
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    13
qed "givenBy_eq_all";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    14
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    15
Goal "givenBy v = {A. EX P. A = {s. P(v s)}}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    16
by (simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    17
by Safe_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    18
by (res_inst_tac [("x", "%n. EX s. v s = n & s : ?A")] exI 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    19
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    20
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    21
qed "givenBy_eq_Collect";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    22
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    23
val prems =
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    24
Goal "(!!x y. [|  x:A;  v x = v y |] ==> y: A) ==> A: givenBy v";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    25
by (stac givenBy_eq_all 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    26
by (blast_tac (claset() addIs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    27
qed "givenByI";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    28
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    29
Goalw [givenBy_def] "[| A: givenBy v;  x:A;  v x = v y |] ==> y: A";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    30
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    31
qed "givenByD";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    32
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    33
Goal "{} : givenBy v";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    34
by (blast_tac (claset() addSIs [givenByI]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    35
qed "empty_mem_givenBy";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    36
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    37
AddIffs [empty_mem_givenBy];
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    38
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    39
Goal "A: givenBy v ==> EX P. A = {s. P(v s)}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    40
by (res_inst_tac [("x", "%n. EX s. v s = n & s : A")] exI 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    41
by (full_simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    42
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    43
qed "givenBy_imp_eq_Collect";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    44
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    45
Goalw [givenBy_def] "EX P. A = {s. P(v s)} ==> A: givenBy v";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    46
by (Best_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    47
qed "eq_Collect_imp_givenBy";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    48
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    49
Goal "givenBy v = {A. EX P. A = {s. P(v s)}}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    50
by (blast_tac (claset() addIs [eq_Collect_imp_givenBy,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    51
			       givenBy_imp_eq_Collect]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    52
qed "givenBy_eq_eq_Collect";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    53
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    54
Goal "(funPair f g) o h = funPair (f o h) (g o h)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    55
by (simp_tac (simpset() addsimps [funPair_def, o_def]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    56
qed "funPair_o_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    57
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    58
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    59
(** Standard leadsTo rules **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    60
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    61
Goalw [leadsETo_def] "[| F: A ensures B;  A-B: CC |] ==> F : A leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    62
by (blast_tac (claset() addIs [elt.Basis]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    63
qed "leadsETo_Basis";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    64
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    65
Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    66
     "[| F : A leadsTo[CC] B;  F : B leadsTo[CC] C |] ==> F : A leadsTo[CC] C";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    67
by (blast_tac (claset() addIs [elt.Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    68
qed "leadsETo_Trans";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    69
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    70
(*Useful with cancellation, disjunction*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    71
Goal "F : A leadsTo[CC] (A' Un A') ==> F : A leadsTo[CC] A'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    72
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    73
qed "leadsETo_Un_duplicate";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    74
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    75
Goal "F : A leadsTo[CC] (A' Un C Un C) ==> F : A leadsTo[CC] (A' Un C)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    76
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    77
qed "leadsETo_Un_duplicate2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    78
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    79
(*The Union introduction rule as we should have liked to state it*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    80
val prems = Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    81
    "(!!A. A : S ==> F : A leadsTo[CC] B) ==> F : (Union S) leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    82
by (blast_tac (claset() addIs [elt.Union] addDs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    83
qed "leadsETo_Union";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    84
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    85
val prems = Goal
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    86
    "(!!i. i : I ==> F : (A i) leadsTo[CC] B) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    87
\    ==> F : (UN i:I. A i) leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    88
by (stac (Union_image_eq RS sym) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    89
by (blast_tac (claset() addIs leadsETo_Union::prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    90
qed "leadsETo_UN";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    91
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    92
(*The INDUCTION rule as we should have liked to state it*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    93
val major::prems = Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    94
  "[| F : za leadsTo[CC] zb;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    95
\     !!A B. [| F : A ensures B;  A-B : CC |] ==> P A B; \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    96
\     !!A B C. [| F : A leadsTo[CC] B; P A B; F : B leadsTo[CC] C; P B C |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    97
\              ==> P A C; \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    98
\     !!B S. ALL A:S. F : A leadsTo[CC] B & P A B ==> P (Union S) B \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    99
\  |] ==> P za zb";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   100
by (rtac (major RS CollectD RS elt.induct) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   101
by (REPEAT (blast_tac (claset() addIs prems) 1));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   102
qed "leadsETo_induct";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   103
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   104
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   105
(** New facts involving leadsETo **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   106
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   107
Goal "CC' <= CC ==> (A leadsTo[CC'] B) <= (A leadsTo[CC] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   108
by Safe_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   109
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   110
by (blast_tac (claset() addIs [leadsETo_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   111
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   112
by (blast_tac (claset() addIs [leadsETo_Basis]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   113
qed "leadsETo_mono";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   114
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   115
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   116
val prems = Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   117
 "(!!A. A : S ==> F : (A Int C) leadsTo[CC] B) ==> F : (Union S Int C) leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   118
by (simp_tac (HOL_ss addsimps [Int_Union_Union]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   119
by (blast_tac (claset() addIs [elt.Union] addDs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   120
qed "leadsETo_Union_Int";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   121
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   122
(*Binary union introduction rule*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   123
Goal "[| F : A leadsTo[CC] C; F : B leadsTo[CC] C |] ==> F : (A Un B) leadsTo[CC] C";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   124
by (stac Un_eq_Union 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   125
by (blast_tac (claset() addIs [leadsETo_Union]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   126
qed "leadsETo_Un";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   127
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   128
val prems = 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   129
Goal "(!!x. x : A ==> F : {x} leadsTo[CC] B) ==> F : A leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   130
by (stac (UN_singleton RS sym) 1 THEN rtac leadsETo_UN 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   131
by (blast_tac (claset() addIs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   132
qed "single_leadsETo_I";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   133
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   134
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   135
Goal "[| A<=B;  {}:CC |]  ==> F : A leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   136
by (asm_simp_tac (simpset() addsimps [subset_imp_ensures RS leadsETo_Basis,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   137
				      Diff_eq_empty_iff RS iffD2]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   138
qed "subset_imp_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   139
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   140
bind_thm ("empty_leadsETo", empty_subsetI RS subset_imp_leadsETo);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   141
Addsimps [empty_leadsETo];
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   142
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   143
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   144
(** Weakening laws all require {}:CC **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   145
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   146
Goal "[| F : A leadsTo[CC] A';  A'<=B';  {}:CC |] ==> F : A leadsTo[CC] B'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   147
by (blast_tac (claset() addIs [subset_imp_leadsETo, leadsETo_Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   148
qed "leadsETo_weaken_R";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   149
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   150
Goal "[| F : A leadsTo[CC] A'; B<=A;  {}:CC |] ==> F : B leadsTo[CC] A'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   151
by (blast_tac (claset() addIs [leadsETo_Trans, subset_imp_leadsETo]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   152
qed_spec_mp "leadsETo_weaken_L";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   153
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   154
(*Distributes over binary unions*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   155
Goal "{} : CC ==> \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   156
\ F : (A Un B) leadsTo[CC] C  =  (F : A leadsTo[CC] C & F : B leadsTo[CC] C)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   157
by (blast_tac (claset() addIs [leadsETo_Un, leadsETo_weaken_L]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   158
qed "leadsETo_Un_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   159
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   160
Goal "{} : CC ==> \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   161
\   F : (UN i:I. A i) leadsTo[CC] B  =  (ALL i : I. F : (A i) leadsTo[CC] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   162
by (blast_tac (claset() addIs [leadsETo_UN, leadsETo_weaken_L]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   163
qed "leadsETo_UN_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   164
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   165
Goal "{} : CC \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   166
\     ==> F : (Union S) leadsTo[CC] B  =  (ALL A : S. F : A leadsTo[CC] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   167
by (blast_tac (claset() addIs [leadsETo_Union, leadsETo_weaken_L]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   168
qed "leadsETo_Union_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   169
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   170
Goal "[| F : A leadsTo[CC'] A'; B<=A; A'<=B';  CC' <= CC;  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   171
\     ==> F : B leadsTo[CC] B'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   172
by (dtac (impOfSubs leadsETo_mono) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   173
by (assume_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   174
by (blast_tac (claset() addIs [leadsETo_weaken_R, leadsETo_weaken_L,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   175
			       leadsETo_Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   176
qed "leadsETo_weaken";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   177
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   178
Goal "[| F : A leadsTo[CC] A';  CC <= givenBy v |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   179
\     ==> F : A leadsTo[givenBy v] A'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   180
by (blast_tac (claset() addIs [empty_mem_givenBy, leadsETo_weaken]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   181
qed "leadsETo_givenBy";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   182
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   183
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   184
(*Set difference*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   185
Goal "[| F : (A-B) leadsTo[CC] C; F : B leadsTo[CC] C;  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   186
\     ==> F : A leadsTo[CC] C";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   187
by (blast_tac (claset() addIs [leadsETo_Un, leadsETo_weaken]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   188
qed "leadsETo_Diff";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   189
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   190
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   191
(** Meta or object quantifier ???
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   192
    see ball_constrains_UN in UNITY.ML***)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   193
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   194
val prems = goal thy
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   195
   "[| !! i. i:I ==> F : (A i) leadsTo[CC] (A' i);  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   196
\   ==> F : (UN i:I. A i) leadsTo[CC] (UN i:I. A' i)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   197
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   198
by (blast_tac (claset() addIs [leadsETo_Union, leadsETo_weaken_R] 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   199
                        addIs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   200
qed "leadsETo_UN_UN";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   201
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   202
(*Binary union version*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   203
Goal "[| F : A leadsTo[CC] A';  F : B leadsTo[CC] B';  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   204
\     ==> F : (A Un B) leadsTo[CC] (A' Un B')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   205
by (blast_tac (claset() addIs [leadsETo_Un, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   206
			       leadsETo_weaken_R]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   207
qed "leadsETo_Un_Un";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   208
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   209
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   210
(** The cancellation law **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   211
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   212
Goal "[| F : A leadsTo[CC] (A' Un B); F : B leadsTo[CC] B';  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   213
\     ==> F : A leadsTo[CC] (A' Un B')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   214
by (blast_tac (claset() addIs [leadsETo_Un_Un, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   215
			       subset_imp_leadsETo, leadsETo_Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   216
qed "leadsETo_cancel2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   217
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   218
Goal "[| F : A leadsTo[CC] (A' Un B); F : (B-A') leadsTo[CC] B';  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   219
\     ==> F : A leadsTo[CC] (A' Un B')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   220
by (rtac leadsETo_cancel2 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   221
by (assume_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   222
by (ALLGOALS Asm_simp_tac);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   223
qed "leadsETo_cancel_Diff2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   224
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   225
Goal "[| F : A leadsTo[CC] (B Un A'); F : B leadsTo[CC] B';  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   226
\   ==> F : A leadsTo[CC] (B' Un A')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   227
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   228
by (blast_tac (claset() addSIs [leadsETo_cancel2]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   229
qed "leadsETo_cancel1";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   230
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   231
Goal "[| F : A leadsTo[CC] (B Un A'); F : (B-A') leadsTo[CC] B';  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   232
\   ==> F : A leadsTo[CC] (B' Un A')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   233
by (rtac leadsETo_cancel1 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   234
by (assume_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   235
by (ALLGOALS Asm_simp_tac);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   236
qed "leadsETo_cancel_Diff1";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   237
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   238
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   239
(** The impossibility law **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   240
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   241
Goal "F : A leadsTo[CC] B ==> B={} --> A={}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   242
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   243
by (ALLGOALS Asm_simp_tac);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   244
by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   245
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   246
val lemma = result() RS mp;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   247
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   248
Goal "F : A leadsTo[CC] {} ==> A={}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   249
by (blast_tac (claset() addSIs [lemma]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   250
qed "leadsETo_empty";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   251
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   252
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   253
(** PSP: Progress-Safety-Progress **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   254
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   255
(*Special case of PSP: Misra's "stable conjunction"*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   256
Goalw [stable_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   257
   "[| F : A leadsTo[CC] A';  F : stable B;  ALL C:CC. C Int B : CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   258
\   ==> F : (A Int B) leadsTo[CC] (A' Int B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   259
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   260
by (blast_tac (claset() addIs [leadsETo_Union_Int]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   261
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   262
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   263
by (asm_full_simp_tac
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   264
    (simpset() addsimps [ensures_def, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   265
			 Diff_Int_distrib2 RS sym, Int_Un_distrib2 RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   266
by (asm_simp_tac (simpset() addsimps [Diff_Int_distrib2 RS sym]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   267
by (blast_tac (claset() addIs [transient_strengthen, constrains_Int]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   268
qed "e_psp_stable";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   269
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   270
Goal "[| F : A leadsTo[CC] A'; F : stable B;  ALL C:CC. C Int B : CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   271
\     ==> F : (B Int A) leadsTo[CC] (B Int A')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   272
by (asm_simp_tac (simpset() addsimps e_psp_stable::Int_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   273
qed "e_psp_stable2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   274
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   275
Goal "[| F : A leadsTo[CC] A'; F : B co B';  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   276
\        ALL C:CC. C Int B Int B' : CC;  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   277
\     ==> F : (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   278
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   279
by (blast_tac (claset() addIs [leadsETo_Union_Int]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   280
(*Transitivity case has a delicate argument involving "cancellation"*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   281
by (rtac leadsETo_Un_duplicate2 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   282
by (etac leadsETo_cancel_Diff1 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   283
by (assume_tac 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   284
by (asm_full_simp_tac (simpset() addsimps [Int_Diff, Diff_triv]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   285
by (blast_tac (claset() addIs [leadsETo_weaken_L] 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   286
                        addDs [constrains_imp_subset]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   287
(*Basis case*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   288
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   289
by (blast_tac (claset() addIs [psp_ensures]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   290
by (subgoal_tac "A Int B' - (Ba Int B Un (B' - B)) = (A - Ba) Int B Int B'" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   291
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   292
qed "e_psp";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   293
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   294
Goal "[| F : A leadsTo[CC] A'; F : B co B';  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   295
\        ALL C:CC. C Int B Int B' : CC;  {}:CC |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   296
\     ==> F : (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   297
by (asm_full_simp_tac (simpset() addsimps e_psp::Int_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   298
qed "e_psp2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   299
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   300
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   301
(*** Special properties involving the parameter [CC] ***)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   302
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   303
(*??IS THIS NEEDED?? or is it just an example of what's provable??*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   304
Goal "[| F: (A leadsTo[givenBy v] B);  F Join G : v localTo[C] F;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   305
\        F Join G : stable C |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   306
\     ==> F Join G : ((C Int A) leadsTo[(%D. C Int D) `` givenBy v] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   307
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   308
by (stac Int_Union 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   309
by (blast_tac (claset() addIs [leadsETo_UN]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   310
by (blast_tac (claset() addIs [e_psp_stable2 RS leadsETo_weaken_L, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   311
			       leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   312
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   313
by (auto_tac (claset(),
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   314
	      simpset() addsimps [Int_Diff, ensures_def, stable_def,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   315
				  givenBy_eq_Collect,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   316
				  Join_localTo, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   317
				  Join_constrains, Join_transient]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   318
by (blast_tac (claset() addIs [transient_strengthen]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   319
by (blast_tac (claset() addDs [constrains_localTo_constrains]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   320
			addIs [constrains_Int RS constrains_weaken]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   321
by (blast_tac (claset() addIs [constrains_Int RS constrains_weaken]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   322
qed "gen_leadsETo_localTo_imp_Join_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   323
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   324
(*USED???
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   325
  Could replace this proof by instantiation of the one above with C=UNIV*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   326
Goal "[| F: (A leadsTo[givenBy v] B);  F Join G : v localTo[UNIV] F |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   327
\     ==> F Join G : (A leadsTo[givenBy v] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   328
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   329
by (blast_tac (claset() addIs [leadsETo_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   330
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   331
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   332
by (auto_tac (claset(),
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   333
	      simpset() addsimps [ensures_def, givenBy_eq_Collect,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   334
				  Join_localTo, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   335
				  Join_constrains, Join_transient]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   336
by (force_tac (claset() addDs [constrains_localTo_constrains], simpset()) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   337
qed "leadsETo_localTo_imp_Join_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   338
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   339
(*useful??*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   340
Goal "[| F Join G : (A leadsTo[CC] B);  ALL C:CC. G : stable C |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   341
\     ==> F: (A leadsTo[CC] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   342
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   343
by (blast_tac (claset() addIs [leadsETo_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   344
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   345
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   346
by (case_tac "A <= B" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   347
by (etac subset_imp_ensures 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   348
by (auto_tac (claset() addIs [constrains_weaken],
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   349
	      simpset() addsimps [stable_def, ensures_def, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   350
				  Join_constrains, Join_transient]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   351
by (REPEAT (thin_tac "?F : ?A co ?B" 1));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   352
by (etac transientE 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   353
by (rewtac constrains_def);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   354
by (blast_tac (claset() addSDs [bspec]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   355
qed "Join_leadsETo_stable_imp_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   356
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   357
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   358
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   359
(**** EXTEND/PROJECT PROPERTIES ****)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   360
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   361
Open_locale "Extend";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   362
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   363
(*Here h and f are locale constants*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   364
Goal "extend_set h `` (givenBy v) <= (givenBy (v o f))";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   365
by (simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   366
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   367
qed "extend_set_givenBy_subset";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   368
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   369
Goal "D : givenBy v ==> extend_set h D : givenBy (v o f)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   370
by (full_simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   371
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   372
qed "extend_set_givenBy_I";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   373
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   374
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   375
Goal "F : A leadsTo[CC] B \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   376
\     ==> extend h F : (extend_set h A) leadsTo[extend_set h `` CC] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   377
\                      (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   378
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   379
by (asm_simp_tac (simpset() addsimps [leadsETo_UN, extend_set_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   380
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   381
by (asm_simp_tac (simpset() addsimps [leadsETo_Basis, extend_ensures,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   382
				      extend_set_Diff_distrib RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   383
qed "leadsETo_imp_extend_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   384
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   385
(*NOW OBSOLETE: SEE BELOW !! Generalizes the version proved in Project.ML*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   386
Goalw [LOCALTO_def, transient_def, Diff_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   387
     "[| G : (v o f) localTo[C] extend h F;  project h C G : transient D;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   388
\        D : givenBy v |]    \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   389
\     ==> F : transient D";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   390
by (auto_tac (claset(), 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   391
	      simpset() addsimps [givenBy_eq_Collect]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   392
by (case_tac "Restrict C act : Restrict C ``extend_act h `` Acts F" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   393
by Auto_tac; 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   394
by (rtac bexI 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   395
by (assume_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   396
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   397
by (case_tac "{s. P (v s)} = {}" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   398
by (auto_tac (claset(),
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   399
	      simpset() addsimps [stable_def, constrains_def]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   400
by (subgoal_tac
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   401
    "ALL z. Restrict C act ^^ {s. v (f s) = z} <= {s. v (f s) = z}" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   402
by (blast_tac (claset() addSDs [bspec]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   403
by (thin_tac "ALL z. ?P z" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   404
by (subgoal_tac "project_act h (Restrict C act) ^^ {s. P (v s)} <= {s. P (v s)}" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   405
by (Clarify_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   406
by (asm_full_simp_tac (simpset() addsimps [project_act_def]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   407
by (force_tac (claset() addSDs [spec, ImageI RSN (2, subsetD)], simpset()) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   408
by (blast_tac (claset() addSDs [subsetD]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   409
qed "localTo_project_transient_transient";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   410
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   411
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   412
Goal "A Int extend_set h ((project_set h A) Int B) = A Int extend_set h B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   413
by (auto_tac (claset() addIs [project_set_I], 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   414
	      simpset()));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   415
qed "Int_extend_set_lemma";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   416
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   417
Goal "G : C co B ==> project h C G : project_set h C co project_set h B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   418
by (full_simp_tac (simpset() addsimps [constrains_def, project_def, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   419
				       project_act_def, project_set_def]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   420
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   421
qed "project_constrains_project_set";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   422
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   423
Goal "G : stable C ==> project h C G : stable (project_set h C)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   424
by (asm_full_simp_tac (simpset() addsimps [stable_def, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   425
					   project_constrains_project_set]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   426
qed "project_stable_project_set";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   427
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   428
(*!! Generalizes the version proved in Project.ML*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   429
Goalw [LOCALTO_def, transient_def, Diff_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   430
     "[| G : (v o f) localTo[C] extend h F;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   431
\        project h C G : transient (C' Int D);  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   432
\        project h C G : stable C';  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   433
\        D : givenBy v;  (C' Int D) <= D |]    \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   434
\     ==> F : transient (C' Int D)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   435
by (auto_tac (claset(), 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   436
	      simpset() addsimps [givenBy_eq_Collect]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   437
by (case_tac "Restrict C act : Restrict C ``extend_act h `` Acts F" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   438
by Auto_tac; 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   439
by (rtac bexI 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   440
by (assume_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   441
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   442
by (case_tac "(C' Int {s. P (v s)}) = {}" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   443
by (auto_tac (claset(),
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   444
	      simpset() addsimps [stable_def, constrains_def]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   445
by (subgoal_tac
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   446
    "ALL z. Restrict C act ^^ {s. v (f s) = z} <= {s. v (f s) = z}" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   447
by (blast_tac (claset() addSDs [bspec]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   448
by (thin_tac "ALL z. ?P z" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   449
by (subgoal_tac "project_act h (Restrict C act) ^^ (C' Int {s. P (v s)}) <= (C' Int {s. P (v s)})" 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   450
by (Clarify_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   451
by (asm_full_simp_tac (simpset() addsimps [project_act_def]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   452
by (thin_tac "(C' Int {s. P (v s)}) <= Domain ?A" 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   453
by (thin_tac "?A <= -C' Un ?B" 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   454
by (rtac conjI 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   455
by (force_tac (claset() addSDs [spec, ImageI RSN (2, subsetD)], simpset()) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   456
by (Blast_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   457
by (blast_tac (claset() addSDs [subsetD]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   458
qed "localTo_project_transient_transient";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   459
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   460
(*This version's stronger in the "ensures" precondition
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   461
  BUT there's no ensures_weaken_L*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   462
Goal "[| project h C G : transient (project_set h C Int (A-B)) --> \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   463
\          F : transient (project_set h C Int (A-B));  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   464
\        extend h F Join G : stable C;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   465
\        F Join project h C G : (project_set h C Int A) ensures B |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   466
\     ==> extend h F Join G : (C Int extend_set h A) ensures (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   467
by (stac (Int_extend_set_lemma RS sym) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   468
by (rtac Join_project_ensures 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   469
by (auto_tac (claset(), simpset() addsimps [Int_Diff]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   470
qed "Join_project_ensures_strong";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   471
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   472
Goal "[| extend h F Join G : stable C;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   473
\        F Join project h C G : (project_set h C Int A) leadsTo[(%D. project_set h C Int D)``givenBy v] B; \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   474
\        G : (v o f) localTo[C] extend h F |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   475
\     ==> extend h F Join G : \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   476
\           (C Int extend_set h (project_set h C Int A)) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   477
\           leadsTo[(%D. C Int extend_set h D)``givenBy v]  (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   478
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   479
by (asm_simp_tac (simpset() delsimps UN_simps
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   480
		  addsimps [Int_UN_distrib, leadsETo_UN, extend_set_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   481
by (blast_tac (claset() addIs [e_psp_stable2 RS leadsETo_weaken_L, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   482
			       leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   483
by (Clarify_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   484
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   485
by (etac rev_image_eqI 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   486
by (asm_simp_tac (simpset() addsimps [Int_Diff, Int_extend_set_lemma,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   487
				      extend_set_Diff_distrib RS sym]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   488
by (rtac Join_project_ensures_strong 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   489
by (auto_tac (claset() addIs [localTo_project_transient_transient,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   490
			      project_stable_project_set], 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   491
	      simpset() addsimps [Int_left_absorb, Join_stable]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   492
by (asm_simp_tac
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   493
    (simpset() addsimps [stable_ensures_Int RS ensures_weaken_R,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   494
			 Int_lower2, project_stable_project_set,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   495
			 Join_stable, extend_stable_project_set]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   496
val lemma = result();
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   497
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   498
Goal "[| extend h F Join G : stable C;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   499
\        F Join project h C G : (project_set h C Int A) leadsTo[(%D. project_set h C Int D)``givenBy v] B; \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   500
\        G : (v o f) localTo[C] extend h F |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   501
\     ==> extend h F Join G : (C Int extend_set h A) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   502
\           leadsTo[(%D. C Int extend_set h D)``givenBy v] (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   503
by (rtac (lemma RS leadsETo_weaken) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   504
by (auto_tac (claset() addIs [project_set_I], simpset()));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   505
qed "project_leadsETo_lemma";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   506
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   507
Goal "[| F Join project h UNIV G : A leadsTo[givenBy v] B;    \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   508
\        G : (v o f) localTo[UNIV] extend h F |]  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   509
\     ==> extend h F Join G : (extend_set h A) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   510
\           leadsTo[givenBy (v o f)] (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   511
by (rtac (make_elim project_leadsETo_lemma) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   512
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   513
by (etac leadsETo_givenBy 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   514
by (rtac extend_set_givenBy_subset 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   515
qed "project_leadsETo_D";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   516
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   517
Goal "[| F Join project h (reachable (extend h F Join G)) G \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   518
\            : A LeadsTo[givenBy v] B;    \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   519
\        G : (v o f) LocalTo extend h F |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   520
\     ==> extend h F Join G : \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   521
\           (extend_set h A) LeadsTo[givenBy (v o f)] (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   522
by (rtac (make_elim (subset_refl RS stable_reachable RS 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   523
		     project_leadsETo_lemma)) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   524
by (auto_tac (claset(), 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   525
	      simpset() addsimps [LeadsETo_def, LocalTo_def]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   526
by (asm_full_simp_tac 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   527
    (simpset() addsimps [project_set_reachable_extend_eq RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   528
by (etac (impOfSubs leadsETo_mono) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   529
by (blast_tac (claset() addIs [extend_set_givenBy_I]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   530
qed "project_LeadsETo_D";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   531
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   532
Goalw [extending_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   533
     "extending (%G. UNIV) h F \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   534
\               ((v o f) localTo[UNIV] extend h F) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   535
\               (extend_set h A leadsTo[givenBy (v o f)] extend_set h B) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   536
\               (A leadsTo[givenBy v] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   537
by (auto_tac (claset(), 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   538
	      simpset() addsimps [project_leadsETo_D, Join_localTo]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   539
qed "extending_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   540
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   541
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   542
Goalw [extending_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   543
     "extending (%G. reachable (extend h F Join G)) h F \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   544
\               ((v o f) LocalTo extend h F) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   545
\               (extend_set h A LeadsTo[givenBy (v o f)] extend_set h B) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   546
\               (A LeadsTo[givenBy v]  B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   547
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   548
by (force_tac (claset() addIs [project_LeadsETo_D],
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   549
	       simpset()addsimps [LocalTo_def, Join_assoc RS sym, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   550
				  Join_localTo]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   551
qed "extending_LeadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   552
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   553
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   554
Close_locale "Extend";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   555
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   556