# HG changeset patch # User paulson # Date 912005700 -3600 # Node ID 2430ccbde87d66de62e66ecc21e01bec3319821b # Parent c5a7a7685826c7dc88c9ca5535f1b28e7104af04 guarantees laws diff -r c5a7a7685826 -r 2430ccbde87d src/HOL/UNITY/PPROD.ML --- a/src/HOL/UNITY/PPROD.ML Wed Nov 25 15:54:41 1998 +0100 +++ b/src/HOL/UNITY/PPROD.ML Wed Nov 25 15:55:00 1998 +0100 @@ -4,6 +4,9 @@ Copyright 1998 University of Cambridge *) + Addsimps [image_id]; + + val rinst = read_instantiate_sg (sign_of thy); (*** General lemmas ***) @@ -183,69 +186,76 @@ by (asm_simp_tac (simpset() addsimps [leadsTo_Basis, Lcopy_ensures]) 1); qed "leadsTo_imp_Lcopy_leadsTo"; -Goal "Lcopy F : ensures A B ==> F : ensures (Domain A) (Domain B)"; -by (full_simp_tac - (simpset() addsimps [ensures_def, Lcopy_constrains, Lcopy_transient, - Domain_Un_eq RS sym, - Sigma_Un_distrib1 RS sym, - Sigma_Diff_distrib1 RS sym]) 1); -by (safe_tac (claset() addSDs [Lcopy_constrains_Domain])); -by (etac constrains_weaken_L 1); -by (Blast_tac 1); -(*NOT able to prove this: -Lcopy F : ensures A B ==> F : ensures (Domain A) (Domain B) - 1. [| Lcopy F : transient (A - B); - F : constrains (Domain (A - B)) (Domain (A Un B)) |] - ==> F : transient (Domain A - Domain B) -**) +(** + Goal "Lcopy F : ensures A B ==> F : ensures (Domain A) (Domain B)"; + by (full_simp_tac + (simpset() addsimps [ensures_def, Lcopy_constrains, Lcopy_transient, + Domain_Un_eq RS sym, + Sigma_Un_distrib1 RS sym, + Sigma_Diff_distrib1 RS sym]) 1); + by (safe_tac (claset() addSDs [Lcopy_constrains_Domain])); + by (etac constrains_weaken_L 1); + by (Blast_tac 1); + (*NOT able to prove this: + Lcopy F : ensures A B ==> F : ensures (Domain A) (Domain B) + 1. [| Lcopy F : transient (A - B); + F : constrains (Domain (A - B)) (Domain (A Un B)) |] + ==> F : transient (Domain A - Domain B) + **) -Goal "Lcopy F : leadsTo AU BU ==> F : leadsTo (Domain AU) (Domain BU)"; -by (etac leadsTo_induct 1); -by (full_simp_tac (simpset() addsimps [Domain_Union]) 3); -by (blast_tac (claset() addIs [leadsTo_UN]) 3); -by (blast_tac (claset() addIs [leadsTo_Trans]) 2); -by (rtac leadsTo_Basis 1); -(*...and so CANNOT PROVE THIS*) + Goal "Lcopy F : leadsTo AU BU ==> F : leadsTo (Domain AU) (Domain BU)"; + by (etac leadsTo_induct 1); + by (full_simp_tac (simpset() addsimps [Domain_Union]) 3); + by (blast_tac (claset() addIs [leadsTo_UN]) 3); + by (blast_tac (claset() addIs [leadsTo_Trans]) 2); + by (rtac leadsTo_Basis 1); + (*...and so CANNOT PROVE THIS*) -(*This also seems impossible to prove??*) -Goal "(Lcopy F : leadsTo (A Times UNIV) (B Times UNIV)) = \ -\ (F : leadsTo A B)"; - + (*This also seems impossible to prove??*) + Goal "(Lcopy F : leadsTo (A Times UNIV) (B Times UNIV)) = \ + \ (F : leadsTo A B)"; +**) (**** PPROD ****) (*** Basic properties ***) -Goalw [PPROD_def, lift_prog_def] - "Init (PPROD I F) = {f. ALL i:I. f i : Init F}"; -by Auto_tac; -qed "Init_PPROD"; - Goalw [lift_act_def] "lift_act i Id = Id"; by Auto_tac; qed "lift_act_Id"; Addsimps [lift_act_Id]; +Goalw [lift_prog_def] "Init (lift_prog i F) = {f. f i : Init F}"; +by Auto_tac; +qed "Init_lift_prog"; +Addsimps [Init_lift_prog]; + +Goalw [lift_prog_def] "Acts (lift_prog i F) = lift_act i `` Acts F"; +by (auto_tac (claset() addIs [Id_in_Acts RSN (2,image_eqI)], simpset())); +qed "Acts_lift_prog"; + +Goalw [PPROD_def] "Init (PPROD I F) = {f. ALL i:I. f i : Init (F i)}"; +by Auto_tac; +qed "Init_PPROD"; +Addsimps [Init_PPROD]; + Goalw [lift_act_def] "((f,f') : lift_act i act) = (EX s'. f' = f(i := s') & (f i, s') : act)"; by (Blast_tac 1); qed "lift_act_eq"; AddIffs [lift_act_eq]; -Goal "Acts (PPROD I F) = insert Id (UN i:I. lift_act i `` Acts F)"; +Goal "Acts (PPROD I F) = insert Id (UN i:I. lift_act i `` Acts (F i))"; by (auto_tac (claset(), - simpset() addsimps [PPROD_def, lift_prog_def, Acts_JN])); + simpset() addsimps [PPROD_def, Acts_lift_prog, Acts_JN])); qed "Acts_PPROD"; -Addsimps [Init_PPROD]; - -Goal "PPROD I SKIP = SKIP"; -by (rtac program_equalityI 1); -by (auto_tac (claset(), - simpset() addsimps [PPROD_def, lift_prog_def, +Goal "(PPI i: I. SKIP) = SKIP"; +by (auto_tac (claset() addSIs [program_equalityI], + simpset() addsimps [PPROD_def, Acts_lift_prog, SKIP_def, Acts_JN])); qed "PPROD_SKIP"; @@ -255,131 +265,325 @@ Addsimps [PPROD_SKIP, PPROD_empty]; -Goalw [PPROD_def] "PPROD (insert i I) F = (lift_prog i F) Join (PPROD I F)"; +Goalw [PPROD_def] + "PPROD (insert i I) F = (lift_prog i (F i)) Join (PPROD I F)"; by Auto_tac; qed "PPROD_insert"; +Goalw [PPROD_def] "i : I ==> component (lift_prog i (F i)) (PPROD I F)"; +(*blast_tac doesn't use HO unification*) +by (fast_tac (claset() addIs [component_JN]) 1); +qed "component_PPROD"; -(*** Safety: constrains, stable ***) + +(*** Safety: constrains, stable, invariant ***) + +(** 1st formulation of lifting **) -val subsetCE' = rinst - [("c", "(%u. ?s)(i:=?s')")] subsetCE; +Goal "(lift_prog i F : constrains {f. f i : A} {f. f i : B}) = \ +\ (F : constrains A B)"; +by (auto_tac (claset(), + simpset() addsimps [constrains_def, Acts_lift_prog])); +by (Blast_tac 2); +by (Force_tac 1); +qed "lift_prog_constrains_eq"; + +Goal "(lift_prog i F : stable {f. f i : A}) = (F : stable A)"; +by (simp_tac (simpset() addsimps [stable_def, lift_prog_constrains_eq]) 1); +qed "lift_prog_stable_eq"; + +(*This one looks strange! Proof probably is by case analysis on i=j.*) +Goal "F i : constrains A B \ +\ ==> lift_prog j (F j) : constrains {f. f i : A} {f. f i : B}"; +by (auto_tac (claset(), + simpset() addsimps [constrains_def, Acts_lift_prog])); +by (REPEAT (Blast_tac 1)); +qed "constrains_imp_lift_prog_constrains"; Goal "i : I ==> \ \ (PPROD I F : constrains {f. f i : A} {f. f i : B}) = \ -\ (F : constrains A B)"; -by (auto_tac (claset(), - simpset() addsimps [constrains_def, lift_prog_def, PPROD_def, - Acts_JN])); -by (REPEAT (Blast_tac 2)); -by (force_tac (claset() addSEs [subsetCE', allE, ballE], simpset()) 1); +\ (F i : constrains A B)"; +by (asm_simp_tac (simpset() addsimps [PPROD_def, constrains_JN]) 1); +by (blast_tac (claset() addIs [lift_prog_constrains_eq RS iffD1, + constrains_imp_lift_prog_constrains]) 1); qed "PPROD_constrains"; -Goal "[| PPROD I F : constrains AA BB; i: I |] \ +Goal "i : I ==> (PPROD I F : stable {f. f i : A}) = (F i : stable A)"; +by (asm_simp_tac (simpset() addsimps [stable_def, PPROD_constrains]) 1); +qed "PPROD_stable"; + + +(** 2nd formulation of lifting **) + +Goal "[| lift_prog i F : constrains AA BB |] \ \ ==> F : constrains (Applyall AA i) (Applyall BB i)"; by (auto_tac (claset(), simpset() addsimps [Applyall_def, constrains_def, - lift_prog_def, PPROD_def, Acts_JN])); -by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] image_eqI] - addSEs [rinst [("c", "?ff(i := ?u)")] subsetCE, ballE], + Acts_lift_prog])); +by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] image_eqI], simpset()) 1); +qed "lift_prog_constrains_projection"; + +Goal "[| PPROD I F : constrains AA BB; i: I |] \ +\ ==> F i : constrains (Applyall AA i) (Applyall BB i)"; +by (rtac lift_prog_constrains_projection 1); +(*rotate this assumption to be last*) +by (dres_inst_tac [("psi", "PPROD I F : ?C")] asm_rl 1); +by (asm_full_simp_tac (simpset() addsimps [PPROD_def, constrains_JN]) 1); qed "PPROD_constrains_projection"; -Goal "i : I ==> (PPROD I F : stable {f. f i : A}) = (F : stable A)"; -by (asm_simp_tac (simpset() addsimps [stable_def, PPROD_constrains]) 1); -qed "PPROD_stable"; +(** invariant **) + +Goal "F : invariant A ==> lift_prog i F : invariant {f. f i : A}"; +by (auto_tac (claset(), + simpset() addsimps [invariant_def, lift_prog_stable_eq])); +qed "invariant_imp_lift_prog_invariant"; -Goal "i : I ==> (PPROD I F : invariant {f. f i : A}) = (F : invariant A)"; +Goal "[| lift_prog i F : invariant {f. f i : A} |] ==> F : invariant A"; +by (auto_tac (claset(), + simpset() addsimps [invariant_def, lift_prog_stable_eq])); +qed "lift_prog_invariant_imp_invariant"; + +(*NOT clear that the previous lemmas help in proving this one.*) +Goal "[| F i : invariant A; i : I |] ==> PPROD I F : invariant {f. f i : A}"; +by (auto_tac (claset(), + simpset() addsimps [invariant_def, PPROD_stable])); +qed "invariant_imp_PPROD_invariant"; + +(*The f0 premise ensures that the product is well-defined.*) +Goal "[| PPROD I F : invariant {f. f i : A}; i : I; \ +\ f0: Init (PPROD I F) |] ==> F i : invariant A"; by (auto_tac (claset(), simpset() addsimps [invariant_def, PPROD_stable])); +by (dres_inst_tac [("c", "f0(i:=x)")] subsetD 1); +by Auto_tac; +qed "PPROD_invariant_imp_invariant"; + +Goal "[| i : I; f0: Init (PPROD I F) |] \ +\ ==> (PPROD I F : invariant {f. f i : A}) = (F i : invariant A)"; +by (blast_tac (claset() addIs [invariant_imp_PPROD_invariant, + PPROD_invariant_imp_invariant]) 1); qed "PPROD_invariant"; +(*The f0 premise isn't needed if F is a constant program because then + we get an initial state by replicating that of F*) +Goal "i : I \ +\ ==> ((PPI x:I. F) : invariant {f. f i : A}) = (F : invariant A)"; +by (auto_tac (claset(), + simpset() addsimps [invariant_def, PPROD_stable])); +qed "PFUN_invariant"; + -(** Substitution Axiom versions: Constrains, Stable **) +(*** Substitution Axiom versions: Constrains, Stable ***) -Goal "[| f : reachable (PPROD I F); i : I |] ==> f i : reachable F"; +(** Reachability **) + +Goal "[| f : reachable (PPROD I F); i : I |] ==> f i : reachable (F i)"; by (etac reachable.induct 1); by (auto_tac (claset() addIs reachable.intrs, simpset() addsimps [Acts_PPROD])); qed "reachable_PPROD"; -Goal "reachable (PPROD I F) <= {f. ALL i:I. f i : reachable F}"; +Goal "reachable (PPROD I F) <= {f. ALL i:I. f i : reachable (F i)}"; by (force_tac (claset() addSDs [reachable_PPROD], simpset()) 1); qed "reachable_PPROD_subset1"; -Goal "[| i ~: I; A : reachable F |] \ -\ ==> ALL f. f : reachable (PPROD I F) \ -\ --> f(i:=A) : reachable (lift_prog i F Join PPROD I F)"; +Goal "[| i ~: I; A : reachable (F i) |] \ +\ ==> ALL f. f : reachable (PPROD I F) \ +\ --> f(i:=A) : reachable (lift_prog i (F i) Join PPROD I F)"; by (etac reachable.induct 1); by (ALLGOALS Clarify_tac); by (etac reachable.induct 1); (*Init, Init case*) by (force_tac (claset() addIs reachable.intrs, - simpset() addsimps [lift_prog_def]) 1); + simpset() addsimps [Acts_lift_prog]) 1); (*Init of F, action of PPROD F case*) -br reachable.Acts 1; +by (rtac reachable.Acts 1); by (force_tac (claset(), simpset() addsimps [Acts_Join]) 1); -ba 1; +by (assume_tac 1); by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PPROD]) 1); (*induction over the 2nd "reachable" assumption*) by (eres_inst_tac [("xa","f")] reachable.induct 1); (*Init of PPROD F, action of F case*) by (res_inst_tac [("act","lift_act i act")] reachable.Acts 1); -by (force_tac (claset(), simpset() addsimps [lift_prog_def, Acts_Join]) 1); +by (force_tac (claset(), simpset() addsimps [Acts_lift_prog, Acts_Join]) 1); by (force_tac (claset() addIs [reachable.Init], simpset()) 1); by (force_tac (claset() addIs [ext], simpset() addsimps [lift_act_def]) 1); (*last case: an action of PPROD I F*) -br reachable.Acts 1; +by (rtac reachable.Acts 1); by (force_tac (claset(), simpset() addsimps [Acts_Join]) 1); -ba 1; +by (assume_tac 1); by (force_tac (claset() addIs [ext], simpset() addsimps [Acts_PPROD]) 1); qed_spec_mp "reachable_lift_Join_PPROD"; (*The index set must be finite: otherwise infinitely many copies of F can perform actions, and PPROD can never catch up in finite time.*) -Goal "finite I ==> {f. ALL i:I. f i : reachable F} <= reachable (PPROD I F)"; +Goal "finite I \ +\ ==> {f. ALL i:I. f i : reachable (F i)} <= reachable (PPROD I F)"; by (etac finite_induct 1); by (Simp_tac 1); by (force_tac (claset() addDs [reachable_lift_Join_PPROD], simpset() addsimps [PPROD_insert]) 1); qed "reachable_PPROD_subset2"; -Goal "finite I ==> reachable (PPROD I F) = {f. ALL i:I. f i : reachable F}"; +Goal "finite I ==> \ +\ reachable (PPROD I F) = {f. ALL i:I. f i : reachable (F i)}"; by (REPEAT_FIRST (ares_tac [equalityI, reachable_PPROD_subset1, reachable_PPROD_subset2])); qed "reachable_PPROD_eq"; -Goal "i: I ==> Applyall {f. (ALL i:I. f i : R) & f i : A} i = R Int A"; -by (force_tac (claset(), simpset() addsimps [Applyall_def]) 1); -qed "Applyall_Int"; +(** Constrains **) - -Goal "[| i: I; finite I |] \ -\ ==> (PPROD I F : Constrains {f. f i : A} {f. f i : B}) = \ -\ (F : Constrains A B)"; +Goal "[| F i : Constrains A B; i: I; finite I |] \ +\ ==> PPROD I F : Constrains {f. f i : A} {f. f i : B}"; by (auto_tac (claset(), simpset() addsimps [Constrains_def, Collect_conj_eq RS sym, reachable_PPROD_eq])); -bd PPROD_constrains_projection 1; -ba 1; -by (asm_full_simp_tac (simpset() addsimps [Applyall_Int]) 1); by (auto_tac (claset(), - simpset() addsimps [constrains_def, lift_prog_def, PPROD_def, + simpset() addsimps [constrains_def, Acts_lift_prog, PPROD_def, Acts_JN])); by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1)); -qed "PPROD_Constrains"; +qed "Constrains_imp_PPROD_Constrains"; + +Goal "[| ALL i:I. f0 i : R i; i: I |] \ +\ ==> Applyall {f. (ALL i:I. f i : R i) & f i : A} i = R i Int A"; +by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] exI], + simpset() addsimps [Applyall_def]) 1); +qed "Applyall_Int_depend"; + +(*Again, we need the f0 premise because otherwise Constrains holds trivially + for PPROD I F.*) +Goal "[| PPROD I F : Constrains {f. f i : A} {f. f i : B}; \ +\ i: I; finite I; f0: Init (PPROD I F) |] \ +\ ==> F i : Constrains A B"; +by (full_simp_tac (simpset() addsimps [Constrains_def]) 1); +by (subgoal_tac "ALL i:I. f0 i : reachable (F i)" 1); +by (blast_tac (claset() addIs [reachable.Init]) 2); +by (dtac PPROD_constrains_projection 1); +by (assume_tac 1); +by (asm_full_simp_tac + (simpset() addsimps [Applyall_Int_depend, Collect_conj_eq RS sym, + reachable_PPROD_eq]) 1); +qed "PPROD_Constrains_imp_Constrains"; -Goal "[| i: I; finite I |] \ -\ ==> (PPROD I F : Stable {f. f i : A}) = (F : Stable A)"; -by (asm_simp_tac (simpset() addsimps [Stable_def, PPROD_Constrains]) 1); +Goal "[| i: I; finite I; f0: Init (PPROD I F) |] \ +\ ==> (PPROD I F : Constrains {f. f i : A} {f. f i : B}) = \ +\ (F i : Constrains A B)"; +by (blast_tac (claset() addIs [Constrains_imp_PPROD_Constrains, + PPROD_Constrains_imp_Constrains]) 1); +qed "PPROD_Constrains"; + +Goal "[| i: I; finite I; f0: Init (PPROD I F) |] \ +\ ==> (PPROD I F : Stable {f. f i : A}) = (F i : Stable A)"; +by (asm_simp_tac (simpset() delsimps [Init_PPROD] + addsimps [Stable_def, PPROD_Constrains]) 1); qed "PPROD_Stable"; +(** PFUN (no dependence on i) doesn't require the f0 premise **) +Goal "i: I ==> Applyall {f. (ALL i:I. f i : R) & f i : A} i = R Int A"; +by (force_tac (claset(), simpset() addsimps [Applyall_def]) 1); +qed "Applyall_Int"; + +Goal "[| (PPI x:I. F) : Constrains {f. f i : A} {f. f i : B}; \ +\ i: I; finite I |] \ +\ ==> F : Constrains A B"; +by (full_simp_tac (simpset() addsimps [Constrains_def]) 1); +by (dtac PPROD_constrains_projection 1); +by (assume_tac 1); +by (asm_full_simp_tac + (simpset() addsimps [Applyall_Int, Collect_conj_eq RS sym, + reachable_PPROD_eq]) 1); +qed "PFUN_Constrains_imp_Constrains"; + +Goal "[| i: I; finite I |] \ +\ ==> ((PPI x:I. F) : Constrains {f. f i : A} {f. f i : B}) = \ +\ (F : Constrains A B)"; +by (blast_tac (claset() addIs [Constrains_imp_PPROD_Constrains, + PFUN_Constrains_imp_Constrains]) 1); +qed "PFUN_Constrains"; + +Goal "[| i: I; finite I |] \ +\ ==> ((PPI x:I. F) : Stable {f. f i : A}) = (F : Stable A)"; +by (asm_simp_tac (simpset() addsimps [Stable_def, PFUN_Constrains]) 1); +qed "PFUN_Stable"; + + + +(*** guarantees properties ***) + +(*We only need act2=Id but the condition used is very weak*) +Goal "(x,y): act2 ==> fst_act (act1 RTimes act2) = act1"; +by (auto_tac (claset(), simpset() addsimps [fst_act_def, RTimes_def])); +qed "fst_act_RTimes"; +Addsimps [fst_act_RTimes]; + + +Goal "(Lcopy F) Join G = Lcopy H ==> EX J. H = F Join J"; +by (etac program_equalityE 1); +by (auto_tac (claset(), simpset() addsimps [Acts_Join])); +by (res_inst_tac + [("x", "mk_program(Domain (Init G), fst_act `` Acts G)")] exI 1); +by (rtac program_equalityI 1); +(*Init*) +by (simp_tac (simpset() addsimps [Acts_Join]) 1); +by (blast_tac (claset() addEs [equalityE]) 1); +(*Now for the Actions*) +by (dres_inst_tac [("f", "op `` fst_act")] arg_cong 1); +by (asm_full_simp_tac + (simpset() addsimps [insert_absorb, Acts_Lcopy, Acts_Join, + image_Un, image_compose RS sym, o_def]) 1); +qed "Lcopy_Join_eq_Lcopy_D"; + + +Goal "F : X guarantees Y \ +\ ==> Lcopy F : (Lcopy `` X) guarantees (Lcopy `` Y)"; +by (rtac guaranteesI 1); +by Auto_tac; +by (blast_tac (claset() addDs [Lcopy_Join_eq_Lcopy_D, guaranteesD]) 1); +qed "Lcopy_guarantees"; + + +Goal "drop_act i (lift_act i act) = act"; +by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] exI], + simpset() addsimps [drop_act_def, lift_act_def]) 1); +qed "lift_act_inverse"; +Addsimps [lift_act_inverse]; + + +Goal "(lift_prog i F) Join G = lift_prog i H \ +\ ==> EX J. H = F Join J"; +by (etac program_equalityE 1); +by (auto_tac (claset(), simpset() addsimps [Acts_lift_prog, Acts_Join])); +by (res_inst_tac [("x", + "mk_program(Applyall(Init G) i, drop_act i `` Acts G)")] + exI 1); +by (rtac program_equalityI 1); +(*Init*) +by (simp_tac (simpset() addsimps [Applyall_def]) 1); +(*Blast_tac can't do HO unification, needed to invent function states*) +by (fast_tac (claset() addEs [equalityE]) 1); +(*Now for the Actions*) +by (dres_inst_tac [("f", "op `` (drop_act i)")] arg_cong 1); +by (asm_full_simp_tac + (simpset() addsimps [insert_absorb, Acts_Join, + image_Un, image_compose RS sym, o_def]) 1); +qed "lift_prog_Join_eq_lift_prog_D"; + + +Goal "F : X guarantees Y \ +\ ==> lift_prog i F : (lift_prog i `` X) guarantees (lift_prog i `` Y)"; +by (rtac guaranteesI 1); +by Auto_tac; +by (blast_tac (claset() addDs [lift_prog_Join_eq_lift_prog_D, guaranteesD]) 1); +qed "lift_prog_guarantees"; + + diff -r c5a7a7685826 -r 2430ccbde87d src/HOL/UNITY/PPROD.thy --- a/src/HOL/UNITY/PPROD.thy Wed Nov 25 15:54:41 1998 +0100 +++ b/src/HOL/UNITY/PPROD.thy Wed Nov 25 15:55:00 1998 +0100 @@ -7,7 +7,7 @@ Also merging of state sets. *) -PPROD = Union + +PPROD = Union + Comp + constdefs (*Cartesian product of two relations*) @@ -22,6 +22,9 @@ *) constdefs + fst_act :: "(('a*'b) * ('a*'b)) set => ('a*'a) set" + "fst_act act == (%((x,y),(x',y')). (x,x')) `` act" + Lcopy :: "'a program => ('a*'b) program" "Lcopy F == mk_program (Init F Times UNIV, (%act. act RTimes Id) `` Acts F)" @@ -29,12 +32,15 @@ lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set" "lift_act i act == {(f,f'). EX s'. f' = f(i:=s') & (f i, s') : act}" + drop_act :: "['a, (('a=>'b) * ('a=>'b)) set] => ('b*'b) set" + "drop_act i act == (%(f,f'). (f i, f' i)) `` act" + lift_prog :: "['a, 'b program] => ('a => 'b) program" "lift_prog i F == mk_program ({f. f i : Init F}, lift_act i `` Acts F)" (*products of programs*) - PPROD :: ['a set, 'b program] => ('a => 'b) program - "PPROD I F == JN i:I. lift_prog i F" + PPROD :: ['a set, 'a => 'b program] => ('a => 'b) program + "PPROD I F == JN i:I. lift_prog i (F i)" syntax "@PPROD" :: [pttrn, 'a set, 'b set] => ('a => 'b) set ("(3PPI _:_./ _)" 10)