src/HOL/UNITY/PPROD.ML
author berghofe
Thu, 07 Oct 1999 15:40:32 +0200
changeset 7786 cf9d07ad62af
parent 7689 affe0c2fdfbf
child 7826 c6a8b73b6c2a
permissions -rw-r--r--
Replaced update_new by update.
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
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
     4
    Copyright   1999  University of Cambridge
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
     5
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
     6
Abstraction over replicated components (PLam)
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
     7
General products of programs (Pi operation)
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
     8
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
     9
It is not clear that any of this is good for anything.
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    10
*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    11
6020
4b109bf75976 towards handling sharing of variables
paulson
parents: 6012
diff changeset
    12
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    13
(*** Basic properties ***)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    14
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    15
Goalw [PLam_def] "Init (PLam I F) = (INT i:I. lift_set i (Init (F i)))";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    16
by Auto_tac;
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    17
qed "Init_PLam";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    18
Addsimps [Init_PLam];
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    19
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    20
Goal "Acts (PLam I F) = insert Id (UN i:I. lift_act i `` Acts (F i))";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    21
by (auto_tac (claset(),
7537
875754b599df working snapshot
paulson
parents: 7523
diff changeset
    22
	      simpset() addsimps [PLam_def]));
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    23
qed "Acts_PLam";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    24
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    25
Goal "PLam {} F = SKIP";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    26
by (simp_tac (simpset() addsimps [PLam_def]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    27
qed "PLam_empty";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    28
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    29
Goal "(plam i: I. SKIP) = SKIP";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    30
by (simp_tac (simpset() addsimps [PLam_def,lift_prog_SKIP,JN_constant]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    31
qed "PLam_SKIP";
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    32
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    33
Addsimps [PLam_SKIP, PLam_empty];
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    34
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    35
Goalw [PLam_def]
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    36
    "PLam (insert i I) F = (lift_prog i (F i)) Join (PLam I F)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    37
by Auto_tac;
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    38
qed "PLam_insert";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    39
7399
cf780c2bcccf changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
paulson
parents: 7361
diff changeset
    40
Goalw [PLam_def] "i : I ==> lift_prog i (F i) <= (PLam I F)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    41
(*blast_tac doesn't use HO unification*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    42
by (fast_tac (claset() addIs [component_JN]) 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    43
qed "component_PLam";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    44
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    45
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    46
(** Safety & Progress **)
6835
588f791ee737 addition of drop_... operators with new results and simplification of old ones
paulson
parents: 6826
diff changeset
    47
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    48
Goal "i : I ==>  \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    49
\     (PLam I F : (lift_set i A) co (lift_set i B))  =  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
    50
\     (F i : A co B)";
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    51
by (asm_simp_tac (simpset() addsimps [PLam_def, JN_constrains]) 1);
7344
d54e871d77e0 new guarantees laws; also better natural deduction style for old ones
paulson
parents: 7188
diff changeset
    52
by (blast_tac (claset() addIs [lift_prog_constrains RS iffD1, 
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    53
			       constrains_imp_lift_prog_constrains]) 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    54
qed "PLam_constrains";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    55
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    56
Goal "i : I ==> (PLam I F : stable (lift_set i A)) = (F i : stable A)";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    57
by (asm_simp_tac (simpset() addsimps [stable_def, PLam_constrains]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    58
qed "PLam_stable";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    59
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    60
Goal "i : I ==> \
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    61
\   PLam I F : transient A = (EX i:I. lift_prog i (F i) : transient A)";
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    62
by (asm_simp_tac (simpset() addsimps [JN_transient, PLam_def]) 1);
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    63
qed "PLam_transient";
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    64
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    65
Addsimps [PLam_constrains, PLam_stable, PLam_transient];
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    66
7538
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    67
Goal "[| i : I;  F i : A ensures B |] ==>  \
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    68
\     PLam I F : (lift_set i A) ensures lift_set i B";
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    69
by (auto_tac (claset(), 
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    70
	      simpset() addsimps [ensures_def, lift_prog_transient_eq_disj]));
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    71
qed "PLam_ensures";
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    72
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    73
Goal "[| i : I;  F i : (A-B) co (A Un B);  F i : transient (A-B) |] ==>  \
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    74
\     PLam I F : (lift_set i A) leadsTo lift_set i B";
7538
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    75
by (rtac (PLam_ensures RS leadsTo_Basis) 1);
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    76
by (rtac ensuresI 2);
357873391561 new rule PLam_ensures
paulson
parents: 7537
diff changeset
    77
by (ALLGOALS assume_tac);
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    78
qed "PLam_leadsTo_Basis";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    79
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    80
Goal "[| PLam I F : AA co BB;  i: I |] \
6835
588f791ee737 addition of drop_... operators with new results and simplification of old ones
paulson
parents: 6826
diff changeset
    81
\     ==> F i : (drop_set i AA) co (drop_set i BB)";
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
    82
by (rtac lift_prog_constrains_drop_set 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    83
(*rotate this assumption to be last*)
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    84
by (dres_inst_tac [("psi", "PLam I F : ?C")] asm_rl 1);
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    85
by (asm_full_simp_tac (simpset() addsimps [PLam_def, JN_constrains]) 1);
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
    86
qed "PLam_constrains_drop_set";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    87
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
    88
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    89
(** invariant **)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    90
6299
1a88db6e7c7e UNITY fully working at last...
paulson
parents: 6295
diff changeset
    91
Goal "[| F i : invariant A;  i : I |] \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    92
\     ==> PLam I F : invariant (lift_set i A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    93
by (auto_tac (claset(),
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
    94
	      simpset() addsimps [invariant_def]));
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    95
qed "invariant_imp_PLam_invariant";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    96
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
    97
(*The f0 premise ensures that the product is well-defined.*)
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    98
Goal "[| PLam I F : invariant (lift_set i A);  i : I;  \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
    99
\        f0: Init (PLam I F) |] ==> F i : invariant A";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   100
by (auto_tac (claset(),
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
   101
	      simpset() addsimps [invariant_def]));
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   102
by (dres_inst_tac [("c", "f0(i:=x)")] subsetD 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   103
by Auto_tac;
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   104
qed "PLam_invariant_imp_invariant";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   105
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   106
Goal "[| i : I;  f0: Init (PLam I F) |] \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   107
\     ==> (PLam I F : invariant (lift_set i A)) = (F i : invariant A)";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   108
by (blast_tac (claset() addIs [invariant_imp_PLam_invariant, 
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   109
			       PLam_invariant_imp_invariant]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   110
qed "PLam_invariant";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   111
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   112
(*The f0 premise isn't needed if F is a constant program because then
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   113
  we get an initial state by replicating that of F*)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   114
Goal "i : I \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   115
\     ==> ((plam x:I. F) : invariant (lift_set i A)) = (F : invariant A)";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   116
by (auto_tac (claset(),
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
   117
	      simpset() addsimps [invariant_def]));
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   118
qed "const_PLam_invariant";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   119
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   120
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   121
(** Reachability **)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   122
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   123
Goal "[| f : reachable (PLam I F);  i : I |] ==> f i : reachable (F i)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   124
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   125
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   126
    (claset() addIs reachable.intrs,
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   127
     simpset() addsimps [Acts_PLam]));
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   128
qed "reachable_PLam";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   129
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   130
(*Result to justify a re-organization of this file*)
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   131
Goal "{f. ALL i:I. f i : R i} = (INT i:I. lift_set i (R i))";
6867
7cb9d3250db7 expandshort
paulson
parents: 6842
diff changeset
   132
by Auto_tac;
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   133
result();
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   134
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   135
Goal "reachable (PLam I F) <= (INT i:I. lift_set i (reachable (F i)))";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   136
by (force_tac (claset() addSDs [reachable_PLam], simpset()) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   137
qed "reachable_PLam_subset1";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   138
6826
02c4dd469ec0 many new results for reachable and lift_prog
paulson
parents: 6646
diff changeset
   139
(*simplify using reachable_lift_prog??*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   140
Goal "[| i ~: I;  A : reachable (F i) |]     \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   141
\  ==> ALL f. f : reachable (PLam I F)      \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   142
\             --> f(i:=A) : reachable (lift_prog i (F i) Join PLam I F)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   143
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   144
by (ALLGOALS Clarify_tac);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   145
by (etac reachable.induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   146
(*Init, Init case*)
6826
02c4dd469ec0 many new results for reachable and lift_prog
paulson
parents: 6646
diff changeset
   147
by (force_tac (claset() addIs reachable.intrs, simpset()) 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   148
(*Init of F, action of PLam F case*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   149
by (rtac reachable.Acts 1);
7537
875754b599df working snapshot
paulson
parents: 7523
diff changeset
   150
by (Force_tac 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   151
by (assume_tac 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   152
by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PLam]) 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   153
(*induction over the 2nd "reachable" assumption*)
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   154
by (eres_inst_tac [("xa","f")] reachable.induct 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   155
(*Init of PLam F, action of F case*)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   156
by (res_inst_tac [("act","lift_act i act")] reachable.Acts 1);
7537
875754b599df working snapshot
paulson
parents: 7523
diff changeset
   157
by (Force_tac 1);
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   158
by (force_tac (claset() addIs [reachable.Init], simpset()) 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   159
by (force_tac (claset() addIs [ext], simpset() addsimps [lift_act_def]) 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   160
(*last case: an action of PLam I F*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   161
by (rtac reachable.Acts 1);
7537
875754b599df working snapshot
paulson
parents: 7523
diff changeset
   162
by (Force_tac 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   163
by (assume_tac 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   164
by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PLam]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   165
qed_spec_mp "reachable_lift_Join_PLam";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   166
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   167
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   168
(*The index set must be finite: otherwise infinitely many copies of F can
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   169
  perform actions, and PLam can never catch up in finite time.*)
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   170
Goal "finite I \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   171
\     ==> (INT i:I. lift_set i (reachable (F i))) <= reachable (PLam I F)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   172
by (etac finite_induct 1);
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   173
by (Simp_tac 1);
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   174
by (force_tac (claset() addDs [reachable_lift_Join_PLam], 
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   175
	       simpset() addsimps [PLam_insert]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   176
qed "reachable_PLam_subset2";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   177
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   178
Goal "finite I ==> \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   179
\     reachable (PLam I F) = (INT i:I. lift_set i (reachable (F i)))";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   180
by (REPEAT_FIRST (ares_tac [equalityI,
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   181
			    reachable_PLam_subset1, 
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   182
			    reachable_PLam_subset2]));
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   183
qed "reachable_PLam_eq";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   184
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   185
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   186
(** Co **)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   187
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   188
Goal "[| F i : A Co B;  i: I;  finite I |]  \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   189
\     ==> PLam I F : (lift_set i A) Co (lift_set i B)";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   190
by (auto_tac
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   191
    (claset(),
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   192
     simpset() addsimps [Constrains_def, Collect_conj_eq RS sym,
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   193
			 reachable_PLam_eq]));
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   194
by (auto_tac (claset(), 
7537
875754b599df working snapshot
paulson
parents: 7523
diff changeset
   195
              simpset() addsimps [constrains_def, PLam_def]));
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   196
by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1));
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   197
qed "Constrains_imp_PLam_Constrains";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   198
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   199
Goal "[| ALL j:I. f0 j : A j;   i: I |] \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   200
\     ==> drop_set i (INT j:I. lift_set j (A j)) = A i";
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
   201
by (force_tac (claset() addSIs [image_eqI'],
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   202
	       simpset() addsimps [drop_set_def]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   203
qed "drop_set_INT_lift_set";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   204
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   205
(*Again, we need the f0 premise so that PLam I F has an initial state;
6575
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   206
  otherwise its Co-property is vacuous.*)
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   207
Goal "[| PLam I F : (lift_set i A) Co (lift_set i B);  \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   208
\        i: I;  finite I;  f0: Init (PLam I F) |]  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   209
\     ==> F i : A Co B";
6575
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   210
by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   211
by (subgoal_tac "ALL i:I. f0 i : reachable (F i)" 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   212
by (blast_tac (claset() addIs [reachable.Init]) 2);
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   213
by (dtac PLam_constrains_drop_set 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   214
by (assume_tac 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   215
by (asm_full_simp_tac
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
   216
    (simpset() addsimps [drop_set_Int_lift_set2,
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   217
			 drop_set_INT_lift_set, reachable_PLam_eq]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   218
qed "PLam_Constrains_imp_Constrains";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   219
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   220
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   221
Goal "[| i: I;  finite I;  f0: Init (PLam I F) |]  \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   222
\     ==> (PLam I F : (lift_set i A) Co (lift_set i B)) =  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   223
\         (F i : A Co B)";
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   224
by (blast_tac (claset() addIs [Constrains_imp_PLam_Constrains, 
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   225
			       PLam_Constrains_imp_Constrains]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   226
qed "PLam_Constrains";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   227
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   228
Goal "[| i: I;  finite I;  f0: Init (PLam I F) |]  \
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   229
\     ==> (PLam I F : Stable (lift_set i A)) = (F i : Stable A)";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   230
by (asm_simp_tac (simpset() delsimps [Init_PLam]
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   231
			    addsimps [Stable_def, PLam_Constrains]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   232
qed "PLam_Stable";
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   233
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   234
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   235
(** const_PLam (no dependence on i) doesn't require the f0 premise **)
5899
13d4753079fe new theory PPROD
paulson
parents:
diff changeset
   236
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   237
Goal "[| (plam x:I. F) : (lift_set i A) Co (lift_set i B);  \
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   238
\        i: I;  finite I |]  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   239
\     ==> F : A Co B";
6575
70d758762c50 new definitions of Co and LeadsTo
paulson
parents: 6536
diff changeset
   240
by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1);
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   241
by (dtac PLam_constrains_drop_set 1);
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   242
by (assume_tac 1);
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   243
by (asm_full_simp_tac
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   244
    (simpset() addsimps [drop_set_INT,
7523
3a716ebc2fc0 more rational theorem names (?)
paulson
parents: 7399
diff changeset
   245
			 drop_set_Int_lift_set2, Collect_conj_eq RS sym,
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   246
			 reachable_PLam_eq]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   247
qed "const_PLam_Constrains_imp_Constrains";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   248
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   249
Goal "[| i: I;  finite I |]  \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   250
\     ==> ((plam x:I. F) : (lift_set i A) Co (lift_set i B)) =  \
6536
281d44905cab made many specification operators infix
paulson
parents: 6451
diff changeset
   251
\         (F : A Co B)";
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   252
by (blast_tac (claset() addIs [Constrains_imp_PLam_Constrains, 
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   253
			       const_PLam_Constrains_imp_Constrains]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   254
qed "const_PLam_Constrains";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   255
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   256
Goal "[| i: I;  finite I |]  \
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   257
\     ==> ((plam x:I. F) : Stable (lift_set i A)) = (F : Stable A)";
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   258
by (asm_simp_tac (simpset() addsimps [Stable_def, const_PLam_Constrains]) 1);
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   259
qed "const_PLam_Stable";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   260
7188
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   261
Goalw [Increasing_def]
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   262
     "[| i: I;  finite I |]  \
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   263
\     ==> ((plam x:I. F) : Increasing (f o sub i)) = (F : Increasing f)";
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   264
by (subgoal_tac "ALL z. {s. z <= (f o sub i) s} = lift_set i {s. z <= f s}" 1);
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   265
by (asm_simp_tac (simpset() addsimps [lift_set_sub]) 2);
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   266
by (asm_full_simp_tac
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   267
    (simpset() addsimps [finite_lessThan, const_PLam_Stable]) 1);
2bc63a44721b re-organization of theorems from Alloc and PPROD, partly into new theory
paulson
parents: 6867
diff changeset
   268
qed "const_PLam_Increasing";
5972
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   269
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   270
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   271
(*** guarantees properties ***)
2430ccbde87d guarantees laws
paulson
parents: 5899
diff changeset
   272
6842
56e08e264961 renamed PPI to plam
paulson
parents: 6835
diff changeset
   273
Goalw [PLam_def]
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7538
diff changeset
   274
    "[| lift_prog i (F i): X guarantees Y;  i : I |]  \
7361
477e1bdf230f changed "guar" back to "guarantees" (sorry)
paulson
parents: 7344
diff changeset
   275
\    ==> (PLam I F) : X guarantees Y";
7344
d54e871d77e0 new guarantees laws; also better natural deduction style for old ones
paulson
parents: 7188
diff changeset
   276
by (asm_simp_tac (simpset() addsimps [guarantees_JN_I]) 1);
d54e871d77e0 new guarantees laws; also better natural deduction style for old ones
paulson
parents: 7188
diff changeset
   277
qed "guarantees_PLam_I";