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