src/HOL/UNITY/PPROD.ML
author paulson
Mon, 16 Nov 1998 13:58:56 +0100
changeset 5899 13d4753079fe
child 5972 2430ccbde87d
permissions -rw-r--r--
new theory PPROD
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/PPROD.ML
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     2
    ID:         $Id$
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     5
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     6
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     7
val rinst = read_instantiate_sg (sign_of thy);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     8
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
     9
(*** General lemmas ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    10
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    11
Goal "x:C ==> (A Times C <= B Times C) = (A <= B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    12
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    13
qed "Times_subset_cancel2";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    14
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    15
Goal "x:C ==> (A Times C = B Times C) = (A = B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    16
by (blast_tac (claset() addEs [equalityE]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    17
qed "Times_eq_cancel2";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    18
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    19
Goal "Union(B) Times A = (UN C:B. C Times A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    20
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    21
qed "Times_Union2";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    22
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    23
Goal "Domain (Union S) = (UN A:S. Domain A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    24
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    25
qed "Domain_Union";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    26
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    27
(** RTimes: the product of two relations **)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    28
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    29
Goal "(((a,b), (a',b')) : A RTimes B) = ((a,a'):A & (b,b'):B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    30
by (simp_tac (simpset() addsimps [RTimes_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    31
qed "mem_RTimes_iff";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    32
AddIffs [mem_RTimes_iff]; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    33
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    34
Goalw [RTimes_def] "[| A<=C;  B<=D |] ==> A RTimes B <= C RTimes D";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    35
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    36
qed "RTimes_mono";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    37
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    38
Goal "{} RTimes B = {}";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    39
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    40
qed "RTimes_empty1"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    41
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    42
Goal "A RTimes {} = {}";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    43
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    44
qed "RTimes_empty2"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    45
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    46
Goal "Id RTimes Id = Id";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    47
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    48
qed "RTimes_Id"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    49
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    50
Addsimps [RTimes_empty1, RTimes_empty2, RTimes_Id]; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    51
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    52
Goal "Domain (r RTimes s) = Domain r Times Domain s";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    53
by (auto_tac (claset(), simpset() addsimps [Domain_iff]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    54
qed "Domain_RTimes"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    55
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    56
Goal "Range (r RTimes s) = Range r Times Range s";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    57
by (auto_tac (claset(), simpset() addsimps [Range_iff]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    58
qed "Range_RTimes"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    59
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    60
Goal "(r RTimes s) ^^ (A Times B) = r^^A Times s^^B";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    61
by (auto_tac (claset(), simpset() addsimps [Image_iff]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    62
qed "Image_RTimes"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    63
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    64
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    65
Goal "- (UNIV Times A) = UNIV Times (-A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    66
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    67
qed "Compl_Times_UNIV1"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    68
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    69
Goal "- (A Times UNIV) = (-A) Times UNIV";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    70
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    71
qed "Compl_Times_UNIV2"; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    72
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    73
Addsimps [Compl_Times_UNIV1, Compl_Times_UNIV2]; 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    74
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    75
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    76
(**** Lcopy ****)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    77
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    78
(*** Basic properties ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    79
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    80
Goal "Init (Lcopy F) = Init F Times UNIV";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    81
by (simp_tac (simpset() addsimps [Lcopy_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    82
qed "Init_Lcopy";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    83
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    84
Goal "Id : (%act. act RTimes Id) `` Acts F";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    85
by (rtac image_eqI 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    86
by (rtac Id_in_Acts 2);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    87
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    88
val lemma = result();
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    89
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    90
Goal "Acts (Lcopy F) = (%act. act RTimes Id) `` Acts F";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    91
by (auto_tac (claset() addSIs [lemma], 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    92
	      simpset() addsimps [Lcopy_def]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    93
qed "Acts_Lcopy";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    94
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    95
Addsimps [Init_Lcopy];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    96
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    97
Goalw [Lcopy_def, SKIP_def] "Lcopy SKIP = SKIP";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    98
by (rtac program_equalityI 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    99
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   100
qed "Lcopy_SKIP";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   101
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   102
Addsimps [Lcopy_SKIP];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   103
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   104
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   105
(*** Safety: constrains, stable ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   106
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   107
(** In most cases, C = UNIV.  The generalization isn't of obvious value. **)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   108
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   109
Goal "x: C ==> \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   110
\     (Lcopy F : constrains (A Times C) (B Times C)) = (F : constrains A B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   111
by (auto_tac (claset(), simpset() addsimps [constrains_def, Lcopy_def]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   112
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   113
qed "Lcopy_constrains";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   114
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   115
Goal "Lcopy F : constrains A B ==> F : constrains (Domain A) (Domain B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   116
by (auto_tac (claset(), simpset() addsimps [constrains_def, Lcopy_def]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   117
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   118
qed "Lcopy_constrains_Domain";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   119
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   120
Goal "x: C ==> (Lcopy F : stable (A Times C)) = (F : stable A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   121
by (asm_simp_tac (simpset() addsimps [stable_def, Lcopy_constrains]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   122
qed "Lcopy_stable";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   123
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   124
Goal "(Lcopy F : invariant (A Times UNIV)) = (F : invariant A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   125
by (asm_simp_tac (simpset() addsimps [Times_subset_cancel2,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   126
				      invariant_def, Lcopy_stable]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   127
qed "Lcopy_invariant";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   128
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   129
(** Substitution Axiom versions: Constrains, Stable **)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   130
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   131
Goal "p : reachable (Lcopy F) ==> fst p : reachable F";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   132
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   133
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   134
    (claset() addIs reachable.intrs,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   135
     simpset() addsimps [Acts_Lcopy]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   136
qed "reachable_Lcopy_fst";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   137
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   138
Goal "(a,b) : reachable (Lcopy F) ==> a : reachable F";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   139
by (force_tac (claset() addSDs [reachable_Lcopy_fst], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   140
qed "reachable_LcopyD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   141
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   142
Goal "reachable (Lcopy F) = reachable F Times UNIV";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   143
by (rtac equalityI 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   144
by (force_tac (claset() addSDs [reachable_LcopyD], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   145
by (Clarify_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   146
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   147
by (ALLGOALS (force_tac (claset() addIs reachable.intrs, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   148
			 simpset() addsimps [Acts_Lcopy])));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   149
qed "reachable_Lcopy_eq";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   150
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   151
Goal "(Lcopy F : Constrains (A Times UNIV) (B Times UNIV)) =  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   152
\     (F : Constrains A B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   153
by (simp_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   154
    (simpset() addsimps [Constrains_def, reachable_Lcopy_eq, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   155
			 Lcopy_constrains, Sigma_Int_distrib1 RS sym]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   156
qed "Lcopy_Constrains";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   157
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   158
Goal "(Lcopy F : Stable (A Times UNIV)) = (F : Stable A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   159
by (simp_tac (simpset() addsimps [Stable_def, Lcopy_Constrains]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   160
qed "Lcopy_Stable";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   161
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   162
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   163
(*** Progress: transient, ensures ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   164
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   165
Goal "(Lcopy F : transient (A Times UNIV)) = (F : transient A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   166
by (auto_tac (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   167
	      simpset() addsimps [transient_def, Times_subset_cancel2, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   168
				  Domain_RTimes, Image_RTimes, Lcopy_def]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   169
qed "Lcopy_transient";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   170
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   171
Goal "(Lcopy F : ensures (A Times UNIV) (B Times UNIV)) = \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   172
\     (F : ensures A B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   173
by (simp_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   174
    (simpset() addsimps [ensures_def, Lcopy_constrains, Lcopy_transient, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   175
			 Sigma_Un_distrib1 RS sym, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   176
			 Sigma_Diff_distrib1 RS sym]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   177
qed "Lcopy_ensures";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   178
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   179
Goal "F : leadsTo A B ==> Lcopy F : leadsTo (A Times UNIV) (B Times UNIV)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   180
by (etac leadsTo_induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   181
by (asm_simp_tac (simpset() addsimps [leadsTo_UN, Times_Union2]) 3);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   182
by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   183
by (asm_simp_tac (simpset() addsimps [leadsTo_Basis, Lcopy_ensures]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   184
qed "leadsTo_imp_Lcopy_leadsTo";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   185
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   186
Goal "Lcopy F : ensures A B ==> F : ensures (Domain A) (Domain B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   187
by (full_simp_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   188
    (simpset() addsimps [ensures_def, Lcopy_constrains, Lcopy_transient, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   189
			 Domain_Un_eq RS sym,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   190
			 Sigma_Un_distrib1 RS sym, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   191
			 Sigma_Diff_distrib1 RS sym]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   192
by (safe_tac (claset() addSDs [Lcopy_constrains_Domain]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   193
by (etac constrains_weaken_L 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   194
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   195
(*NOT able to prove this:
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   196
Lcopy F : ensures A B ==> F : ensures (Domain A) (Domain B)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   197
 1. [| Lcopy F : transient (A - B);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   198
       F : constrains (Domain (A - B)) (Domain (A Un B)) |]
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   199
    ==> F : transient (Domain A - Domain B)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   200
**)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   201
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   202
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   203
Goal "Lcopy F : leadsTo AU BU ==> F : leadsTo (Domain AU) (Domain BU)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   204
by (etac leadsTo_induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   205
by (full_simp_tac (simpset() addsimps [Domain_Union]) 3);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   206
by (blast_tac (claset() addIs [leadsTo_UN]) 3);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   207
by (blast_tac (claset() addIs [leadsTo_Trans]) 2);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   208
by (rtac leadsTo_Basis 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   209
(*...and so CANNOT PROVE THIS*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   210
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   211
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   212
(*This also seems impossible to prove??*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   213
Goal "(Lcopy F : leadsTo (A Times UNIV) (B Times UNIV)) = \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   214
\     (F : leadsTo A B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   215
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   216
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   217
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   218
(**** PPROD ****)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   219
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   220
(*** Basic properties ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   221
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   222
Goalw [PPROD_def, lift_prog_def]
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   223
     "Init (PPROD I F) = {f. ALL i:I. f i : Init F}";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   224
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   225
qed "Init_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   226
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   227
Goalw [lift_act_def] "lift_act i Id = Id";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   228
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   229
qed "lift_act_Id";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   230
Addsimps [lift_act_Id];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   231
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   232
Goalw [lift_act_def]
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   233
    "((f,f') : lift_act i act) = (EX s'. f' = f(i := s') & (f i, s') : act)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   234
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   235
qed "lift_act_eq";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   236
AddIffs [lift_act_eq];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   237
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   238
Goal "Acts (PPROD I F) = insert Id (UN i:I. lift_act i `` Acts F)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   239
by (auto_tac (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   240
	      simpset() addsimps [PPROD_def, lift_prog_def, Acts_JN]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   241
qed "Acts_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   242
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   243
Addsimps [Init_PPROD];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   244
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   245
Goal "PPROD I SKIP = SKIP";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   246
by (rtac program_equalityI 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   247
by (auto_tac (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   248
	      simpset() addsimps [PPROD_def, lift_prog_def, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   249
				  SKIP_def, Acts_JN]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   250
qed "PPROD_SKIP";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   251
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   252
Goal "PPROD {} F = SKIP";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   253
by (simp_tac (simpset() addsimps [PPROD_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   254
qed "PPROD_empty";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   255
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   256
Addsimps [PPROD_SKIP, PPROD_empty];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   257
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   258
Goalw [PPROD_def]  "PPROD (insert i I) F = (lift_prog i F) Join (PPROD I F)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   259
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   260
qed "PPROD_insert";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   261
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   262
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   263
(*** Safety: constrains, stable ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   264
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   265
val subsetCE' = rinst
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   266
            [("c", "(%u. ?s)(i:=?s')")] subsetCE;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   267
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   268
Goal "i : I ==>  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   269
\     (PPROD I F : constrains {f. f i : A} {f. f i : B})  =  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   270
\     (F : constrains A B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   271
by (auto_tac (claset(), 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   272
	      simpset() addsimps [constrains_def, lift_prog_def, PPROD_def,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   273
				  Acts_JN]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   274
by (REPEAT (Blast_tac 2));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   275
by (force_tac (claset() addSEs [subsetCE', allE, ballE], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   276
qed "PPROD_constrains";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   277
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   278
Goal "[| PPROD I F : constrains AA BB;  i: I |] \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   279
\     ==> F : constrains (Applyall AA i) (Applyall BB i)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   280
by (auto_tac (claset(), 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   281
	      simpset() addsimps [Applyall_def, constrains_def, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   282
				  lift_prog_def, PPROD_def, Acts_JN]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   283
by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] image_eqI]
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   284
			addSEs [rinst [("c", "?ff(i := ?u)")] subsetCE, ballE],
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   285
	       simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   286
qed "PPROD_constrains_projection";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   287
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   288
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   289
Goal "i : I ==> (PPROD I F : stable {f. f i : A}) = (F : stable A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   290
by (asm_simp_tac (simpset() addsimps [stable_def, PPROD_constrains]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   291
qed "PPROD_stable";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   292
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   293
Goal "i : I ==> (PPROD I F : invariant {f. f i : A}) = (F : invariant A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   294
by (auto_tac (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   295
	      simpset() addsimps [invariant_def, PPROD_stable]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   296
qed "PPROD_invariant";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   297
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   298
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   299
(** Substitution Axiom versions: Constrains, Stable **)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   300
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   301
Goal "[| f : reachable (PPROD I F);  i : I |] ==> f i : reachable F";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   302
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   303
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   304
    (claset() addIs reachable.intrs,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   305
     simpset() addsimps [Acts_PPROD]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   306
qed "reachable_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   307
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   308
Goal "reachable (PPROD I F) <= {f. ALL i:I. f i : reachable F}";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   309
by (force_tac (claset() addSDs [reachable_PPROD], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   310
qed "reachable_PPROD_subset1";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   311
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   312
Goal "[| i ~: I;  A : reachable F |]     \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   313
\  ==> ALL f. f : reachable (PPROD I F)  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   314
\             --> f(i:=A) : reachable (lift_prog i F Join PPROD I F)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   315
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   316
by (ALLGOALS Clarify_tac);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   317
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   318
(*Init, Init case*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   319
by (force_tac (claset() addIs reachable.intrs,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   320
	       simpset() addsimps [lift_prog_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   321
(*Init of F, action of PPROD F case*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   322
br reachable.Acts 1;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   323
by (force_tac (claset(), simpset() addsimps [Acts_Join]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   324
ba 1;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   325
by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PPROD]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   326
(*induction over the 2nd "reachable" assumption*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   327
by (eres_inst_tac [("xa","f")] reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   328
(*Init of PPROD F, action of F case*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   329
by (res_inst_tac [("act","lift_act i act")] reachable.Acts 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   330
by (force_tac (claset(), simpset() addsimps [lift_prog_def, Acts_Join]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   331
by (force_tac (claset() addIs [reachable.Init], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   332
by (force_tac (claset() addIs [ext], simpset() addsimps [lift_act_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   333
(*last case: an action of PPROD I F*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   334
br reachable.Acts 1;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   335
by (force_tac (claset(), simpset() addsimps [Acts_Join]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   336
ba 1;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   337
by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PPROD]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   338
qed_spec_mp "reachable_lift_Join_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   339
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   340
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   341
(*The index set must be finite: otherwise infinitely many copies of F can
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   342
  perform actions, and PPROD can never catch up in finite time.*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   343
Goal "finite I ==> {f. ALL i:I. f i : reachable F} <= reachable (PPROD I F)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   344
by (etac finite_induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   345
by (Simp_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   346
by (force_tac (claset() addDs [reachable_lift_Join_PPROD], 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   347
	       simpset() addsimps [PPROD_insert]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   348
qed "reachable_PPROD_subset2";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   349
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   350
Goal "finite I ==> reachable (PPROD I F) = {f. ALL i:I. f i : reachable F}";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   351
by (REPEAT_FIRST (ares_tac [equalityI,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   352
			    reachable_PPROD_subset1, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   353
			    reachable_PPROD_subset2]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   354
qed "reachable_PPROD_eq";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   355
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   356
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   357
Goal "i: I ==> Applyall {f. (ALL i:I. f i : R) & f i : A} i = R Int A";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   358
by (force_tac (claset(), simpset() addsimps [Applyall_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   359
qed "Applyall_Int";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   360
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   361
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   362
Goal "[| i: I;  finite I |]  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   363
\     ==> (PPROD I F : Constrains {f. f i : A} {f. f i : B}) =  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   364
\         (F : Constrains A B)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   365
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   366
    (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   367
     simpset() addsimps [Constrains_def, Collect_conj_eq RS sym,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   368
			 reachable_PPROD_eq]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   369
bd PPROD_constrains_projection 1;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   370
ba 1;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   371
by (asm_full_simp_tac (simpset() addsimps [Applyall_Int]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   372
by (auto_tac (claset(), 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   373
              simpset() addsimps [constrains_def, lift_prog_def, PPROD_def,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   374
                                  Acts_JN]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   375
by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   376
qed "PPROD_Constrains";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   377
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   378
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   379
Goal "[| i: I;  finite I |]  \
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   380
\     ==> (PPROD I F : Stable {f. f i : A}) = (F : Stable A)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   381
by (asm_simp_tac (simpset() addsimps [Stable_def, PPROD_Constrains]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   382
qed "PPROD_Stable";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   383
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   384
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   385