src/HOL/UNITY/ELT.ML
author paulson
Wed, 22 Dec 1999 17:16:23 +0100
changeset 8072 5b95377d7538
parent 8069 19b9f92ca503
child 8110 f7651ede12b7
permissions -rw-r--r--
removing the "{} : CC" requirement for leadsTo[CC]
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
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
     9
Goalw [givenBy_def] "givenBy id = UNIV";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    10
by Auto_tac;
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    11
qed "givenBy_id";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    12
Addsimps [givenBy_id];
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    13
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    14
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
    15
by Safe_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    16
by (res_inst_tac [("x", "v `` ?u")] image_eqI 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    17
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    18
qed "givenBy_eq_all";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    19
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    20
Goal "givenBy v = {A. EX P. A = {s. P(v s)}}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    21
by (simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    22
by Safe_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    23
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
    24
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    25
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    26
qed "givenBy_eq_Collect";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    27
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    28
val prems =
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
    29
Goal "(!!x y. [| x:A;  v x = v y |] ==> y: A) ==> A: givenBy v";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    30
by (stac givenBy_eq_all 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    31
by (blast_tac (claset() addIs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    32
qed "givenByI";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    33
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    34
Goalw [givenBy_def] "[| A: givenBy v;  x:A;  v x = v y |] ==> y: A";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    35
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    36
qed "givenByD";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    37
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    38
Goal "{} : givenBy v";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    39
by (blast_tac (claset() addSIs [givenByI]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    40
qed "empty_mem_givenBy";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    41
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    42
AddIffs [empty_mem_givenBy];
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    43
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    44
Goal "A: givenBy v ==> EX P. A = {s. P(v s)}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    45
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
    46
by (full_simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    47
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    48
qed "givenBy_imp_eq_Collect";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    49
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    50
Goalw [givenBy_def] "EX P. A = {s. P(v s)} ==> A: givenBy v";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    51
by (Best_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    52
qed "eq_Collect_imp_givenBy";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    53
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    54
Goal "givenBy v = {A. EX P. A = {s. P(v s)}}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    55
by (blast_tac (claset() addIs [eq_Collect_imp_givenBy,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    56
			       givenBy_imp_eq_Collect]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    57
qed "givenBy_eq_eq_Collect";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    58
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    59
(*preserving v preserves properties given by v*)
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    60
Goal "[| F : preserves v;  D : givenBy v |] ==> F : stable D";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    61
by (force_tac (claset(), 
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    62
	       simpset() addsimps [impOfSubs preserves_subset_stable, 
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    63
				   givenBy_eq_Collect]) 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
    64
qed "preserves_givenBy_imp_stable";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    65
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
    66
Goal "givenBy (w o v) <= givenBy v";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
    67
by (simp_tac (simpset() addsimps [givenBy_eq_Collect]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
    68
by (Deepen_tac 0 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
    69
qed "givenBy_o_subset";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
    70
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    71
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    72
(** Standard leadsTo rules **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    73
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
    74
Goalw [leadsETo_def]
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
    75
     "[| F: A ensures B;  A-B: insert {} CC |] ==> F : A leadsTo[CC] B";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    76
by (blast_tac (claset() addIs [elt.Basis]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    77
qed "leadsETo_Basis";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    78
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    79
Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    80
     "[| 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
    81
by (blast_tac (claset() addIs [elt.Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    82
qed "leadsETo_Trans";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    83
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    84
(*Useful with cancellation, disjunction*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    85
Goal "F : A leadsTo[CC] (A' Un A') ==> F : A leadsTo[CC] A'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    86
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    87
qed "leadsETo_Un_duplicate";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    88
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    89
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
    90
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    91
qed "leadsETo_Un_duplicate2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    92
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    93
(*The Union introduction rule as we should have liked to state it*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    94
val prems = Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    95
    "(!!A. A : S ==> F : A leadsTo[CC] B) ==> F : (Union S) leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    96
by (blast_tac (claset() addIs [elt.Union] addDs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    97
qed "leadsETo_Union";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    98
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
    99
val prems = Goal
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   100
    "(!!i. i : I ==> F : (A i) leadsTo[CC] B) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   101
\    ==> F : (UN i:I. A i) leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   102
by (stac (Union_image_eq RS sym) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   103
by (blast_tac (claset() addIs leadsETo_Union::prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   104
qed "leadsETo_UN";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   105
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   106
(*The INDUCTION rule as we should have liked to state it*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   107
val major::prems = Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   108
  "[| F : za leadsTo[CC] zb;  \
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   109
\     !!A B. [| F : A ensures B;  A-B : insert {} CC |] ==> P A B; \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   110
\     !!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
   111
\              ==> P A C; \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   112
\     !!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
   113
\  |] ==> P za zb";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   114
by (rtac (major RS CollectD RS elt.induct) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   115
by (REPEAT (blast_tac (claset() addIs prems) 1));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   116
qed "leadsETo_induct";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   117
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   118
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   119
(** New facts involving leadsETo **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   120
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   121
Goal "CC' <= CC ==> (A leadsTo[CC'] B) <= (A leadsTo[CC] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   122
by Safe_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   123
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   124
by (blast_tac (claset() addIs [leadsETo_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   125
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   126
by (blast_tac (claset() addIs [leadsETo_Basis]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   127
qed "leadsETo_mono";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   128
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   129
val prems = Goalw [leadsETo_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   130
 "(!!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
   131
by (simp_tac (HOL_ss addsimps [Int_Union_Union]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   132
by (blast_tac (claset() addIs [elt.Union] addDs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   133
qed "leadsETo_Union_Int";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   134
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   135
(*Binary union introduction rule*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   136
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
   137
by (stac Un_eq_Union 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   138
by (blast_tac (claset() addIs [leadsETo_Union]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   139
qed "leadsETo_Un";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   140
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   141
val prems = 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   142
Goal "(!!x. x : A ==> F : {x} leadsTo[CC] B) ==> F : A leadsTo[CC] B";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   143
by (stac (UN_singleton RS sym) 1 THEN rtac leadsETo_UN 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   144
by (blast_tac (claset() addIs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   145
qed "single_leadsETo_I";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   146
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   147
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   148
Goal "A<=B ==> F : A leadsTo[CC] B";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   149
by (asm_simp_tac (simpset() addsimps [subset_imp_ensures RS leadsETo_Basis,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   150
				      Diff_eq_empty_iff RS iffD2]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   151
qed "subset_imp_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   152
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   153
bind_thm ("empty_leadsETo", empty_subsetI RS subset_imp_leadsETo);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   154
Addsimps [empty_leadsETo];
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   155
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   156
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   157
(** Weakening laws **)
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   158
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   159
Goal "[| F : A leadsTo[CC] A';  A'<=B' |] ==> F : A leadsTo[CC] B'";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   160
by (blast_tac (claset() addIs [subset_imp_leadsETo, leadsETo_Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   161
qed "leadsETo_weaken_R";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   162
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   163
Goal "[| F : A leadsTo[CC] A'; B<=A |] ==> F : B leadsTo[CC] A'";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   164
by (blast_tac (claset() addIs [leadsETo_Trans, subset_imp_leadsETo]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   165
qed_spec_mp "leadsETo_weaken_L";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   166
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   167
(*Distributes over binary unions*)
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   168
Goal "F : (A Un B) leadsTo[CC] C  =  \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   169
\     (F : A leadsTo[CC] C & F : B leadsTo[CC] C)";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   170
by (blast_tac (claset() addIs [leadsETo_Un, leadsETo_weaken_L]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   171
qed "leadsETo_Un_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   172
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   173
Goal "F : (UN i:I. A i) leadsTo[CC] B  =  \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   174
\     (ALL i : I. F : (A i) leadsTo[CC] B)";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   175
by (blast_tac (claset() addIs [leadsETo_UN, leadsETo_weaken_L]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   176
qed "leadsETo_UN_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   177
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   178
Goal "F : (Union S) leadsTo[CC] B  =  (ALL A : S. F : A leadsTo[CC] B)";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   179
by (blast_tac (claset() addIs [leadsETo_Union, leadsETo_weaken_L]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   180
qed "leadsETo_Union_distrib";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   181
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   182
Goal "[| F : A leadsTo[CC'] A'; B<=A; A'<=B';  CC' <= CC |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   183
\     ==> F : B leadsTo[CC] B'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   184
by (dtac (impOfSubs leadsETo_mono) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   185
by (assume_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   186
by (blast_tac (claset() addIs [leadsETo_weaken_R, leadsETo_weaken_L,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   187
			       leadsETo_Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   188
qed "leadsETo_weaken";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   189
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   190
Goal "[| F : A leadsTo[CC] A';  CC <= givenBy v |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   191
\     ==> F : A leadsTo[givenBy v] A'";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   192
by (blast_tac (claset() addIs [empty_mem_givenBy, leadsETo_weaken]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   193
qed "leadsETo_givenBy";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   194
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   195
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   196
(*Set difference*)
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   197
Goal "[| F : (A-B) leadsTo[CC] C; F : B leadsTo[CC] C |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   198
\     ==> F : A leadsTo[CC] C";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   199
by (blast_tac (claset() addIs [leadsETo_Un, leadsETo_weaken]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   200
qed "leadsETo_Diff";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   201
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   202
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   203
(** Meta or object quantifier ???
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   204
    see ball_constrains_UN in UNITY.ML***)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   205
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   206
val prems = goal thy
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   207
   "[| !! i. i:I ==> F : (A i) leadsTo[CC] (A' i) |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   208
\   ==> F : (UN i:I. A i) leadsTo[CC] (UN i:I. A' i)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   209
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   210
by (blast_tac (claset() addIs [leadsETo_Union, leadsETo_weaken_R] 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   211
                        addIs prems) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   212
qed "leadsETo_UN_UN";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   213
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   214
(*Binary union version*)
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   215
Goal "[| F : A leadsTo[CC] A';  F : B leadsTo[CC] B' |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   216
\     ==> F : (A Un B) leadsTo[CC] (A' Un B')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   217
by (blast_tac (claset() addIs [leadsETo_Un, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   218
			       leadsETo_weaken_R]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   219
qed "leadsETo_Un_Un";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   220
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   221
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   222
(** The cancellation law **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   223
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   224
Goal "[| F : A leadsTo[CC] (A' Un B); F : B leadsTo[CC] B' |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   225
\     ==> F : A leadsTo[CC] (A' Un B')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   226
by (blast_tac (claset() addIs [leadsETo_Un_Un, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   227
			       subset_imp_leadsETo, leadsETo_Trans]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   228
qed "leadsETo_cancel2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   229
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   230
Goal "[| F : A leadsTo[CC] (A' Un B); F : (B-A') leadsTo[CC] B' |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   231
\     ==> F : A leadsTo[CC] (A' Un B')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   232
by (rtac leadsETo_cancel2 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   233
by (assume_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   234
by (ALLGOALS Asm_simp_tac);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   235
qed "leadsETo_cancel_Diff2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   236
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   237
Goal "[| F : A leadsTo[CC] (B Un A'); F : B leadsTo[CC] B' |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   238
\   ==> F : A leadsTo[CC] (B' Un A')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   239
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   240
by (blast_tac (claset() addSIs [leadsETo_cancel2]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   241
qed "leadsETo_cancel1";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   242
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   243
Goal "[| F : A leadsTo[CC] (B Un A'); F : (B-A') leadsTo[CC] B' |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   244
\   ==> F : A leadsTo[CC] (B' Un A')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   245
by (rtac leadsETo_cancel1 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   246
by (assume_tac 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   247
by (ALLGOALS Asm_simp_tac);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   248
qed "leadsETo_cancel_Diff1";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   249
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   250
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   251
(** The impossibility law **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   252
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   253
Goal "F : A leadsTo[CC] B ==> B={} --> A={}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   254
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   255
by (ALLGOALS Asm_simp_tac);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   256
by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   257
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   258
val lemma = result() RS mp;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   259
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   260
Goal "F : A leadsTo[CC] {} ==> A={}";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   261
by (blast_tac (claset() addSIs [lemma]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   262
qed "leadsETo_empty";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   263
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   264
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   265
(** PSP: Progress-Safety-Progress **)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   266
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   267
(*Special case of PSP: Misra's "stable conjunction"*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   268
Goalw [stable_def]
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   269
   "[| 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
   270
\   ==> F : (A Int B) leadsTo[CC] (A' Int B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   271
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   272
by (blast_tac (claset() addIs [leadsETo_Union_Int]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   273
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   274
by (rtac leadsETo_Basis 1);
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   275
by (force_tac (claset(),
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   276
	       simpset() addsimps [Diff_Int_distrib2 RS sym]) 2);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   277
by (asm_full_simp_tac
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   278
    (simpset() addsimps [ensures_def, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   279
			 Diff_Int_distrib2 RS sym, Int_Un_distrib2 RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   280
by (blast_tac (claset() addIs [transient_strengthen, constrains_Int]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   281
qed "e_psp_stable";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   282
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   283
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
   284
\     ==> F : (B Int A) leadsTo[CC] (B Int A')";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   285
by (asm_simp_tac (simpset() addsimps e_psp_stable::Int_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   286
qed "e_psp_stable2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   287
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   288
Goal "[| F : A leadsTo[CC] A'; F : B co B';  \
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   289
\        ALL C:CC. C Int B Int B' : CC |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   290
\     ==> F : (A Int B') leadsTo[CC] ((A' Int B) Un (B' - B))";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   291
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   292
by (blast_tac (claset() addIs [leadsETo_Union_Int]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   293
(*Transitivity case has a delicate argument involving "cancellation"*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   294
by (rtac leadsETo_Un_duplicate2 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   295
by (etac leadsETo_cancel_Diff1 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   296
by (asm_full_simp_tac (simpset() addsimps [Int_Diff, Diff_triv]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   297
by (blast_tac (claset() addIs [leadsETo_weaken_L] 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   298
                        addDs [constrains_imp_subset]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   299
(*Basis case*)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   300
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   301
by (blast_tac (claset() addIs [psp_ensures]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   302
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
   303
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   304
qed "e_psp";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   305
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   306
Goal "[| F : A leadsTo[CC] A'; F : B co B';  \
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   307
\        ALL C:CC. C Int B Int B' : CC |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   308
\     ==> F : (B' Int A) leadsTo[CC] ((B Int A') Un (B' - B))";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   309
by (asm_full_simp_tac (simpset() addsimps e_psp::Int_ac) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   310
qed "e_psp2";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   311
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   312
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   313
(*** Special properties involving the parameter [CC] ***)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   314
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   315
(*??IS THIS NEEDED?? or is it just an example of what's provable??*)
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   316
Goal "[| F: (A leadsTo[givenBy v] B);  G : preserves v;  \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   317
\        F Join G : stable C |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   318
\     ==> F Join G : ((C Int A) leadsTo[(%D. C Int D) `` givenBy v] B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   319
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   320
by (stac Int_Union 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   321
by (blast_tac (claset() addIs [leadsETo_UN]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   322
by (blast_tac (claset() addIs [e_psp_stable2 RS leadsETo_weaken_L, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   323
			       leadsETo_Trans]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   324
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   325
by (auto_tac (claset(),
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   326
	      simpset() addsimps [Diff_eq_empty_iff RS iffD2,
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   327
				  Int_Diff, ensures_def,
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   328
				  givenBy_eq_Collect, Join_stable,
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   329
				  Join_constrains, Join_transient]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   330
by (blast_tac (claset() addIs [transient_strengthen]) 3);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   331
by (ALLGOALS (dres_inst_tac [("P1","P")] (impOfSubs preserves_subset_stable)));
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   332
by (rewtac stable_def);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   333
by (blast_tac (claset() addSEs [equalityE]
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   334
			addIs [constrains_Int RS constrains_weaken]) 2);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   335
by (blast_tac (claset() addSEs [equalityE]
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   336
			addIs [constrains_Int RS constrains_weaken]) 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   337
qed "gen_leadsETo_imp_Join_leadsETo";
8044
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],
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   349
              simpset() addsimps [stable_def, ensures_def, 
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   350
                                  Join_constrains, Join_transient]));
8044
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
8067
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   357
(**** Relationship with traditional "leadsTo", strong & weak ****)
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   358
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   359
(** strong **)
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   360
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   361
Goal "(A leadsTo[CC] B) <= (A leadsTo B)";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   362
by Safe_tac;
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   363
by (etac leadsETo_induct 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   364
by (blast_tac (claset() addIs [leadsTo_Union]) 3);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   365
by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   366
by (blast_tac (claset() addIs [leadsTo_Basis]) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   367
qed "leadsETo_subset_leadsTo";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   368
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   369
Goal "(A leadsTo[UNIV] B) = (A leadsTo B)";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   370
by Safe_tac;
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   371
by (etac (impOfSubs leadsETo_subset_leadsTo) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   372
(*right-to-left case*)
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   373
by (etac leadsTo_induct 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   374
by (blast_tac (claset() addIs [leadsETo_Union]) 3);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   375
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   376
by (blast_tac (claset() addIs [leadsETo_Basis]) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   377
qed "leadsETo_UNIV_eq_leadsTo";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   378
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   379
(**** weak ****)
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   380
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   381
Goalw [LeadsETo_def]
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   382
     "A LeadsTo[CC] B = \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   383
\       {F. F : (reachable F Int A) leadsTo[(%C. reachable F Int C) `` CC] \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   384
\       (reachable F Int B)}";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   385
by (blast_tac (claset() addDs [e_psp_stable2] addIs [leadsETo_weaken]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   386
qed "LeadsETo_eq_leadsETo";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   387
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   388
(*** Introduction rules: Basis, Trans, Union ***)
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   389
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   390
Goal "[| F : A LeadsTo[CC] B;  F : B LeadsTo[CC] C |] \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   391
\     ==> F : A LeadsTo[CC] C";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   392
by (asm_full_simp_tac (simpset() addsimps [LeadsETo_eq_leadsETo]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   393
by (blast_tac (claset() addIs [leadsETo_Trans]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   394
qed "LeadsETo_Trans";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   395
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   396
val prems = Goalw [LeadsETo_def]
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   397
     "(!!A. A : S ==> F : A LeadsTo[CC] B) ==> F : (Union S) LeadsTo[CC] B";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   398
by (Simp_tac 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   399
by (stac Int_Union 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   400
by (blast_tac (claset() addIs [leadsETo_UN] addDs prems) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   401
qed "LeadsETo_Union";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   402
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   403
val prems = 
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   404
Goal "(!!i. i : I ==> F : (A i) LeadsTo[CC] B) \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   405
\     ==> F : (UN i:I. A i) LeadsTo[CC] B";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   406
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   407
by (blast_tac (claset() addIs (LeadsETo_Union::prems)) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   408
qed "LeadsETo_UN";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   409
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   410
(*Binary union introduction rule*)
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   411
Goal "[| F : A LeadsTo[CC] C; F : B LeadsTo[CC] C |] \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   412
\     ==> F : (A Un B) LeadsTo[CC] C";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   413
by (stac Un_eq_Union 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   414
by (blast_tac (claset() addIs [LeadsETo_Union]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   415
qed "LeadsETo_Un";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   416
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   417
(*Lets us look at the starting state*)
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   418
val prems = 
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   419
Goal "(!!s. s : A ==> F : {s} LeadsTo[CC] B) ==> F : A LeadsTo[CC] B";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   420
by (stac (UN_singleton RS sym) 1 THEN rtac LeadsETo_UN 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   421
by (blast_tac (claset() addIs prems) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   422
qed "single_LeadsETo_I";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   423
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   424
Goal "A <= B ==> F : A LeadsTo[CC] B";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   425
by (simp_tac (simpset() addsimps [LeadsETo_def]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   426
by (blast_tac (claset() addIs [subset_imp_leadsETo]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   427
qed "subset_imp_LeadsETo";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   428
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   429
bind_thm ("empty_LeadsETo", empty_subsetI RS subset_imp_LeadsETo);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   430
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   431
Goal "[| F : A LeadsTo[CC] A';  A' <= B' |] ==> F : A LeadsTo[CC] B'";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   432
by (full_simp_tac (simpset() addsimps [LeadsETo_def]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   433
by (blast_tac (claset() addIs [leadsETo_weaken_R]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   434
qed_spec_mp "LeadsETo_weaken_R";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   435
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   436
Goal "[| F : A LeadsTo[CC] A';  B <= A |] ==> F : B LeadsTo[CC] A'";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   437
by (full_simp_tac (simpset() addsimps [LeadsETo_def]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   438
by (blast_tac (claset() addIs [leadsETo_weaken_L]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   439
qed_spec_mp "LeadsETo_weaken_L";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   440
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   441
Goal "[| F : A LeadsTo[CC'] A';   \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   442
\        B <= A;  A' <= B';  CC' <= CC |] \
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   443
\     ==> F : B LeadsTo[CC] B'";
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   444
by (full_simp_tac (simpset() addsimps [LeadsETo_def]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   445
by (blast_tac (claset() addIs [leadsETo_weaken]) 1);
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   446
qed "LeadsETo_weaken";
8067
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   447
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   448
Goalw [LeadsETo_def, LeadsTo_def] "(A LeadsTo[CC] B) <= (A LeadsTo B)";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   449
by (blast_tac (claset() addIs [impOfSubs leadsETo_subset_leadsTo]) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   450
qed "LeadsETo_subset_LeadsTo";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   451
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   452
(*Postcondition can be strengthened to (reachable F Int B) *)
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   453
Goal "F : A ensures B ==> F : (reachable F Int A) ensures B";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   454
by (rtac (stable_ensures_Int RS ensures_weaken_R) 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   455
by Auto_tac;
8067
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   456
qed "reachable_ensures";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   457
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   458
Goal "F : A leadsTo B ==> F : (reachable F Int A) leadsTo[Pow(reachable F)] B";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   459
by (etac leadsTo_induct 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   460
by (stac Int_Union 3);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   461
by (blast_tac (claset() addIs [leadsETo_UN]) 3);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   462
by (blast_tac (claset() addDs [e_psp_stable2] 
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   463
                        addIs [leadsETo_Trans, leadsETo_weaken_L]) 2);
8067
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   464
by (blast_tac (claset() addIs [reachable_ensures, leadsETo_Basis]) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   465
val lemma = result();
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   466
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   467
Goal "(A LeadsTo[UNIV] B) = (A LeadsTo B)";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   468
by Safe_tac;
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   469
by (etac (impOfSubs LeadsETo_subset_LeadsTo) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   470
(*right-to-left case*)
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   471
by (rewrite_goals_tac [LeadsETo_def, LeadsTo_def]);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   472
by (fast_tac (claset() addEs [lemma RS leadsETo_weaken]) 1);
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   473
qed "LeadsETo_UNIV_eq_LeadsTo";
225e3b45b766 now workign as far as System_Alloc_Progress
paulson
parents: 8055
diff changeset
   474
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   475
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   476
(**** EXTEND/PROJECT PROPERTIES ****)
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   477
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   478
Open_locale "Extend";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   479
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   480
Goal "givenBy (v o f) = extend_set h `` (givenBy v)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   481
by (simp_tac (simpset() addsimps [givenBy_eq_Collect]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   482
by (Deepen_tac 0 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   483
qed "givenBy_o_eq_extend_set";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   484
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   485
Goal "givenBy f = range (extend_set h)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   486
by (simp_tac (simpset() addsimps [givenBy_eq_Collect]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   487
by (Deepen_tac 0 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   488
qed "givenBy_eq_extend_set";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   489
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   490
Goal "D : givenBy v ==> extend_set h D : givenBy (v o f)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   491
by (full_simp_tac (simpset() addsimps [givenBy_eq_all]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   492
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   493
qed "extend_set_givenBy_I";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   494
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   495
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   496
Goal "F : A leadsTo[CC] B \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   497
\     ==> extend h F : (extend_set h A) leadsTo[extend_set h `` CC] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   498
\                      (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   499
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   500
by (asm_simp_tac (simpset() addsimps [leadsETo_UN, extend_set_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   501
by (blast_tac (claset() addIs [leadsETo_Trans]) 2);
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   502
by (force_tac (claset() addIs [leadsETo_Basis, subset_imp_ensures],
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   503
	       simpset() addsimps [extend_ensures,
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   504
				   extend_set_Diff_distrib RS sym]) 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   505
qed "leadsETo_imp_extend_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   506
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   507
(*MOVE to Extend.ML?*)
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   508
Goal "A Int extend_set h ((project_set h A) Int B) = A Int extend_set h B";
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   509
by (auto_tac (claset(), simpset() addsimps [split_extended_all]));
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   510
qed "Int_extend_set_lemma";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   511
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   512
(*MOVE to extend.ML??*)
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   513
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
   514
by (full_simp_tac (simpset() addsimps [constrains_def, project_def, 
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   515
				       project_act_def]) 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   516
by (Blast_tac 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   517
qed "project_constrains_project_set";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   518
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   519
(*MOVE to extend.ML??*)
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   520
Goal "G : stable C ==> project h C G : stable (project_set h C)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   521
by (asm_full_simp_tac (simpset() addsimps [stable_def, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   522
					   project_constrains_project_set]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   523
qed "project_stable_project_set";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   524
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   525
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   526
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   527
(*NOT USED, but analogous to preserves_project_transient_empty in Project.ML*)
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   528
Goal "[| G : preserves (v o f);  project h C G : transient D;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   529
\        D : givenBy v |] ==> D={}";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   530
by (rtac stable_transient_empty 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   531
by (assume_tac 2);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   532
(*If addIs then PROOF FAILED at depth 2*)
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   533
by (blast_tac (claset() addSIs [preserves_givenBy_imp_stable,
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   534
				project_preserves_I]) 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   535
result();
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   536
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   537
(*Generalizes the version proved in Project.ML*)
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   538
Goal "[| G : preserves (v o f);  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   539
\        project h C G : transient (C' Int D);  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   540
\        project h C G : stable C';  D : givenBy v |]    \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   541
\     ==> C' Int D = {}";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   542
by (rtac stable_transient_empty 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   543
by (assume_tac 2);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   544
(*If addIs then PROOF FAILED at depth 3*)
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   545
by (blast_tac (claset() addSIs [stable_Int, preserves_givenBy_imp_stable,
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   546
				project_preserves_I]) 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   547
qed "preserves_o_project_transient_empty";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   548
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   549
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   550
(*This version's stronger in the "ensures" precondition
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   551
  BUT there's no ensures_weaken_L*)
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   552
Goal "[| project h C G ~: transient (project_set h C Int (A-B)) | \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   553
\          project_set h C Int (A - B) = {};  \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   554
\        extend h F Join G : stable C;  \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   555
\        F Join project h C G : (project_set h C Int A) ensures B |] \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   556
\     ==> 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
   557
by (stac (Int_extend_set_lemma RS sym) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   558
by (rtac Join_project_ensures 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   559
by (auto_tac (claset(), simpset() addsimps [Int_Diff]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   560
qed "Join_project_ensures_strong";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   561
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   562
Goal "[| extend h F Join G : stable C;  \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   563
\        F Join project h C G : (project_set h C Int A) leadsTo[(%D. project_set h C Int D)``givenBy v] B;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   564
\        G : preserves (v o f) |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   565
\     ==> extend h F Join G : \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   566
\           (C Int extend_set h (project_set h C Int A)) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   567
\           leadsTo[(%D. C Int extend_set h D)``givenBy v]  (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   568
by (etac leadsETo_induct 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   569
by (asm_simp_tac (simpset() delsimps UN_simps
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   570
		  addsimps [Int_UN_distrib, leadsETo_UN, extend_set_Union]) 3);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   571
by (blast_tac (claset() addIs [e_psp_stable2 RS leadsETo_weaken_L, 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   572
			       leadsETo_Trans]) 2);
8072
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   573
auto();
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   574
by (force_tac (claset() addIs [leadsETo_Basis, subset_imp_ensures],
5b95377d7538 removing the "{} : CC" requirement for leadsTo[CC]
paulson
parents: 8069
diff changeset
   575
	       simpset()) 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   576
by (rtac leadsETo_Basis 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   577
by (asm_simp_tac (simpset() addsimps [Int_Diff, Int_extend_set_lemma,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   578
				      extend_set_Diff_distrib RS sym]) 2);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   579
by (rtac Join_project_ensures_strong 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   580
by (auto_tac (claset() addDs [preserves_o_project_transient_empty]
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   581
		       addIs [project_stable_project_set], 
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   582
	      simpset() addsimps [Int_left_absorb, Join_stable]));
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   583
by (asm_simp_tac
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   584
    (simpset() addsimps [stable_ensures_Int RS ensures_weaken_R,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   585
			 Int_lower2, project_stable_project_set,
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   586
			 Join_stable, extend_stable_project_set]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   587
val lemma = result();
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   588
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   589
Goal "[| extend h F Join G : stable C;  \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   590
\        F Join project h C G : \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   591
\            (project_set h C Int A) \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   592
\            leadsTo[(%D. project_set h C Int D)``givenBy v] B;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   593
\        G : preserves (v o f) |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   594
\     ==> extend h F Join G : (C Int extend_set h A) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   595
\           leadsTo[(%D. C Int extend_set h D)``givenBy v] (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   596
by (rtac (lemma RS leadsETo_weaken) 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   597
by (auto_tac (claset(), 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   598
	      simpset() addsimps [split_extended_all]));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   599
qed "project_leadsETo_D_lemma";
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   600
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   601
Goal "[| F Join project h UNIV G : A leadsTo[givenBy v] B;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   602
\        G : preserves (v o f) |]  \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   603
\     ==> extend h F Join G : (extend_set h A) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   604
\           leadsTo[givenBy (v o f)] (extend_set h B)";
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   605
by (rtac (make_elim project_leadsETo_D_lemma) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   606
by (stac stable_UNIV 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   607
by Auto_tac;
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   608
by (etac leadsETo_givenBy 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   609
by (rtac (givenBy_o_eq_extend_set RS equalityD2) 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   610
qed "project_leadsETo_D";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   611
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   612
Goal "[| F Join project h (reachable (extend h F Join G)) G \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   613
\            : A LeadsTo[givenBy v] B;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   614
\        G : preserves (v o f) |] \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   615
\     ==> extend h F Join G : \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   616
\           (extend_set h A) LeadsTo[givenBy (v o f)] (extend_set h B)";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   617
by (rtac (make_elim (subset_refl RS stable_reachable RS 
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   618
		     project_leadsETo_D_lemma)) 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   619
by (auto_tac (claset(), 
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   620
	      simpset() addsimps [LeadsETo_def]));
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   621
by (asm_full_simp_tac 
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   622
    (simpset() addsimps [project_set_reachable_extend_eq RS sym]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   623
by (etac (impOfSubs leadsETo_mono) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   624
by (blast_tac (claset() addIs [extend_set_givenBy_I]) 1);
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   625
qed "project_LeadsETo_D";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   626
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   627
Goalw [extending_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   628
     "extending (v o f) (%G. UNIV) h F \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   629
\               (extend_set h A leadsTo[givenBy (v o f)] extend_set h B) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   630
\               (A leadsTo[givenBy v] B)";
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   631
by (auto_tac (claset(), simpset() addsimps [project_leadsETo_D]));
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   632
qed "extending_leadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   633
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   634
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   635
Goalw [extending_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   636
     "extending (v o f) (%G. reachable (extend h F Join G)) h F \
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   637
\               (extend_set h A LeadsTo[givenBy (v o f)] extend_set h B) \
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   638
\               (A LeadsTo[givenBy v]  B)";
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8044
diff changeset
   639
by (blast_tac (claset() addIs [project_LeadsETo_D]) 1);
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   640
qed "extending_LeadsETo";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   641
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   642
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   643
(*** leadsETo in the precondition ***)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   644
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   645
Goalw [transient_def]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   646
     "[| G : transient (C Int extend_set h A);  G : stable C |]  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   647
\     ==> project h C G : transient (project_set h C Int A)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   648
by (auto_tac (claset(), simpset() addsimps [Domain_project_act]));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   649
by (subgoal_tac "act ^^ (C Int extend_set h A) <= - extend_set h A" 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   650
by (asm_full_simp_tac 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   651
    (simpset() addsimps [stable_def, constrains_def]) 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   652
by (Blast_tac 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   653
(*back to main goal*)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   654
by (thin_tac "?AA <= -C Un ?BB" 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   655
by (ball_tac 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   656
by (asm_full_simp_tac 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   657
    (simpset() addsimps [extend_set_def, project_act_def]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   658
by (Blast_tac 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   659
qed "transient_extend_set_imp_project_transient";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   660
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   661
(*converse might hold too?*)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   662
Goalw [transient_def]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   663
     "project h C (extend h F) : transient (project_set h C Int D) \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   664
\     ==> F : transient (project_set h C Int D)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   665
by (auto_tac (claset(), simpset() addsimps [Domain_project_act]));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   666
by (rtac bexI 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   667
by (assume_tac 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   668
by Auto_tac;
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   669
by (rewtac extend_act_def);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   670
by (Blast_tac 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   671
qed "project_extend_transient_D";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   672
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   673
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   674
Goal "[| extend h F : stable C;  G : stable C;  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   675
\        extend h F Join G : A ensures B;  A-B = C Int extend_set h D |]  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   676
\     ==> F Join project h C G  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   677
\           : (project_set h C Int project_set h A) ensures (project_set h B)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   678
by (asm_full_simp_tac
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   679
    (simpset() addsimps [ensures_def, Join_constrains, project_constrains, 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   680
			 Join_transient, extend_transient]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   681
by (Clarify_tac 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   682
by (REPEAT_FIRST (rtac conjI));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   683
(*first subgoal*)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   684
by (blast_tac (claset() addIs [extend_stable_project_set RS stableD RS 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   685
			       constrains_Int RS constrains_weaken]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   686
	                addSDs [extend_constrains_project_set]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   687
			addSEs [equalityE]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   688
(*2nd subgoal*)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   689
by (etac (stableD RS constrains_Int RS constrains_weaken) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   690
by (assume_tac 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   691
by (Blast_tac 3);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   692
by (full_simp_tac (simpset() addsimps [extend_set_Int_distrib,
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   693
				       extend_set_Un_distrib]) 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   694
by (blast_tac (claset() addSIs [impOfSubs extend_set_project_set]) 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   695
by (full_simp_tac (simpset() addsimps [extend_set_def]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   696
by (blast_tac (claset() addSEs [equalityE]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   697
(*The transient part*)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   698
by Auto_tac;
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   699
by (force_tac (claset() addSEs [equalityE]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   700
	                addIs [transient_extend_set_imp_project_transient RS
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   701
			       transient_strengthen], 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   702
              simpset()) 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   703
by (full_simp_tac (simpset() addsimps [Int_Diff]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   704
by (force_tac (claset() addSEs [equalityE]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   705
	                addIs [transient_extend_set_imp_project_transient RS 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   706
			 project_extend_transient_D RS transient_strengthen], 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   707
              simpset()) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   708
qed "ensures_extend_set_imp_project_ensures";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   709
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   710
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   711
(*Lemma for the Trans case*)
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   712
Goal "[| extend h F Join G : stable C;    \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   713
\        F Join project h C G    \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   714
\          : project_set h C Int project_set h A leadsTo project_set h B |] \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   715
\     ==> F Join project h C G    \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   716
\           : project_set h C Int project_set h A leadsTo    \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   717
\             project_set h C Int project_set h B";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   718
by (rtac (psp_stable2 RS leadsTo_weaken_L) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   719
by (auto_tac (claset(),
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   720
	      simpset() addsimps [project_stable_project_set, Join_stable, 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   721
				  extend_stable_project_set]));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   722
val lemma = result();
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   723
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   724
Goal "[| extend h F Join G : stable C;  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   725
\        extend h F Join G : \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   726
\          (C Int A) leadsTo[(%D. C Int D)``givenBy f]  B |]  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   727
\ ==> F Join project h C G  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   728
\   : (project_set h C Int project_set h (C Int A)) leadsTo (project_set h B)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   729
by (etac leadsETo_induct 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   730
by (asm_simp_tac (HOL_ss addsimps [Int_UN_distrib, project_set_Union]) 3);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   731
by (blast_tac (claset() addIs [leadsTo_UN]) 3);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   732
by (blast_tac (claset() addIs [leadsTo_Trans, lemma]) 2);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   733
by (asm_full_simp_tac 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   734
    (simpset() addsimps [givenBy_eq_extend_set, Join_stable]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   735
by (rtac leadsTo_Basis 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   736
by (blast_tac (claset() addIs [leadsTo_Basis,
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   737
			       ensures_extend_set_imp_project_ensures]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   738
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   739
qed "project_leadsETo_I_lemma";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   740
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   741
Goal "extend h F Join G : (extend_set h A) leadsTo[givenBy f] (extend_set h B)\
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   742
\     ==> F Join project h UNIV G : A leadsTo B";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   743
by (rtac (project_leadsETo_I_lemma RS leadsTo_weaken) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   744
by Auto_tac;
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   745
qed "project_leadsETo_I";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   746
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   747
Goal "extend h F Join G : (extend_set h A) LeadsTo[givenBy f] (extend_set h B)\
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   748
\     ==> F Join project h (reachable (extend h F Join G)) G  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   749
\          : A LeadsTo B";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   750
by (full_simp_tac (simpset() addsimps [LeadsTo_def, LeadsETo_def]) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   751
by (rtac (project_leadsETo_I_lemma RS leadsTo_weaken) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   752
by (auto_tac (claset(), 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   753
	      simpset() addsimps [project_set_reachable_extend_eq RS sym]));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   754
qed "project_LeadsETo_I";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   755
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   756
Goalw [projecting_def]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   757
     "projecting (%G. UNIV) h F \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   758
\                (extend_set h A leadsTo[givenBy f] extend_set h B) \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   759
\                (A leadsTo B)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   760
by (force_tac (claset() addDs [project_leadsETo_I], simpset()) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   761
qed "projecting_leadsTo";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   762
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   763
Goalw [projecting_def]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   764
     "projecting (%G. reachable (extend h F Join G)) h F \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   765
\                (extend_set h A LeadsTo[givenBy f] extend_set h B) \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   766
\                (A LeadsTo B)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   767
by (force_tac (claset() addDs [project_LeadsETo_I], simpset()) 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   768
qed "projecting_LeadsTo";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8067
diff changeset
   769
8044
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   770
Close_locale "Extend";
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   771
296b03b79505 new generalized leads-to theory
paulson
parents:
diff changeset
   772