src/HOL/UNITY/WFair.ML
author paulson
Thu, 13 Jan 2000 17:30:23 +0100
changeset 8122 b43ad07660b9
parent 8112 efbe50e2bef9
child 8251 9be357df93d4
permissions -rw-r--r--
working version, with Alloc now working on the same state space as the whole system. Partial removal of ELT.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/WFair
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
Weak Fairness versions of transient, ensures, leadsTo.
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming", 1994
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    12
overload_1st_set "WFair.transient";
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    13
overload_1st_set "WFair.ensures";
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
    14
overload_1st_set "WFair.op leadsTo";
5340
d75c03cf77b5 Misc changes
paulson
parents: 5277
diff changeset
    15
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
(*** transient ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    18
Goalw [stable_def, constrains_def, transient_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    19
    "[| F : stable A; F : transient A |] ==> A = {}";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    21
qed "stable_transient_empty";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    22
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    23
Goalw [transient_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    24
    "[| F : transient A; B<=A |] ==> F : transient B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
by (Clarify_tac 1);
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5971
diff changeset
    26
by (blast_tac (claset() addSIs [rev_bexI]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    27
qed "transient_strengthen";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    28
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    29
Goalw [transient_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    30
    "[| act: Acts F;  A <= Domain act;  act^^A <= -A |] ==> F : transient A";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
by (Blast_tac 1);
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    32
qed "transientI";
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    33
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    34
val major::prems = 
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    35
Goalw [transient_def]
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    36
    "[| F : transient A;  \
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    37
\       !!act. [| act: Acts F;  A <= Domain act;  act^^A <= -A |] ==> P |] \
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    38
\    ==> P";
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    39
by (rtac (major RS CollectD RS bexE) 1);
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    40
by (blast_tac (claset() addIs prems) 1);
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    41
qed "transientE";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    43
Goalw [transient_def] "transient UNIV = {}";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    44
by Auto_tac;
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    45
qed "transient_UNIV";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    46
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    47
Goalw [transient_def] "transient {} = UNIV";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    48
by Auto_tac;
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    49
qed "transient_empty";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    50
Addsimps [transient_UNIV, transient_empty];
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7594
diff changeset
    51
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    52
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    53
(*** ensures ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    55
Goalw [ensures_def]
7524
paulson
parents: 6801
diff changeset
    56
    "[| F : (A-B) co (A Un B); F : transient (A-B) |] ==> F : A ensures B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    58
qed "ensuresI";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    60
Goalw [ensures_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
    61
    "F : A ensures B ==> F : (A-B) co (A Un B) & F : transient (A-B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    62
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    63
qed "ensuresD";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    64
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    65
Goalw [ensures_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
    66
    "[| F : A ensures A'; A'<=B' |] ==> F : A ensures B'";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    67
by (blast_tac (claset() addIs [constrains_weaken, transient_strengthen]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    68
qed "ensures_weaken_R";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    69
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    70
(*The L-version (precondition strengthening) fails, but we have this*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    71
Goalw [ensures_def]
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    72
    "[| F : stable C;  F : A ensures B |]   \
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    73
\   ==> F : (C Int A) ensures (C Int B)";
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    74
by (auto_tac (claset(),
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    75
	      simpset() addsimps [ensures_def,
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    76
				  Int_Un_distrib RS sym,
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    77
				  Diff_Int_distrib RS sym]));
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    78
by (blast_tac (claset() addIs [transient_strengthen]) 2);
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
    79
by (blast_tac (claset() addIs [stable_constrains_Int, constrains_weaken]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    80
qed "stable_ensures_Int";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    81
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
    82
Goal "[| F : stable A;  F : transient C;  A <= B Un C |] ==> F : A ensures B";
5640
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
    83
by (asm_full_simp_tac (simpset() addsimps [ensures_def, stable_def]) 1);
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
    84
by (blast_tac (claset() addIs [constrains_weaken, transient_strengthen]) 1);
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
    85
qed "stable_transient_ensures";
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
    86
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8112
diff changeset
    87
Goal "(A ensures B) = (A unless B) Int transient (A-B)";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8112
diff changeset
    88
by (simp_tac (simpset() addsimps [ensures_def, unless_def]) 1);
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8112
diff changeset
    89
qed "ensures_eq";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8112
diff changeset
    90
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    91
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    92
(*** leadsTo ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    93
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
    94
Goalw [leadsTo_def] "F : A ensures B ==> F : A leadsTo B";
6801
9e0037839d63 renamed the underlying relation of leadsTo from "leadsto"
paulson
parents: 6714
diff changeset
    95
by (blast_tac (claset() addIs [leads.Basis]) 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    96
qed "leadsTo_Basis";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    97
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
    98
Goalw [leadsTo_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
    99
     "[| F : A leadsTo B;  F : B leadsTo C |] ==> F : A leadsTo C";
6801
9e0037839d63 renamed the underlying relation of leadsTo from "leadsto"
paulson
parents: 6714
diff changeset
   100
by (blast_tac (claset() addIs [leads.Trans]) 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   101
qed "leadsTo_Trans";
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   102
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   103
Goal "F : transient A ==> F : A leadsTo (-A)";
5640
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
   104
by (asm_simp_tac 
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
   105
    (simpset() addsimps [leadsTo_Basis, ensuresI, Compl_partition]) 1);
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
   106
qed "transient_imp_leadsTo";
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
   107
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   108
(*Useful with cancellation, disjunction*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   109
Goal "F : A leadsTo (A' Un A') ==> F : A leadsTo A'";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   110
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   111
qed "leadsTo_Un_duplicate";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   112
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   113
Goal "F : A leadsTo (A' Un C Un C) ==> F : A leadsTo (A' Un C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   114
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   115
qed "leadsTo_Un_duplicate2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   116
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   117
(*The Union introduction rule as we should have liked to state it*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   118
val prems = Goalw [leadsTo_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   119
    "(!!A. A : S ==> F : A leadsTo B) ==> F : (Union S) leadsTo B";
6801
9e0037839d63 renamed the underlying relation of leadsTo from "leadsto"
paulson
parents: 6714
diff changeset
   120
by (blast_tac (claset() addIs [leads.Union] addDs prems) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   121
qed "leadsTo_Union";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   122
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   123
val prems = Goalw [leadsTo_def]
7524
paulson
parents: 6801
diff changeset
   124
 "(!!A. A : S ==> F : (A Int C) leadsTo B) ==> F : (Union S Int C) leadsTo B";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   125
by (simp_tac (HOL_ss addsimps [Int_Union_Union]) 1);
6801
9e0037839d63 renamed the underlying relation of leadsTo from "leadsto"
paulson
parents: 6714
diff changeset
   126
by (blast_tac (claset() addIs [leads.Union] addDs prems) 1);
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   127
qed "leadsTo_Union_Int";
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   128
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   129
val prems = Goal
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   130
    "(!!i. i : I ==> F : (A i) leadsTo B) ==> F : (UN i:I. A i) leadsTo B";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   131
by (stac (Union_image_eq RS sym) 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   132
by (blast_tac (claset() addIs leadsTo_Union::prems) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   133
qed "leadsTo_UN";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   134
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   135
(*Binary union introduction rule*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   136
Goal "[| F : A leadsTo C; F : B leadsTo C |] ==> F : (A Un B) leadsTo C";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   137
by (stac Un_eq_Union 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   138
by (blast_tac (claset() addIs [leadsTo_Union]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   139
qed "leadsTo_Un";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   140
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   141
val prems = 
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   142
Goal "(!!x. x : A ==> F : {x} leadsTo B) ==> F : A leadsTo B";
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   143
by (stac (UN_singleton RS sym) 1 THEN rtac leadsTo_UN 1);
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   144
by (blast_tac (claset() addIs prems) 1);
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   145
qed "single_leadsTo_I";
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   146
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   147
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   148
(*The INDUCTION rule as we should have liked to state it*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   149
val major::prems = Goalw [leadsTo_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   150
  "[| F : za leadsTo zb;  \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   151
\     !!A B. F : A ensures B ==> P A B; \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   152
\     !!A B C. [| F : A leadsTo B; P A B; F : B leadsTo C; P B C |] \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   153
\              ==> P A C; \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   154
\     !!B S. ALL A:S. F : A leadsTo B & P A B ==> P (Union S) B \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   155
\  |] ==> P za zb";
6801
9e0037839d63 renamed the underlying relation of leadsTo from "leadsto"
paulson
parents: 6714
diff changeset
   156
by (rtac (major RS CollectD RS leads.induct) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   157
by (REPEAT (blast_tac (claset() addIs prems) 1));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   158
qed "leadsTo_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   159
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   160
7594
8a188ef6545e working version with co-guarantees-leadsto results
paulson
parents: 7524
diff changeset
   161
Goal "A<=B ==> F : A ensures B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   162
by (rewrite_goals_tac [ensures_def, constrains_def, transient_def]);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   163
by (Blast_tac 1);
7594
8a188ef6545e working version with co-guarantees-leadsto results
paulson
parents: 7524
diff changeset
   164
qed "subset_imp_ensures";
8a188ef6545e working version with co-guarantees-leadsto results
paulson
parents: 7524
diff changeset
   165
8a188ef6545e working version with co-guarantees-leadsto results
paulson
parents: 7524
diff changeset
   166
bind_thm ("subset_imp_leadsTo", subset_imp_ensures RS leadsTo_Basis);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   167
8112
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   168
bind_thm ("leadsTo_refl", subset_refl RS subset_imp_leadsTo);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   169
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   170
bind_thm ("empty_leadsTo", empty_subsetI RS subset_imp_leadsTo);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   171
Addsimps [empty_leadsTo];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   172
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
   173
bind_thm ("leadsTo_UNIV", subset_UNIV RS subset_imp_leadsTo);
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
   174
Addsimps [leadsTo_UNIV];
e3237d8c18d6 working version with new theory ELT
paulson
parents: 7963
diff changeset
   175
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   176
8112
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   177
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   178
(** Variant induction rule: on the preconditions for B **)
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   179
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   180
(*Lemma is the weak version: can't see how to do it in one step*)
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   181
val major::prems = Goal
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   182
  "[| F : za leadsTo zb;  \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   183
\     P zb; \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   184
\     !!A B. [| F : A ensures B;  P B |] ==> P A; \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   185
\     !!S. ALL A:S. P A ==> P (Union S) \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   186
\  |] ==> P za";
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   187
(*by induction on this formula*)
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   188
by (subgoal_tac "P zb --> P za" 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   189
(*now solve first subgoal: this formula is sufficient*)
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   190
by (blast_tac (claset() addIs leadsTo_refl::prems) 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   191
by (rtac (major RS leadsTo_induct) 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   192
by (REPEAT (blast_tac (claset() addIs prems) 1));
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   193
val lemma = result();
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   194
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   195
val major::prems = Goal
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   196
  "[| F : za leadsTo zb;  \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   197
\     P zb; \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   198
\     !!A B. [| F : A ensures B;  F : B leadsTo zb;  P B |] ==> P A; \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   199
\     !!S. ALL A:S. F : A leadsTo zb & P A ==> P (Union S) \
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   200
\  |] ==> P za";
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   201
by (subgoal_tac "F : za leadsTo zb & P za" 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   202
by (etac conjunct2 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   203
by (rtac (major RS lemma) 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   204
by (blast_tac (claset() addIs [leadsTo_Union]@prems) 3);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   205
by (blast_tac (claset() addIs [leadsTo_Basis,leadsTo_Trans]@prems) 2);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   206
by (blast_tac (claset() addIs [leadsTo_refl]@prems) 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   207
qed "leadsTo_induct_pre";
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   208
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   209
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   210
Goal "[| F : A leadsTo A'; A'<=B' |] ==> F : A leadsTo B'";
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   211
by (blast_tac (claset() addIs [subset_imp_leadsTo, leadsTo_Trans]) 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   212
qed "leadsTo_weaken_R";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   213
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   214
Goal "[| F : A leadsTo A'; B<=A |] ==> F : B leadsTo A'";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   215
by (blast_tac (claset() addIs [leadsTo_Trans, subset_imp_leadsTo]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   216
qed_spec_mp "leadsTo_weaken_L";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   217
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   218
(*Distributes over binary unions*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   219
Goal "F : (A Un B) leadsTo C  =  (F : A leadsTo C & F : B leadsTo C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   220
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   221
qed "leadsTo_Un_distrib";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   222
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   223
Goal "F : (UN i:I. A i) leadsTo B  =  (ALL i : I. F : (A i) leadsTo B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   224
by (blast_tac (claset() addIs [leadsTo_UN, leadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   225
qed "leadsTo_UN_distrib";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   226
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   227
Goal "F : (Union S) leadsTo B  =  (ALL A : S. F : A leadsTo B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   228
by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   229
qed "leadsTo_Union_distrib";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   230
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   231
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   232
Goal "[| F : A leadsTo A'; B<=A; A'<=B' |] ==> F : B leadsTo B'";
5340
d75c03cf77b5 Misc changes
paulson
parents: 5277
diff changeset
   233
by (blast_tac (claset() addIs [leadsTo_weaken_R, leadsTo_weaken_L,
d75c03cf77b5 Misc changes
paulson
parents: 5277
diff changeset
   234
			       leadsTo_Trans]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   235
qed "leadsTo_weaken";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   236
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   237
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   238
(*Set difference: maybe combine with leadsTo_weaken_L??*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   239
Goal "[| F : (A-B) leadsTo C; F : B leadsTo C |]   ==> F : A leadsTo C";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   240
by (blast_tac (claset() addIs [leadsTo_Un, leadsTo_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   241
qed "leadsTo_Diff";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   242
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   243
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   244
(** Meta or object quantifier ???
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   245
    see ball_constrains_UN in UNITY.ML***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   246
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   247
val prems = goal thy
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   248
   "(!! i. i:I ==> F : (A i) leadsTo (A' i)) \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   249
\   ==> F : (UN i:I. A i) leadsTo (UN i:I. A' i)";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   250
by (simp_tac (HOL_ss addsimps [Union_image_eq RS sym]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   251
by (blast_tac (claset() addIs [leadsTo_Union, leadsTo_weaken_R] 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   252
                        addIs prems) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   253
qed "leadsTo_UN_UN";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   254
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   255
(*Binary union version*)
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   256
Goal "[| F : A leadsTo A'; F : B leadsTo B' |] \
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   257
\     ==> F : (A Un B) leadsTo (A' Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   258
by (blast_tac (claset() addIs [leadsTo_Un, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   259
			       leadsTo_weaken_R]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   260
qed "leadsTo_Un_Un";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   261
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   262
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   263
(** The cancellation law **)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   264
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   265
Goal "[| F : A leadsTo (A' Un B); F : B leadsTo B' |] \
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   266
\     ==> F : A leadsTo (A' Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   267
by (blast_tac (claset() addIs [leadsTo_Un_Un, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   268
			       subset_imp_leadsTo, leadsTo_Trans]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   269
qed "leadsTo_cancel2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   270
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   271
Goal "[| F : A leadsTo (A' Un B); F : (B-A') leadsTo B' |] \
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   272
\     ==> F : A leadsTo (A' Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   273
by (rtac leadsTo_cancel2 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   274
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   275
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   276
qed "leadsTo_cancel_Diff2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   277
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   278
Goal "[| F : A leadsTo (B Un A'); F : B leadsTo B' |] \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   279
\   ==> F : A leadsTo (B' Un A')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   280
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   281
by (blast_tac (claset() addSIs [leadsTo_cancel2]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   282
qed "leadsTo_cancel1";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   283
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   284
Goal "[| F : A leadsTo (B Un A'); F : (B-A') leadsTo B' |] \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   285
\   ==> F : A leadsTo (B' Un A')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   286
by (rtac leadsTo_cancel1 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   287
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   288
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   289
qed "leadsTo_cancel_Diff1";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   290
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   291
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   292
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   293
(** The impossibility law **)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   294
8112
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   295
Goal "F : A leadsTo {} ==> A={}";
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   296
by (etac leadsTo_induct_pre 1);
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   297
by (ALLGOALS
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   298
    (asm_full_simp_tac
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   299
     (simpset() addsimps [ensures_def, constrains_def, transient_def])));
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   300
by (Blast_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   301
qed "leadsTo_empty";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   302
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   303
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   304
(** PSP: Progress-Safety-Progress **)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   305
5640
4a59d99b5ca3 new theorems
paulson
parents: 5620
diff changeset
   306
(*Special case of PSP: Misra's "stable conjunction"*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   307
Goalw [stable_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   308
   "[| F : A leadsTo A'; F : stable B |] \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   309
\   ==> F : (A Int B) leadsTo (A' Int B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   310
by (etac leadsTo_induct 1);
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   311
by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   312
by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   313
by (rtac leadsTo_Basis 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   314
by (asm_full_simp_tac
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   315
    (simpset() addsimps [ensures_def, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   316
			 Diff_Int_distrib2 RS sym, Int_Un_distrib2 RS sym]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   317
by (blast_tac (claset() addIs [transient_strengthen, constrains_Int]) 1);
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   318
qed "psp_stable";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   319
7524
paulson
parents: 6801
diff changeset
   320
Goal
paulson
parents: 6801
diff changeset
   321
   "[| F : A leadsTo A'; F : stable B |] ==> F : (B Int A) leadsTo (B Int A')";
5536
130f3d891fb2 tidying and deleting needless parentheses
paulson
parents: 5521
diff changeset
   322
by (asm_simp_tac (simpset() addsimps psp_stable::Int_ac) 1);
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   323
qed "psp_stable2";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   324
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   325
Goalw [ensures_def, constrains_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   326
   "[| F : A ensures A'; F : B co B' |] \
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   327
\   ==> F : (A Int B') ensures ((A' Int B) Un (B' - B))";
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   328
by (Clarify_tac 1);  (*speeds up the proof*)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   329
by (blast_tac (claset() addIs [transient_strengthen]) 1);
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   330
qed "psp_ensures";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   331
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   332
Goal "[| F : A leadsTo A'; F : B co B' |] \
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   333
\     ==> F : (A Int B') leadsTo ((A' Int B) Un (B' - B))";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   334
by (etac leadsTo_induct 1);
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6012
diff changeset
   335
by (blast_tac (claset() addIs [leadsTo_Union_Int]) 3);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   336
(*Transitivity case has a delicate argument involving "cancellation"*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   337
by (rtac leadsTo_Un_duplicate2 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   338
by (etac leadsTo_cancel_Diff1 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   339
by (asm_full_simp_tac (simpset() addsimps [Int_Diff, Diff_triv]) 2);
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   340
by (blast_tac (claset() addIs [leadsTo_weaken_L] 
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   341
                        addDs [constrains_imp_subset]) 2);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   342
(*Basis case*)
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   343
by (blast_tac (claset() addIs [leadsTo_Basis, psp_ensures]) 1);
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   344
qed "psp";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   345
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   346
Goal "[| F : A leadsTo A'; F : B co B' |] \
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   347
\   ==> F : (B' Int A) leadsTo ((B Int A') Un (B' - B))";
5536
130f3d891fb2 tidying and deleting needless parentheses
paulson
parents: 5521
diff changeset
   348
by (asm_simp_tac (simpset() addsimps psp::Int_ac) 1);
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   349
qed "psp2";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   350
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   351
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   352
Goalw [unless_def]
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   353
   "[| F : A leadsTo A';  F : B unless B' |] \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   354
\   ==> F : (A Int B) leadsTo ((A' Int B) Un B')";
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   355
by (dtac psp 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   356
by (assume_tac 1);
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   357
by (blast_tac (claset() addIs [leadsTo_weaken]) 1);
5277
e4297d03e5d2 A higher-level treatment of LeadsTo, minimizing use of "reachable"
paulson
parents: 5257
diff changeset
   358
qed "psp_unless";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   359
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   360
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   361
(*** Proving the induction rules ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   362
5257
c03e3ba9cbcc Indentation, comments
paulson
parents: 5253
diff changeset
   363
(** The most general rule: r is any wf relation; f is any variant function **)
c03e3ba9cbcc Indentation, comments
paulson
parents: 5253
diff changeset
   364
5239
2fd94efb9d64 Tidying
paulson
parents: 5232
diff changeset
   365
Goal "[| wf r;     \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   366
\        ALL m. F : (A Int f-``{m}) leadsTo                     \
7524
paulson
parents: 6801
diff changeset
   367
\                   ((A Int f-``(r^-1 ^^ {m})) Un B) |] \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   368
\     ==> F : (A Int f-``{m}) leadsTo B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   369
by (eres_inst_tac [("a","m")] wf_induct 1);
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   370
by (subgoal_tac "F : (A Int (f -`` (r^-1 ^^ {x}))) leadsTo B" 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   371
by (stac vimage_eq_UN 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   372
by (asm_simp_tac (HOL_ss addsimps (UN_simps RL [sym])) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   373
by (blast_tac (claset() addIs [leadsTo_UN]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   374
by (blast_tac (claset() addIs [leadsTo_cancel1, leadsTo_Un_duplicate]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   375
val lemma = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   376
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   377
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   378
(** Meta or object quantifier ????? **)
5239
2fd94efb9d64 Tidying
paulson
parents: 5232
diff changeset
   379
Goal "[| wf r;     \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   380
\        ALL m. F : (A Int f-``{m}) leadsTo                     \
7524
paulson
parents: 6801
diff changeset
   381
\                   ((A Int f-``(r^-1 ^^ {m})) Un B) |] \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   382
\     ==> F : A leadsTo B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   383
by (res_inst_tac [("t", "A")] subst 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   384
by (rtac leadsTo_UN 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   385
by (etac lemma 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   386
by (REPEAT (assume_tac 2));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   387
by (Fast_tac 1);    (*Blast_tac: Function unknown's argument not a parameter*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   388
qed "leadsTo_wf_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   389
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   390
5239
2fd94efb9d64 Tidying
paulson
parents: 5232
diff changeset
   391
Goal "[| wf r;     \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   392
\        ALL m:I. F : (A Int f-``{m}) leadsTo                   \
7524
paulson
parents: 6801
diff changeset
   393
\                     ((A Int f-``(r^-1 ^^ {m})) Un B) |] \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   394
\     ==> F : A leadsTo ((A - (f-``I)) Un B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   395
by (etac leadsTo_wf_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   396
by Safe_tac;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   397
by (case_tac "m:I" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   398
by (blast_tac (claset() addIs [leadsTo_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   399
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   400
qed "bounded_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   401
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   402
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   403
(*Alternative proof is via the lemma F : (A Int f-``(lessThan m)) leadsTo B*)
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   404
Goal "[| ALL m. F : (A Int f-``{m}) leadsTo                     \
7524
paulson
parents: 6801
diff changeset
   405
\                   ((A Int f-``(lessThan m)) Un B) |] \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   406
\     ==> F : A leadsTo B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   407
by (rtac (wf_less_than RS leadsTo_wf_induct) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   408
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   409
qed "lessThan_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   410
7524
paulson
parents: 6801
diff changeset
   411
Goal "[| ALL m:(greaterThan l).    \
paulson
parents: 6801
diff changeset
   412
\           F : (A Int f-``{m}) leadsTo ((A Int f-``(lessThan m)) Un B) |] \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   413
\     ==> F : A leadsTo ((A Int (f-``(atMost l))) Un B)";
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   414
by (simp_tac (HOL_ss addsimps [Diff_eq RS sym, vimage_Compl, 
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   415
			       Compl_greaterThan RS sym]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   416
by (rtac (wf_less_than RS bounded_induct) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   417
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   418
qed "lessThan_bounded_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   419
7524
paulson
parents: 6801
diff changeset
   420
Goal "[| ALL m:(lessThan l).    \
paulson
parents: 6801
diff changeset
   421
\           F : (A Int f-``{m}) leadsTo ((A Int f-``(greaterThan m)) Un B) |] \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   422
\     ==> F : A leadsTo ((A Int (f-``(atLeast l))) Un B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   423
by (res_inst_tac [("f","f"),("f1", "%k. l - k")]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   424
    (wf_less_than RS wf_inv_image RS leadsTo_wf_induct) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   425
by (simp_tac (simpset() addsimps [inv_image_def, Image_singleton]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   426
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   427
by (case_tac "m<l" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   428
by (blast_tac (claset() addIs [not_leE, subset_imp_leadsTo]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   429
by (blast_tac (claset() addIs [leadsTo_weaken_R, diff_less_mono2]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   430
qed "greaterThan_bounded_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   431
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   432
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   433
(*** wlt ****)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   434
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   435
(*Misra's property W3*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   436
Goalw [wlt_def] "F : (wlt F B) leadsTo B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   437
by (blast_tac (claset() addSIs [leadsTo_Union]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   438
qed "wlt_leadsTo";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   439
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   440
Goalw [wlt_def] "F : A leadsTo B ==> A <= wlt F B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   441
by (blast_tac (claset() addSIs [leadsTo_Union]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   442
qed "leadsTo_subset";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   443
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   444
(*Misra's property W2*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   445
Goal "F : A leadsTo B = (A <= wlt F B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   446
by (blast_tac (claset() addSIs [leadsTo_subset, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   447
				wlt_leadsTo RS leadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   448
qed "leadsTo_eq_subset_wlt";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   449
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   450
(*Misra's property W4*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   451
Goal "B <= wlt F B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   452
by (asm_simp_tac (simpset() addsimps [leadsTo_eq_subset_wlt RS sym,
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   453
				      subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   454
qed "wlt_increasing";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   455
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   456
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   457
(*Used in the Trans case below*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   458
Goalw [constrains_def]
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   459
   "[| B <= A2;  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   460
\      F : (A1 - B) co (A1 Un B); \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   461
\      F : (A2 - C) co (A2 Un C) |] \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   462
\   ==> F : (A1 Un A2 - C) co (A1 Un A2 Un C)";
5669
f5d9caafc3bd added Clarify_tac to speed up proofs
paulson
parents: 5648
diff changeset
   463
by (Clarify_tac 1);
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5608
diff changeset
   464
by (Blast_tac 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   465
val lemma1 = result();
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   466
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   467
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   468
(*Lemma (1,2,3) of Misra's draft book, Chapter 4, "Progress"*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   469
Goal "F : A leadsTo A' ==> \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   470
\     EX B. A<=B & F : B leadsTo A' & F : (B-A') co (B Un A')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   471
by (etac leadsTo_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   472
(*Basis*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   473
by (blast_tac (claset() addIs [leadsTo_Basis]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   474
                        addDs [ensuresD]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   475
(*Trans*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   476
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   477
by (res_inst_tac [("x", "Ba Un Bb")] exI 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   478
by (blast_tac (claset() addIs [lemma1, leadsTo_Un_Un, leadsTo_cancel1,
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   479
			       leadsTo_Un_duplicate]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   480
(*Union*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   481
by (clarify_tac (claset() addSDs [ball_conj_distrib RS iffD1,
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   482
				  bchoice, ball_constrains_UN]) 1);;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   483
by (res_inst_tac [("x", "UN A:S. f A")] exI 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   484
by (blast_tac (claset() addIs [leadsTo_UN, constrains_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   485
qed "leadsTo_123";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   486
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   487
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   488
(*Misra's property W5*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   489
Goal "F : (wlt F B - B) co (wlt F B)";
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   490
by (cut_inst_tac [("F","F")] (wlt_leadsTo RS leadsTo_123) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   491
by (Clarify_tac 1);
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   492
by (subgoal_tac "Ba = wlt F B" 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   493
by (blast_tac (claset() addDs [leadsTo_eq_subset_wlt RS iffD1]) 2);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   494
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   495
by (asm_full_simp_tac (simpset() addsimps [wlt_increasing, Un_absorb2]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   496
qed "wlt_constrains_wlt";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   497
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   498
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   499
(*** Completion: Binary and General Finite versions ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   500
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5640
diff changeset
   501
Goal "[| W = wlt F (B' Un C);     \
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   502
\      F : A leadsTo (A' Un C);  F : A' co (A' Un C);   \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   503
\      F : B leadsTo (B' Un C);  F : B' co (B' Un C) |] \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   504
\   ==> F : (A Int B) leadsTo ((A' Int B') Un C)";
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   505
by (subgoal_tac "F : (W-C) co (W Un B' Un C)" 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   506
by (blast_tac (claset() addIs [[asm_rl, wlt_constrains_wlt] 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   507
			       MRS constrains_Un RS constrains_weaken]) 2);
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   508
by (subgoal_tac "F : (W-C) co W" 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   509
by (asm_full_simp_tac 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   510
    (simpset() addsimps [wlt_increasing, Un_assoc, Un_absorb2]) 2);
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   511
by (subgoal_tac "F : (A Int W - C) leadsTo (A' Int W Un C)" 1);
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   512
by (blast_tac (claset() addIs [wlt_leadsTo, psp RS leadsTo_weaken]) 2);
7963
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   513
(** LEVEL 6 **)
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   514
by (subgoal_tac "F : (A' Int W Un C) leadsTo (A' Int B' Un C)" 1);
6714
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   515
by (rtac leadsTo_Un_duplicate2 2);
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   516
by (blast_tac (claset() addIs [leadsTo_Un_Un, 
6b2b4ec58178 new rule single_leadsTo_I; stronger PSP rule
paulson
parents: 6564
diff changeset
   517
                               wlt_leadsTo RS psp2 RS leadsTo_weaken, 
8112
efbe50e2bef9 new theorem leadsTo_refl and induction rule leadsTo_induct_pre
paulson
parents: 8041
diff changeset
   518
                               leadsTo_refl]) 2);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   519
by (dtac leadsTo_Diff 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   520
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   521
by (subgoal_tac "A Int B <= A Int W" 1);
5456
paulson
parents: 5340
diff changeset
   522
by (blast_tac (claset() addSDs [leadsTo_subset]
paulson
parents: 5340
diff changeset
   523
			addSIs [subset_refl RS Int_mono]) 2);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   524
by (blast_tac (claset() addIs [leadsTo_Trans, subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   525
bind_thm("completion", refl RS result());
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   526
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   527
6536
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   528
Goal "finite I ==> (ALL i:I. F : (A i) leadsTo (A' i Un C)) -->  \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   529
\                  (ALL i:I. F : (A' i) co (A' i Un C)) --> \
281d44905cab made many specification operators infix
paulson
parents: 6295
diff changeset
   530
\                  F : (INT i:I. A i) leadsTo ((INT i:I. A' i) Un C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   531
by (etac finite_induct 1);
7963
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   532
by Auto_tac;
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   533
by (dtac ball_constrains_INT 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   534
by (asm_full_simp_tac (simpset() addsimps [completion]) 1); 
6564
paulson
parents: 6536
diff changeset
   535
qed_spec_mp "finite_completion";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   536
7963
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   537
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   538
Goalw [stable_def]
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   539
     "[| F : A leadsTo A';  F : stable A';   \
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   540
\        F : B leadsTo B';  F : stable B' |] \
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   541
\   ==> F : (A Int B) leadsTo (A' Int B')";
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   542
by (res_inst_tac [("C1", "{}")] (completion RS leadsTo_weaken_R) 1);
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   543
by (REPEAT (Force_tac 1));
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   544
qed "stable_completion";
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   545
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   546
Goalw [stable_def]
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   547
     "[| finite I;  \
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   548
\        ALL i:I. F : (A i) leadsTo (A' i);  \
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   549
\        ALL i:I. F : stable (A' i) |]       \
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   550
\     ==> F : (INT i:I. A i) leadsTo (INT i:I. A' i)";
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   551
by (res_inst_tac [("C1", "{}")] (finite_completion RS leadsTo_weaken_R) 1);
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   552
by (REPEAT (Force_tac 1));
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   553
qed_spec_mp "finite_stable_completion";
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   554
e7beff82e1ba simplified the stable_completion proofs
paulson
parents: 7826
diff changeset
   555