src/HOL/UNITY/PPROD.ML
author paulson
Mon, 17 May 1999 10:38:08 +0200
changeset 6646 3ea726909fff
parent 6575 70d758762c50
child 6826 02c4dd469ec0
permissions -rw-r--r--
"component" now an infix
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
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
     7
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
     8
val rinst = read_instantiate_sg (sign_of thy);
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
     9
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    10
(**** PPROD ****)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    11
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    12
(*** Basic properties ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    13
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    14
Goalw [lift_set_def] "(f : lift_set i A) = (f i : A)";
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    15
by Auto_tac;
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    16
qed "lift_set_iff";
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    17
AddIffs [lift_set_iff];
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    18
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6020
diff changeset
    19
Goalw [lift_act_def] "lift_act i Id = Id";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    20
by Auto_tac;
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6020
diff changeset
    21
qed "lift_act_Id";
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6020
diff changeset
    22
Addsimps [lift_act_Id];
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    23
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    24
Goalw [lift_prog_def] "Init (lift_prog i F) = lift_set i (Init F)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    25
by Auto_tac;
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    26
qed "Init_lift_prog";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    27
Addsimps [Init_lift_prog];
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    28
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    29
Goalw [lift_prog_def] "Acts (lift_prog i F) = lift_act i `` Acts F";
6295
351b3c2b0d83 removed the infernal States, eqStates, compatible, etc.
paulson
parents: 6020
diff changeset
    30
by (auto_tac (claset() addIs [Id_in_Acts RSN (2,image_eqI)], simpset()));
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    31
qed "Acts_lift_prog";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    32
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    33
Goalw [PPROD_def] "Init (PPROD I F) = (INT i:I. lift_set i (Init (F i)))";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    34
by Auto_tac;
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    35
qed "Init_PPROD";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    36
Addsimps [Init_PPROD];
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    37
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    38
Goalw [lift_act_def]
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    39
    "((f,f') : lift_act i act) = (EX s'. f' = f(i := s') & (f i, s') : act)";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    40
by (Blast_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    41
qed "lift_act_eq";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    42
AddIffs [lift_act_eq];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    43
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    44
Goal "Acts (PPROD I F) = insert Id (UN i:I. lift_act i `` Acts (F i))";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    45
by (auto_tac (claset(),
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    46
	      simpset() addsimps [PPROD_def, Acts_lift_prog, Acts_JN]));
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    47
qed "Acts_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    48
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    49
Goal "PPROD {} F = SKIP";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    50
by (simp_tac (simpset() addsimps [PPROD_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    51
qed "PPROD_empty";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    52
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    53
Goal "(PPI i: I. SKIP) = SKIP";
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    54
by (auto_tac (claset() addSIs [program_equalityI],
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    55
	      simpset() addsimps [Acts_lift_prog, SKIP_def, Acts_PPROD]));
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    56
qed "PPROD_SKIP";
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    57
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    58
Addsimps [PPROD_SKIP, PPROD_empty];
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    59
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    60
Goalw [PPROD_def]
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    61
    "PPROD (insert i I) F = (lift_prog i (F i)) Join (PPROD I F)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    62
by Auto_tac;
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    63
qed "PPROD_insert";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    64
6646
3ea726909fff "component" now an infix
paulson
parents: 6575
diff changeset
    65
Goalw [PPROD_def] "i : I ==> (lift_prog i (F i)) component (PPROD I F)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    66
(*blast_tac doesn't use HO unification*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    67
by (fast_tac (claset() addIs [component_JN]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    68
qed "component_PPROD";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    69
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    70
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    71
(*** Safety: co, stable, invariant ***)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    72
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    73
(** 1st formulation of lifting **)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    74
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    75
Goal "(lift_prog i F : (lift_set i A) co (lift_set i B))  =  \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    76
\     (F : A co B)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    77
by (auto_tac (claset(), 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    78
	      simpset() addsimps [constrains_def, Acts_lift_prog]));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    79
by (Blast_tac 2);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    80
by (Force_tac 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    81
qed "lift_prog_constrains_eq";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    82
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    83
Goal "(lift_prog i F : stable (lift_set i A)) = (F : stable A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    84
by (simp_tac (simpset() addsimps [stable_def, lift_prog_constrains_eq]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    85
qed "lift_prog_stable_eq";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    86
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    87
(*This one looks strange!  Proof probably is by case analysis on i=j.*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    88
Goal "F i : A co B  \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    89
\     ==> lift_prog j (F j) : (lift_set i A) co (lift_set i B)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    90
by (auto_tac (claset(), 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    91
	      simpset() addsimps [constrains_def, Acts_lift_prog]));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    92
by (REPEAT (Blast_tac 1));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    93
qed "constrains_imp_lift_prog_constrains";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    94
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    95
Goal "i : I ==>  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    96
\     (PPROD I F : (lift_set i A) co (lift_set i B))  =  \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    97
\     (F i : A co B)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    98
by (asm_simp_tac (simpset() addsimps [PPROD_def, constrains_JN]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    99
by (blast_tac (claset() addIs [lift_prog_constrains_eq RS iffD1, 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   100
			       constrains_imp_lift_prog_constrains]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   101
qed "PPROD_constrains";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   102
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
   103
Goal "i : I ==> (PPROD I F : stable (lift_set i A)) = (F i : stable A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   104
by (asm_simp_tac (simpset() addsimps [stable_def, PPROD_constrains]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   105
qed "PPROD_stable";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   106
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   107
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   108
(** 2nd formulation of lifting **)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   109
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   110
Goal "[| lift_prog i F : AA co BB |] \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   111
\     ==> F : (Applyall AA i) co (Applyall BB i)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   112
by (auto_tac (claset(), 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   113
	      simpset() addsimps [Applyall_def, constrains_def, 
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   114
				  Acts_lift_prog]));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   115
by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] image_eqI],
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   116
	       simpset()) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   117
qed "lift_prog_constrains_projection";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   118
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   119
Goal "[| PPROD I F : AA co BB;  i: I |] \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   120
\     ==> F i : (Applyall AA i) co (Applyall BB i)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   121
by (rtac lift_prog_constrains_projection 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   122
(*rotate this assumption to be last*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   123
by (dres_inst_tac [("psi", "PPROD I F : ?C")] asm_rl 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   124
by (asm_full_simp_tac (simpset() addsimps [PPROD_def, constrains_JN]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   125
qed "PPROD_constrains_projection";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   126
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   127
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   128
(** invariant **)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   129
6451
paulson
parents: 6299
diff changeset
   130
(*UNUSED*)
paulson
parents: 6299
diff changeset
   131
Goal "(lift_prog i F : invariant (lift_set i A)) = (F : invariant A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   132
by (auto_tac (claset(),
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   133
	      simpset() addsimps [invariant_def, lift_prog_stable_eq]));
6451
paulson
parents: 6299
diff changeset
   134
qed "lift_prog_invariant_eq";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   135
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   136
Goal "[| F i : invariant A;  i : I |] \
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   137
\     ==> PPROD I F : invariant (lift_set i A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   138
by (auto_tac (claset(),
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   139
	      simpset() addsimps [invariant_def, PPROD_stable]));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   140
qed "invariant_imp_PPROD_invariant";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   141
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   142
(*The f0 premise ensures that the product is well-defined.*)
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
   143
Goal "[| PPROD I F : invariant (lift_set i A);  i : I;  \
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   144
\        f0: Init (PPROD I F) |] ==> F i : invariant A";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   145
by (auto_tac (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   146
	      simpset() addsimps [invariant_def, PPROD_stable]));
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   147
by (dres_inst_tac [("c", "f0(i:=x)")] subsetD 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   148
by Auto_tac;
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   149
qed "PPROD_invariant_imp_invariant";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   150
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   151
Goal "[| i : I;  f0: Init (PPROD I F) |] \
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
   152
\     ==> (PPROD I F : invariant (lift_set i A)) = (F i : invariant A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   153
by (blast_tac (claset() addIs [invariant_imp_PPROD_invariant, 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   154
			       PPROD_invariant_imp_invariant]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   155
qed "PPROD_invariant";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   156
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   157
(*The f0 premise isn't needed if F is a constant program because then
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   158
  we get an initial state by replicating that of F*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   159
Goal "i : I \
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
   160
\     ==> ((PPI x:I. F) : invariant (lift_set i A)) = (F : invariant A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   161
by (auto_tac (claset(),
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   162
	      simpset() addsimps [invariant_def, PPROD_stable]));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   163
qed "PFUN_invariant";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   164
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   165
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   166
(*** Substitution Axiom versions: Co, Stable ***)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   167
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   168
(** Reachability **)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   169
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   170
Goal "[| f : reachable (PPROD I F);  i : I |] ==> f i : reachable (F i)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   171
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   172
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   173
    (claset() addIs reachable.intrs,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   174
     simpset() addsimps [Acts_PPROD]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   175
qed "reachable_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   176
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   177
Goal "reachable (PPROD I F) <= {f. ALL i:I. f i : reachable (F i)}";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   178
by (force_tac (claset() addSDs [reachable_PPROD], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   179
qed "reachable_PPROD_subset1";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   180
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   181
Goal "[| i ~: I;  A : reachable (F i) |]     \
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   182
\  ==> ALL f. f : reachable (PPROD I F)      \
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   183
\             --> f(i:=A) : reachable (lift_prog i (F i) Join PPROD I F)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   184
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   185
by (ALLGOALS Clarify_tac);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   186
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   187
(*Init, Init case*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   188
by (force_tac (claset() addIs reachable.intrs,
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   189
	       simpset() addsimps [Acts_lift_prog]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   190
(*Init of F, action of PPROD F case*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   191
by (rtac reachable.Acts 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   192
by (force_tac (claset(), simpset() addsimps [Acts_Join]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   193
by (assume_tac 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   194
by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PPROD]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   195
(*induction over the 2nd "reachable" assumption*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   196
by (eres_inst_tac [("xa","f")] reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   197
(*Init of PPROD F, action of F case*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   198
by (res_inst_tac [("act","lift_act i act")] reachable.Acts 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   199
by (force_tac (claset(), simpset() addsimps [Acts_lift_prog, Acts_Join]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   200
by (force_tac (claset() addIs [reachable.Init], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   201
by (force_tac (claset() addIs [ext], simpset() addsimps [lift_act_def]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   202
(*last case: an action of PPROD I F*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   203
by (rtac reachable.Acts 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   204
by (force_tac (claset(), simpset() addsimps [Acts_Join]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   205
by (assume_tac 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   206
by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PPROD]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   207
qed_spec_mp "reachable_lift_Join_PPROD";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   208
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   209
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   210
(*The index set must be finite: otherwise infinitely many copies of F can
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   211
  perform actions, and PPROD can never catch up in finite time.*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   212
Goal "finite I \
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   213
\     ==> {f. ALL i:I. f i : reachable (F i)} <= reachable (PPROD I F)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   214
by (etac finite_induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   215
by (Simp_tac 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   216
by (force_tac (claset() addDs [reachable_lift_Join_PPROD], 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   217
	       simpset() addsimps [PPROD_insert]) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   218
qed "reachable_PPROD_subset2";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   219
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   220
Goal "finite I ==> \
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   221
\     reachable (PPROD I F) = {f. ALL i:I. f i : reachable (F i)}";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   222
by (REPEAT_FIRST (ares_tac [equalityI,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   223
			    reachable_PPROD_subset1, 
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   224
			    reachable_PPROD_subset2]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   225
qed "reachable_PPROD_eq";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   226
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   227
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   228
(** Co **)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   229
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   230
Goal "[| F i : A Co B;  i: I;  finite I |]  \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   231
\     ==> PPROD I F : (lift_set i A) Co (lift_set i B)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   232
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   233
    (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   234
     simpset() addsimps [Constrains_def, Collect_conj_eq RS sym,
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   235
			 reachable_PPROD_eq]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   236
by (auto_tac (claset(), 
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   237
              simpset() addsimps [constrains_def, Acts_lift_prog, PPROD_def,
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   238
                                  Acts_JN]));
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   239
by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1));
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   240
qed "Constrains_imp_PPROD_Constrains";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   241
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   242
Goal "[| ALL i:I. f0 i : R i;   i: I |] \
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   243
\     ==> Applyall ({f. (ALL i:I. f i : R i)} Int lift_set i A) i = R i Int A";
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   244
by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] image_eqI],
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   245
	       simpset() addsimps [Applyall_def, lift_set_def]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   246
qed "Applyall_Int_depend";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   247
6575
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   248
(*Again, we need the f0 premise so that PPROD I F has an initial state;
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   249
  otherwise its Co-property is vacuous.*)
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   250
Goal "[| PPROD I F : (lift_set i A) Co (lift_set i B);  \
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   251
\        i: I;  finite I;  f0: Init (PPROD I F) |]  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   252
\     ==> F i : A Co B";
6575
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   253
by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   254
by (subgoal_tac "ALL i:I. f0 i : reachable (F i)" 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   255
by (blast_tac (claset() addIs [reachable.Init]) 2);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   256
by (dtac PPROD_constrains_projection 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   257
by (assume_tac 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   258
by (asm_full_simp_tac
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   259
    (simpset() addsimps [Applyall_Int_depend, reachable_PPROD_eq]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   260
qed "PPROD_Constrains_imp_Constrains";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   261
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   262
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   263
Goal "[| i: I;  finite I;  f0: Init (PPROD I F) |]  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   264
\     ==> (PPROD I F : (lift_set i A) Co (lift_set i B)) =  \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   265
\         (F i : A Co B)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   266
by (blast_tac (claset() addIs [Constrains_imp_PPROD_Constrains, 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   267
			       PPROD_Constrains_imp_Constrains]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   268
qed "PPROD_Constrains";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   269
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   270
Goal "[| i: I;  finite I;  f0: Init (PPROD I F) |]  \
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
   271
\     ==> (PPROD I F : Stable (lift_set i A)) = (F i : Stable A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   272
by (asm_simp_tac (simpset() delsimps [Init_PPROD]
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   273
			    addsimps [Stable_def, PPROD_Constrains]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   274
qed "PPROD_Stable";
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   275
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   276
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   277
(** PFUN (no dependence on i) doesn't require the f0 premise **)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   278
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   279
Goal "i: I \
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
   280
\     ==> Applyall ({f. (ALL i:I. f i : R)} Int lift_set i A) i = R Int A";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   281
by (force_tac (claset(), simpset() addsimps [Applyall_def]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   282
qed "Applyall_Int";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   283
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   284
Goal "[| (PPI x:I. F) : (lift_set i A) Co (lift_set i B);  \
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   285
\        i: I;  finite I |]  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   286
\     ==> F : A Co B";
6575
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   287
by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   288
by (dtac PPROD_constrains_projection 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   289
by (assume_tac 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   290
by (asm_full_simp_tac
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   291
    (simpset() addsimps [Applyall_Int, Collect_conj_eq RS sym,
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   292
			 reachable_PPROD_eq]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   293
qed "PFUN_Constrains_imp_Constrains";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   294
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   295
Goal "[| i: I;  finite I |]  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   296
\     ==> ((PPI x:I. F) : (lift_set i A) Co (lift_set i B)) =  \
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   297
\         (F : A Co B)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   298
by (blast_tac (claset() addIs [Constrains_imp_PPROD_Constrains, 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   299
			       PFUN_Constrains_imp_Constrains]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   300
qed "PFUN_Constrains";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   301
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   302
Goal "[| i: I;  finite I |]  \
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
   303
\     ==> ((PPI x:I. F) : Stable (lift_set i A)) = (F : Stable A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   304
by (asm_simp_tac (simpset() addsimps [Stable_def, PFUN_Constrains]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   305
qed "PFUN_Stable";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   306
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   307
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   308
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   309
(*** guarantees properties ***)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   310
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   311
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   312
Goal "drop_act i (lift_act i act) = act";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   313
by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] exI],
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   314
	       simpset() addsimps [drop_act_def, lift_act_def]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   315
qed "lift_act_inverse";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   316
Addsimps [lift_act_inverse];
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   317
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   318
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   319
Goal "(lift_prog i F) Join G = lift_prog i H \
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   320
\     ==> EX J. H = F Join J";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   321
by (etac program_equalityE 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   322
by (auto_tac (claset(), simpset() addsimps [Acts_lift_prog, Acts_Join]));
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   323
by (res_inst_tac [("x", 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   324
		   "mk_program(Applyall(Init G) i, drop_act i `` Acts G)")] 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   325
    exI 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   326
by (rtac program_equalityI 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   327
(*Init*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   328
by (simp_tac (simpset() addsimps [Applyall_def]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   329
(*Blast_tac can't do HO unification, needed to invent function states*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   330
by (fast_tac (claset() addEs [equalityE]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   331
(*Now for the Actions*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   332
by (dres_inst_tac [("f", "op `` (drop_act i)")] arg_cong 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   333
by (asm_full_simp_tac 
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   334
    (simpset() addsimps [insert_absorb, Acts_Join,
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   335
			 image_Un, image_compose RS sym, o_def]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   336
qed "lift_prog_Join_eq_lift_prog_D";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   337
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   338
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   339
Goal "F : X guarantees Y \
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   340
\     ==> lift_prog i F : (lift_prog i `` X) guarantees (lift_prog i `` Y)";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   341
by (rtac guaranteesI 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   342
by Auto_tac;
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   343
by (blast_tac (claset() addDs [lift_prog_Join_eq_lift_prog_D, guaranteesD]) 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   344
qed "lift_prog_guarantees";
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   345
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   346