# HG changeset patch # User paulson # Date 933953358 -7200 # Node ID 2bc63a44721b2b1631f821bf1fd9c5e3f39f1082 # Parent 676027b1d770f82c53ac729e5d8752775215cbc1 re-organization of theorems from Alloc and PPROD, partly into new theory Lift_prog diff -r 676027b1d770 -r 2bc63a44721b src/HOL/UNITY/Alloc.ML --- a/src/HOL/UNITY/Alloc.ML Fri Aug 06 17:28:45 1999 +0200 +++ b/src/HOL/UNITY/Alloc.ML Fri Aug 06 17:29:18 1999 +0200 @@ -20,8 +20,6 @@ -Addsimps [sub_def]; - Goalw [sysOfAlloc_def] "inj sysOfAlloc"; by (rtac injI 1); by Auto_tac; @@ -47,554 +45,6 @@ AddIffs [finite_lessThan]; -(*could move to PPROD.ML, but function "sub" is needed there*) -Goal "lift_set i {s. P s} = {s. P (sub i s)}"; -by (asm_simp_tac (simpset() addsimps [lift_set_def]) 1); -qed "lift_set_sub"; - -(*ONE OF THESE IS REDUNDANT!*) -Goal "lift_set i {s. P s} = {s. P (s i)}"; -by (asm_simp_tac (simpset() addsimps [lift_set_def]) 1); -qed "lift_set_sub2"; - -Goalw [Increasing_def] - "[| i: I; finite I |] \ -\ ==> ((plam x:I. F) : Increasing (f o sub i)) = (F : Increasing f)"; -by (subgoal_tac "ALL z. {s. z <= (f o sub i) s} = lift_set i {s. z <= f s}" 1); -by (asm_simp_tac (simpset() addsimps [lift_set_sub]) 2); -by (asm_full_simp_tac (simpset() addsimps [const_PLam_Stable]) 1); -qed "const_PLam_Increasing"; - - -(*The other direction fails: having FF : Stable {s. z <= f (s i)} does not - ensure that F has the form lift_prog i F for some F.*) -Goal "lift_prog i `` Stable {s. P (f s)} <= Stable {s. P (f (s i))}"; -by Auto_tac; -by (stac (lift_set_sub2 RS sym) 1); -by (asm_simp_tac (simpset() addsimps [lift_prog_Stable_eq]) 1); -qed "image_lift_prog_Stable"; - -Goal "lift_prog i `` Increasing f <= Increasing (f o sub i)"; -by (simp_tac (simpset() addsimps [Increasing_def, - inj_lift_prog RS image_INT]) 1); -by (blast_tac (claset() addIs [impOfSubs image_lift_prog_Stable]) 1); -qed "image_lift_prog_Increasing"; - - -(*****************PPROD.ML******************) - -(*???????????????*) -Goal "[| F : (drop_prog i `` X) guar (drop_prog i `` Y); modular i Y |] \ -\ ==> lift_prog i F : X guar Y"; -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 1); -by (rtac (drop_prog_lift_prog_Join RS sym) 1); -by (assume_tac 1); -by Auto_tac; -by (full_simp_tac (simpset() addsimps [modular_def, - drop_prog_lift_prog_Join RS sym]) 1); -by (dtac sym 1); -by (Blast_tac 1); -qed "drop_prog_guarantees"; - -Goalw [constrains_def] - "(drop_prog i F : A co B) = \ -\ (F : (lift_set i A) co (lift_set i B))"; -by Auto_tac; -by (force_tac (claset(), - simpset() addsimps [drop_act_def]) 2); -by (blast_tac (claset() addIs [drop_act_I]) 1); -qed "drop_prog_constrains_eq"; - -Goal "(drop_prog i F : stable A) = (F : stable (lift_set i A))"; -by (simp_tac (simpset() addsimps [stable_def, drop_prog_constrains_eq]) 1); -qed "drop_prog_stable_eq"; - -Goal "modular i (stable {s. P(s i)})"; -by (simp_tac (simpset() addsimps [modular_def, lift_set_sub2 RS sym, - drop_prog_stable_eq RS sym]) 1); -by Auto_tac; -qed "modular_stable_i"; - - -(** Weak Constrains and Stable **) - -Goal "f : reachable F ==> f i : reachable (drop_prog i F)"; -by (etac reachable.induct 1); -by (force_tac (claset() addIs [reachable.Acts, drop_act_I], - simpset()) 2); -by (force_tac (claset() addIs [reachable.Init, drop_set_I], - simpset()) 1); -qed "reachable_imp_reachable_drop_prog"; - -Goalw [Constrains_def] - "drop_prog i F : A Co B ==> F : (lift_set i A) Co (lift_set i B)"; -by (full_simp_tac (simpset() addsimps [drop_prog_constrains_eq]) 1); -by (etac constrains_weaken_L 1); -by Auto_tac; -by (etac reachable_imp_reachable_drop_prog 1); -qed "drop_prog_Constrains_D"; - -Goalw [Stable_def] - "drop_prog i F : Stable A ==> F : Stable (lift_set i A)"; -by (asm_simp_tac (simpset() addsimps [drop_prog_Constrains_D]) 1); -qed "drop_prog_Stable_D"; - -(***????????????? -Goal "i ~= j ==> lift_prog i `` UNIV <= stable {s. P (f (s j))}"; -by Auto_tac; -by (stac (lift_set_sub2 RS sym) 1); -by (asm_simp_tac (simpset() addsimps [lift_prog_Stable_eq]) 1); -qed "image_lift_prog_Stable"; -??????????????*) - - - -(****UNITY.ML? BUT it's not clear that Idle is good for anything. - Could equally want to forbid actions that are subsets of Id - OR can simple constraints on actions prevent this?****) -Goalw [constrains_def, Idle_def] "F : Idle ==> (F : A co B) = (A<=B)"; -by (Blast_tac 1); -qed "constrains_Idle_iff"; - -Goalw [stable_def,constrains_def,Idle_def] "F : Idle ==> F : stable A"; -by (Blast_tac 1); -qed "stable_Idle"; - -Goalw [drop_act_def, lift_act_def] - "i~=j ==> drop_act i (lift_act j act) <= Id"; -by Auto_tac; -by (etac subst 1); -by (Asm_simp_tac 1); -qed "neq_drop_act_lift_act"; - -Goal "drop_prog i (JN j:I - {i}. lift_prog j F) : Idle"; -by (simp_tac (simpset() addsimps [Acts_JN,Idle_def]) 1); -by Auto_tac; -by (rtac ccontr 1); - (*otherwise PROOF FAILED because tactics don't get negated conclusion*) -by (blast_tac (claset() addSDs [impOfSubs neq_drop_act_lift_act]) 1); -qed "drop_prog_JN_other_Idle"; - -Goal "[| F component G; G : Idle |] ==> F : Idle"; -by (full_simp_tac (simpset() addsimps [Idle_def, component_eq_subset]) 1); -by Auto_tac; -qed "component_Idle"; - -Goal "G component (JN j: I-{i}. JN H: UNIV. lift_prog j H) \ -\ ==> drop_prog i G : Idle"; -by (dtac drop_prog_mono 1); -by (etac component_Idle 1); -by (simp_tac (simpset() addsimps [drop_prog_JN_other_Idle, - lift_prog_JN RS sym]) 1); -qed "component_JN_neq_imp_Idle"; - -Goal "drop_prog i G : Idle ==> G : stable (lift_set i A)"; -by (simp_tac (simpset() addsimps [drop_prog_stable_eq RS sym]) 1); -by (blast_tac (claset() addIs [stable_Idle]) 1); -qed "Idle_imp_stable_lift_set"; - -(*like neq_drop_act_lift_act but stronger premises and conclusion*) -Goal "[| i~=j; act ~= {} |] ==> drop_act i (lift_act j act) = Id"; -by (rtac equalityI 1); -by (etac neq_drop_act_lift_act 1); -by (asm_simp_tac (simpset() addsimps [drop_act_def, lift_act_def]) 1); -by Auto_tac; -ren "s s'" 1; -by (dres_inst_tac [("x", "f(i:=x,j:=s)")] spec 1); -by (Asm_full_simp_tac 1); -by (dres_inst_tac [("x", "f(i:=x,j:=s')")] spec 1); -by (Asm_full_simp_tac 1); -by (swap_res_tac [ext] 1); -by (Asm_simp_tac 1); -qed "neq_drop_act_lift_act_eq"; - - -Goal "act ~= {} ==> drop_act i (lift_act j act) = (if i=j then act else Id)"; -by (asm_simp_tac (simpset() addsimps [neq_drop_act_lift_act_eq]) 1); -qed "drop_act_lift_act_eq"; - - -(*first premise says that the system has an initial state*) -Goalw [PLam_def] - "[| ALL j:I. f0 j : Init (F j); \ -\ ALL j:I. {} ~: Acts (F j); i: I |] \ -\ ==> drop_prog i (plam j:I. F j) = F i"; -by (simp_tac (simpset() addsimps [Acts_JN, lift_prog_def, drop_prog_def]) 1); -by (rtac program_equalityI 1); -by (asm_simp_tac (simpset() addsimps [drop_set_INT_lift_set]) 1); -by (Asm_simp_tac 1); -by (subgoal_tac - "drop_act i `` (UN i:I. lift_act i `` Acts (F i)) = Acts (F i)" 1); -by (Asm_simp_tac 1); -by Auto_tac; -ren "act" 1; -by (subgoal_tac "act ~= {}" 1); -by (Blast_tac 2); -by (asm_simp_tac (simpset() addsimps [drop_act_lift_act_eq]) 1); -by (rtac image_eqI 1); -by (rtac (lift_act_inverse RS sym) 1); -by Auto_tac; -qed "drop_prog_plam_eq"; - -xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx; - -(**STILL trying to lift the guarantees propery from one client to a family. - first with a "localTo" precondition. Then to look at the void (UNIV) - precondition.**) - -(*CANNOT be proved because the premise tells us NOTHING AT ALL about - actions outside i [rather than telling us there aren't any] - MAY UNIVERSALLY QUANTIFY both premise and conclusion???**) -Goal "F : stable (lift_set i A) ==> x = lift_prog i (drop_prog i x)"; - -(*I.E TRY THIS*) -Goal "{F. F = (plam j:UNIV. drop_prog j F)} Int (INT i. stable (lift_set i A)) \ -\ <= (INT i. lift_prog i `` stable A)"; - - -Goal "{F. F = (plam j:UNIV. drop_prog j F)} Int stable (lift_set i A) \ -\ <= lift_prog i `` stable A"; -by Auto_tac; -fr image_eqI; -by (etac (drop_prog_stable_eq RS iffD2) 2); -by (etac ssubst 1); -by (stac drop_prog_plam_eq 1); -by Auto_tac; - - -by (asm_simp_tac (simpset() addsimps [lift_prog_Stable_eq]) 1); -qed "image_lift_prog_Stable"; - - -Goal "{F. F = (plam i:UNIV. drop_prog i F)} Int stable {s. P (f (s i))} \ -\ <= lift_prog i `` stable {s. P (f s)}"; -by (stac (lift_set_sub2 RS sym) 1); - - -Goal "Cl : (f localTo F) guar Increasing f ==> \ -\ lift_prog i Cl : (f o sub i) localTo (lift_prog i F) \ -\ guar Increasing (f o sub i)"; -by (dtac lift_prog_guarantees 1); -by (etac guarantees_weaken 1); -by (rtac image_lift_prog_Increasing 2); - - - - - - - -Goal "(f o sub i) localTo (lift_prog i F) <= lift_prog i `` (f localTo F)"; - - - - - - -Goal "(f o sub i) localTo (lift_prog i F) <= lift_prog i `` (f localTo F)"; -by (simp_tac (simpset() addsimps [localTo_def]) 1); -by Auto_tac; -by (dres_inst_tac [("x","z")] spec 1); -by Auto_tac; -by (asm_simp_tac (simpset() addsimps [lift_set_sub2 RS sym, - lift_prog_stable_eq]) 1); -qed "image_lift_prog_?"; - -Goal "{GX. EX G: UNIV. lift_prog i G component GX} = UNIV"; -by Auto_tac; -by (asm_full_simp_tac (simpset() addsimps [component_eq_subset]) 1); -by (asm_full_simp_tac (simpset() addsimps [lift_set_def]) 1); -ren "GX" 1; -by (res_inst_tac [("x", "mk_program((%f. f i)``(Init GX), drop_act i `` Acts GX)")] exI 1); -by Auto_tac; - - - - -(*quantified formula's too strong and yet too weak. - Close to what's needed, but maybe need further restrictions on X, Y*) -Goal "[| F : (drop_prog i `` X) guar (drop_prog i `` Y); \ -\ ALL j: -{i}. drop_prog j `` X <= drop_prog j `` Y |] \ -\ ==> lift_prog i F : X guar Y"; -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 1); -by (rtac (drop_prog_lift_prog_Join RS sym) 1); -by (assume_tac 1); -by Auto_tac; -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join RS sym]) 1); -by (subgoal_tac "ALL k. drop_prog k (lift_prog i F Join G) = drop_prog k x" 1); -by (Clarify_tac 2); -by (case_tac "k=i" 2); -by (Blast_tac 2); - -by (dres_inst_tac [("f", "lift_prog i")] arg_cong 1); -by (asm_full_simp_tac (simpset() addsimps [singleton_drop_prog_inverse]) 1); - -by (dtac sym 1); -by (Blast_tac 1); -qed "drop_prog_guarantees"; - -Goalw [guarantees_def] - "[| F : (X Int A) guar (Y Int A); \ -\ F : (X Int B) guar (Y Int B); \ -\ UNIV <= A Un B |] \ -\ ==> F : X guar Y"; -by (Blast_tac 1); -qed "guarantees_Un_partition_I"; - -Goalw [guarantees_def] - "[| ALL i:I. F : (X Int A i) guar (Y Int A i); UNIV <= UNION I A |] \ -\ ==> F : X guar Y"; -by (Blast_tac 1); -qed "guarantees_UN_partition_I"; - -Goal "{G. G component (JN j: I-{i}. JN H: UNIV. lift_prog j H)} \ -\ <= stable (lift_set i A)"; -by (Clarify_tac 1); -by (etac (component_JN_neq_imp_Idle RS Idle_imp_stable_lift_set) 1); -result(); - -Goal "i~=j ==> (lift_prog j `` UNIV) <= Increasing (f o sub i)"; -by (rewtac Increasing_def); -by Auto_tac; -by (stac (lift_set_sub2 RS sym) 1); -by (blast_tac (claset() addIs [neq_imp_lift_prog_Stable]) 1); -result(); - -Goal "Cl : UNIV guar Increasing f ==> \ -\ lift_prog i Cl : {G. G component (JN j: I. JN H: UNIV. lift_prog j H)} \ -\ guar Increasing (f o sub i)"; -by (rtac guarantees_partition_I 1); -by (simp_tac (simpset() addsimps []) 1); - -by (dtac lift_prog_guarantees 1); -by (etac guarantees_weaken 1); -by Auto_tac; -by (rtac (impOfSubs image_lift_prog_Increasing) 1); -by Auto_tac; -result(); - - -Goal "drop_prog i `` UNIV = UNIV"; -by Auto_tac; -fr image_eqI; -by (rtac (lift_prog_inverse RS sym) 1); -by Auto_tac; -result(); - - -Goal "x : reachable (drop_prog i F) ==> x : drop_set i (reachable F)"; -by (etac reachable.induct 1); -by Auto_tac; -by (rewtac drop_set_def); -by (force_tac (claset() addIs [reachable.Init, drop_set_I], - simpset()) 1); -by (rewtac drop_act_def); -by Auto_tac; -by (rtac imageI 1); -by (rtac reachable.Acts 1); -by (assume_tac 1); -by (assume_tac 1); - - -Goal "reachable (drop_prog i F) = drop_set i (reachable F)"; -by (rewtac drop_set_def); -by Auto_tac; -by (etac reachable_imp_reachable_drop_prog 2); - - - - - -Goalw [Constrains_def] - "(F : (lift_set i A) Co (lift_set i B)) \ -\ ==> (drop_prog i F : A Co B)"; -by (full_simp_tac (simpset() addsimps [drop_prog_constrains_eq]) 1); -by (etac constrains_weaken_L 1); - -by Auto_tac; -by (force_tac (claset(), - simpset() addsimps [drop_act_def]) 2); -by (blast_tac (claset() addIs [drop_act_I]) 1); -qed "drop_prog_Constrains_eq"; - - -(*Does NOT appear to be provable, so this notion of modular is too strong*) -Goal "lift_prog i `` drop_prog i `` Increasing (f o sub i) \ -\ <= Increasing (f o sub i)"; -by (simp_tac (simpset() addsimps [Increasing_def]) 1); -by Auto_tac; -by (stac (lift_set_sub2 RS sym) 1); -by (simp_tac (simpset() addsimps [lift_prog_Stable_eq]) 1); - - - - - -Goal "[| F : (drop_prog i `` X) guar (drop_prog i `` Y); \ -\ X <= (UN F. {lift_prog i F}); \ -\ lift_prog i `` drop_prog i `` Y <= Y |] \ -\ ==> lift_prog i F : X guar Y"; -by (rtac guaranteesI 1); -by (set_mp_tac 1); -by (Clarify_tac 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 1); -by (rtac (drop_prog_lift_prog_Join RS sym) 1); -by (assume_tac 1); -by Auto_tac; -by (asm_full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join RS sym]) 1); -by Auto_tac; -qed "drop_prog_guarantees"; - - - -NEW NOTION OF MODULAR - - -Goal "(ALL H:X. lift_prog i (drop_prog i H) : X) = \ -\ (lift_prog i `` drop_prog i `` X <= X)"; -by Auto_tac; -qed "modular_equiv"; - -Goal "lift_prog i `` drop_prog i `` (lift_prog i `` F) <= lift_prog i `` F"; -by (simp_tac (simpset() addsimps [image_compose RS sym, o_def]) 1); -result(); - -Goal "[| F : (drop_prog i `` X) guar (drop_prog i `` Y); \ -\ lift_prog i `` drop_prog i `` X <= X; \ -\ lift_prog i `` drop_prog i `` Y <= Y|] \ -\ ==> lift_prog i F : X guar Y"; -by (full_simp_tac (simpset() addsimps [modular_equiv RS sym]) 1); -by (rtac guaranteesI 1); -by (ball_tac 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 1); -by (assume_tac 2);back(); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by (ALLGOALS Clarify_tac); -by (dres_inst_tac [("f", "lift_prog i")] arg_cong 1); - -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (ALLGOALS Clarify_tac); -by (rtac image_eqI 1); -by (rtac (drop_prog_lift_prog_Join RS sym) 1); -by (assume_tac 1); -by (full_simp_tac (simpset() addsimps [modular_equiv RS sym]) 1); -by (REPEAT (ball_tac 1)); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -ren "H" 1; -by (dtac sym 1); -by (asm_full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -????????????????; - -by Auto_tac; - -by (dres_inst_tac [("f", "drop_prog i")] arg_cong 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by (assume_tac 1); -by (ALLGOALS Clarify_tac); -by (set_mp_tac 1); -by (ALLGOALS Clarify_tac); -by Auto_tac; -by (dres_inst_tac [("f", "drop_prog i")] arg_cong 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by Auto_tac; -qed "drop_prog_guarantees"; - - - - - - - - -Goal "(lift_prog i `` F) <= lift_prog i `` drop_prog i `` (lift_prog i `` F)"; -by (simp_tac (simpset() addsimps [image_compose RS sym, o_def]) 1); -by Auto_tac; -result(); - - - - - -(*FAILS*) -Goal "modular i (lift_prog i `` Y)"; - - - - - -Goal "Cl : UNIV guar Increasing f \ -\ ==>lift_prog i Cl : UNIV guar Increasing (f o sub i)"; -by (rtac drop_prog_guarantees 1); -by (rewtac Increasing_def); -by (etac guarantees_weaken 1); -by Auto_tac; -by (rtac image_eqI 1); -by (rtac (lift_prog_inverse RS sym) 1); -by Auto_tac; -by (stac (lift_set_sub2 RS sym) 1); -by (asm_simp_tac (simpset() addsimps [lift_prog_Stable_eq]) 1); - - - -(*could move to PPROD.ML, but function "sub" is needed there*) -Goalw [drop_set_def] "drop_set i {s. P (f (sub i s))} = {s. P (f s)}"; -by Auto_tac; -qed "drop_set_sub"; - -Goal "lift_set i (reachable (drop_prog i F)) = reachable F"; - -(*False?*) -Goal "modular i ((f o sub i) localTo (lift_prog i F))"; -by (simp_tac (simpset() addsimps [localTo_def, modular_def]) 1); -by Auto_tac; - - - -(*FAILS*) -Goal "(drop_prog i F : A Co B) = \ -\ (F : (lift_set i A) Co (lift_set i B))"; -by (simp_tac (simpset() addsimps [Constrains_def, drop_prog_constrains_eq, - lift_set_Int]) 1); - -?? -by (simp_tac (simpset() addsimps [Constrains_def, reachable_drop_prog, - drop_set_Int RS sym, - drop_prog_constrains_eq]) 1); -qed "drop_prog_Constrains_eq"; - -(*FAILS*) -Goal "(drop_prog i F : Stable A) = (F : Stable (lift_set i A))"; -by (simp_tac (simpset() addsimps [Stable_def, drop_prog_Constrains_eq]) 1); -qed "drop_prog_Stable_eq"; - - -(*FAILS*) -Goal "modular i (Stable {s. P(s i)})"; -by (full_simp_tac (simpset() addsimps [modular_def, - drop_prog_lift_prog_Join RS sym]) 1); -by (full_simp_tac (simpset() addsimps [lift_set_sub2 RS sym, - lift_prog_Stable_eq]) 1); -by Auto_tac; - - - -Goal "modular i (Increasing (rel o sub i))"; -by (full_simp_tac (simpset() addsimps [modular_def, Increasing_def, - drop_prog_lift_prog_Join RS sym]) 1); - - -Goal "modular i (lift_prog i `` X)"; -by (full_simp_tac (simpset() addsimps [modular_def, - drop_prog_lift_prog_Join RS sym]) 1); - Goal "i < Nclients ==> \ @@ -643,90 +93,6 @@ yyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyyy; -!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!; -(*Precondition is TOO STRONG*) -Goal "Cl : UNIV guar Increasing f \ -\ ==>lift_prog i Cl : (lift_prog i `` UNIV) guar Increasing (f o sub i)"; -by (rtac drop_prog_guarantees 1); -by (etac guarantees_weaken 1); -by Auto_tac; -by (rtac (impOfSubs image_lift_prog_Increasing) 2); -by (rtac image_eqI 1); -by (rtac (lift_prog_inverse RS sym) 1); -by (rtac (impOfSubs image_lift_prog_Increasing) 1); -by Auto_tac; -fr imageI; - -by (stac (lift_set_sub2 RS sym) 1); -by (asm_simp_tac (simpset() addsimps [lift_prog_Stable_eq]) 1); - - - -by (full_simp_tac - (simpset() addsimps [alloc_spec_def, alloc_safety_def, System_def]) 1); - - - - -Goal "[| F : (drop_prog i `` XX) guar Y; \ -\ ALL Z:XX. Z component (JN i:UNIV. JN G:UNIV. lift_prog i G) |] \ -\ ==> lift_prog i F : XX guar (lift_prog i `` Y)"; -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (ball_tac 1); -by (full_simp_tac (simpset() addsimps [Join_component_iff]) 1); -by (full_simp_tac (simpset() addsimps [component_eq_subset, Acts_JN]) 1); - -by (rtac image_eqI 2); -by (assume_tac 3); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join RS sym]) 2); - -by (rtac (drop_prog_lift_prog_Join RS sym) 1); -by (assume_tac 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join RS sym]) 1); -by (rtac image_eqI 1); -by (assume_tac 2); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); - -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 1); -by (rtac (drop_prog_lift_prog_Join RS sym) 1); -by (assume_tac 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join RS sym]) 1); -by (rtac image_eqI 1); -by (assume_tac 2); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); - - -by (dtac sym 1); -by (Blast_tac 1); - - -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 2); -by (assume_tac 3); - - -Goal "F : X guar Y \ -\ ==> lift_prog i F : XX guar (lift_prog i `` Y)"; -by (rtac guaranteesI 1); -by (dtac guaranteesD 1); -by (rtac image_eqI 2); -by (assume_tac 3); - -by Auto_tac; -by (blast_tac (claset() addDs [lift_prog_Join_eq_lift_prog_D, guaranteesD]) 1); -qed "lift_prog_guarantees"; - - - - - - - - (*partial system...*) Goal "[| Alloc : alloc_spec; Network : network_spec |] \ \ ==> (extend sysOfAlloc Alloc) Join Network : system_spec"; @@ -761,69 +127,3 @@ by Auto_tac; by (Simp_tac 1); - - - - -Goal "[| F : (drop_prog i `` X) guar (drop_prog i `` Y); \ -\ X <= lift_prog i `` drop_prog i `` X; \ -\ Y <= lift_prog i `` drop_prog i `` Y |] \ -\ ==> lift_prog i F : X guar Y"; -by (rtac guaranteesI 1); -by (set_mp_tac 1); -by (ALLGOALS Clarify_tac); -by (dtac guaranteesD 1); -by (rtac image_eqI 1); -by (dres_inst_tac [("f", "drop_prog i")] arg_cong 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by (assume_tac 1); -by (ALLGOALS Clarify_tac); -by (set_mp_tac 1); -by (ALLGOALS Clarify_tac); -by Auto_tac; -by (dres_inst_tac [("f", "drop_prog i")] arg_cong 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by Auto_tac; -qed "drop_prog_guarantees"; - - - - -(*****SINGLETON LEMMAS: GOOD FOR ANYTHING??????***) - -(*Because A and B could differ outside i, cannot generalize result to - drop_set i (A Int B) = drop_set i A Int drop_set i B -*) -Goal "[| ALL j. j = i; f i = g i |] ==> f = g"; -by (rtac ext 1); -by (dres_inst_tac [("x", "x")] spec 1); -by Auto_tac; -qed "singleton_ext"; - -Goalw [lift_set_def, drop_set_def] - "ALL j. j = i ==> lift_set i (drop_set i A) = A"; -by (blast_tac (claset() addSDs [singleton_ext]) 1); -qed "singleton_drop_set_inverse"; - -Goal "ALL j. j = i ==> f(i:= g i) = g"; -by (dres_inst_tac [("x", "x")] spec 1); -by Auto_tac; -qed "singleton_update_eq"; - -Goalw [lift_act_def, drop_act_def] - "ALL j. j = i ==> lift_act i (drop_act i act) = act"; -by Auto_tac; -by (asm_simp_tac (simpset() addsimps [singleton_update_eq]) 1); -by (dtac singleton_ext 1); -by (assume_tac 1); -by (auto_tac (claset() addSIs [exI, image_eqI], - simpset() addsimps [singleton_update_eq])); -qed "singleton_drop_act_inverse"; - -Goal "ALL j. j = i ==> lift_prog i (drop_prog i F) = F"; -by (rtac program_equalityI 1); -by (asm_simp_tac (simpset() addsimps [singleton_drop_set_inverse]) 1); -by (asm_simp_tac (simpset() addsimps [image_compose RS sym, o_def, - singleton_drop_act_inverse]) 1); -qed "singleton_drop_prog_inverse"; - diff -r 676027b1d770 -r 2bc63a44721b src/HOL/UNITY/Alloc.thy --- a/src/HOL/UNITY/Alloc.thy Fri Aug 06 17:28:45 1999 +0200 +++ b/src/HOL/UNITY/Alloc.thy Fri Aug 06 17:29:18 1999 +0200 @@ -8,23 +8,6 @@ Alloc = Follows + Extend + PPROD + -constdefs - (**TOO STRONG: DELETE**) - modular :: "['a, ('a=>'b) program set] => bool" - "modular i X == - ALL F G. F : X --> drop_prog i F = drop_prog i G --> G : X" - - - (*UNITY.thy????????????????*) - (*An Idle program can do nothing*) - Idle :: 'a program set - "Idle == {F. Union (Acts F) <= Id}" - -(*simplifies the expression of some specifications*) -constdefs - sub :: ['a, 'a=>'b] => 'b - "sub i f == f i" - (*Duplicated from HOL/ex/NatSum.thy. Maybe type should be [nat=>int, nat] => int**) consts sum :: [nat=>nat, nat] => nat diff -r 676027b1d770 -r 2bc63a44721b src/HOL/UNITY/PPROD.ML --- a/src/HOL/UNITY/PPROD.ML Fri Aug 06 17:28:45 1999 +0200 +++ b/src/HOL/UNITY/PPROD.ML Fri Aug 06 17:29:18 1999 +0200 @@ -1,183 +1,25 @@ (* Title: HOL/UNITY/PPROD.ML ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1998 University of Cambridge + Copyright 1999 University of Cambridge + +Abstraction over replicated components (PLam) +General products of programs (Pi operation) + +It is not clear that any of this is good for anything. *) val rinst = read_instantiate_sg (sign_of thy); -(**** PPROD ****) (*** Basic properties ***) -(** lift_set and drop_set **) - -Goalw [lift_set_def] "(f : lift_set i A) = (f i : A)"; -by Auto_tac; -qed "lift_set_iff"; -AddIffs [lift_set_iff]; - -Goalw [lift_set_def] "lift_set i (A Int B) = lift_set i A Int lift_set i B"; -by Auto_tac; -qed "lift_set_Int"; - -(*USED?? converse fails*) -Goalw [drop_set_def] "f : A ==> f i : drop_set i A"; -by Auto_tac; -qed "drop_set_I"; - -(** lift_act and drop_act **) - -Goalw [lift_act_def] "lift_act i Id = Id"; -by Auto_tac; -by (etac rev_mp 1); -by (dtac sym 1); -by (asm_simp_tac (simpset() addsimps [fun_upd_idem_iff]) 1); -qed "lift_act_Id"; -Addsimps [lift_act_Id]; - -Goalw [drop_act_def] "(s, s') : act ==> (s i, s' i) : drop_act i act"; -by (auto_tac (claset() addSIs [image_eqI], simpset())); -qed "drop_act_I"; - -Goalw [drop_act_def] "drop_act i Id = Id"; -by Auto_tac; -by (res_inst_tac [("x", "((%u. x), (%u. x))")] image_eqI 1); -by Auto_tac; -qed "drop_act_Id"; -Addsimps [drop_act_Id]; - -(** lift_prog and drop_prog **) - -Goalw [lift_prog_def] "Init (lift_prog i F) = lift_set 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"; -Addsimps [Acts_lift_prog]; - -Goalw [drop_prog_def] "Init (drop_prog i F) = drop_set i (Init F)"; -by Auto_tac; -qed "Init_drop_prog"; -Addsimps [Init_drop_prog]; - -Goalw [drop_prog_def] "Acts (drop_prog i F) = drop_act i `` Acts F"; -by (auto_tac (claset() addIs [Id_in_Acts RSN (2,image_eqI)], simpset())); -qed "Acts_drop_prog"; -Addsimps [Acts_drop_prog]; - -Goal "F component G ==> lift_prog i F component lift_prog i G"; -by (full_simp_tac (simpset() addsimps [component_eq_subset]) 1); -by Auto_tac; -qed "lift_prog_mono"; - -Goal "F component G ==> drop_prog i F component drop_prog i G"; -by (full_simp_tac (simpset() addsimps [component_eq_subset, drop_set_def]) 1); -by Auto_tac; -qed "drop_prog_mono"; - -Goal "lift_prog i SKIP = SKIP"; -by (auto_tac (claset() addSIs [program_equalityI], - simpset() addsimps [SKIP_def, lift_prog_def])); -qed "lift_prog_SKIP"; - -Goal "lift_prog i (F Join G) = (lift_prog i F) Join (lift_prog i G)"; -by (rtac program_equalityI 1); -by (auto_tac (claset(), simpset() addsimps [Acts_Join])); -qed "lift_prog_Join"; - -Goal "lift_prog i (JOIN I F) = (JN j:I. lift_prog i (F j))"; -by (rtac program_equalityI 1); -by (auto_tac (claset(), simpset() addsimps [Acts_JN])); -qed "lift_prog_JN"; - -Goal "drop_prog i SKIP = SKIP"; -by (auto_tac (claset() addSIs [program_equalityI], - simpset() addsimps [SKIP_def, drop_set_def, drop_prog_def])); -qed "drop_prog_SKIP"; - - -(** Injectivity of lift_set, lift_act, lift_prog **) - -Goalw [lift_set_def, drop_set_def] "drop_set i (lift_set i F) = F"; -by Auto_tac; -qed "lift_set_inverse"; -Addsimps [lift_set_inverse]; - -Goal "inj (lift_set i)"; -by (rtac inj_on_inverseI 1); -by (rtac lift_set_inverse 1); -qed "inj_lift_set"; - -(*Because A and B could differ outside i, cannot generalize result to - drop_set i (A Int B) = drop_set i A Int drop_set i B -*) -Goalw [lift_set_def, drop_set_def] - "drop_set i ((lift_set i A) Int B) = A Int (drop_set i B)"; -by Auto_tac; -qed "drop_set_lift_set_Int"; - -Goalw [lift_set_def, drop_set_def] - "drop_set i (B Int (lift_set i A)) = (drop_set i B) Int A"; -by Auto_tac; -qed "drop_set_lift_set_Int2"; - -Goalw [drop_set_def] - "i : I ==> drop_set i (INT j:I. lift_set j A) = A"; -by Auto_tac; -qed "drop_set_INT"; - -Goalw [lift_act_def, drop_act_def] "drop_act i (lift_act i act) = act"; -by Auto_tac; -by (res_inst_tac [("x", "f(i:=a)")] exI 1); -by (Asm_simp_tac 1); -by (res_inst_tac [("x", "f(i:=b)")] exI 1); -by (Asm_simp_tac 1); -by (rtac ext 1); -by (Asm_simp_tac 1); -qed "lift_act_inverse"; -Addsimps [lift_act_inverse]; - -Goal "inj (lift_act i)"; -by (rtac inj_on_inverseI 1); -by (rtac lift_act_inverse 1); -qed "inj_lift_act"; - -Goal "drop_prog i (lift_prog i F) = F"; -by (simp_tac (simpset() addsimps [lift_prog_def, drop_prog_def]) 1); -by (rtac program_equalityI 1); -by (simp_tac (simpset() addsimps [image_compose RS sym, o_def]) 2); -by (Simp_tac 1); -qed "lift_prog_inverse"; -Addsimps [lift_prog_inverse]; - -Goal "inj (lift_prog i)"; -by (rtac inj_on_inverseI 1); -by (rtac lift_prog_inverse 1); -qed "inj_lift_prog"; - - -(** PLam **) - Goalw [PLam_def] "Init (PLam I F) = (INT i:I. lift_set i (Init (F i)))"; by Auto_tac; qed "Init_PLam"; Addsimps [Init_PLam]; -(*For compatibility with the original definition and perhaps simpler proofs*) -Goalw [lift_act_def] - "((f,f') : lift_act i act) = (EX s'. f' = f(i := s') & (f i, s') : act)"; -by Auto_tac; -by (rtac exI 1); -by Auto_tac; -qed "lift_act_eq"; -AddIffs [lift_act_eq]; - - Goal "Acts (PLam I F) = insert Id (UN i:I. lift_act i `` Acts (F i))"; by (auto_tac (claset(), simpset() addsimps [PLam_def, Acts_JN])); @@ -204,51 +46,7 @@ qed "component_PLam"; -(*** Safety: co, stable, invariant ***) - -(** Safety and lift_prog **) - -Goal "(lift_prog i F : (lift_set i A) co (lift_set i B)) = \ -\ (F : A co B)"; -by (auto_tac (claset(), - simpset() addsimps [constrains_def])); -by (Blast_tac 2); -by (Force_tac 1); -qed "lift_prog_constrains_eq"; - -Goal "(lift_prog i F : stable (lift_set 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. - If i~=j then lift_prog j (F j) does nothing to lift_set i, and the - premise ensures A<=B.*) -Goal "F i : A co B \ -\ ==> lift_prog j (F j) : (lift_set i A) co (lift_set i B)"; -by (auto_tac (claset(), - simpset() addsimps [constrains_def])); -by (REPEAT (Blast_tac 1)); -qed "constrains_imp_lift_prog_constrains"; - -(** safety properties for program unit j hold trivially of unit i **) -Goalw [constrains_def] - "[| i~=j; A<= B|] ==> lift_prog i F : (lift_set j A) co (lift_set j B)"; -by Auto_tac; -qed "neq_imp_lift_prog_constrains"; - -Goalw [stable_def] - "i~=j ==> lift_prog i F : stable (lift_set j A)"; -by (blast_tac (claset() addIs [neq_imp_lift_prog_constrains]) 1); -qed "neq_imp_lift_prog_stable"; - -bind_thm ("neq_imp_lift_prog_Constrains", - neq_imp_lift_prog_constrains RS constrains_imp_Constrains); - -bind_thm ("neq_imp_lift_prog_Stable", - neq_imp_lift_prog_stable RS stable_imp_Stable); - - -(** Safety and PLam **) +(** Safety **) Goal "i : I ==> \ \ (PLam I F : (lift_set i A) co (lift_set i B)) = \ @@ -262,33 +60,26 @@ by (asm_simp_tac (simpset() addsimps [stable_def, PLam_constrains]) 1); qed "PLam_stable"; - -(** Safety, lift_prog + drop_set **) - +(*Possibly useful in Lift_set.ML?*) Goal "[| lift_prog i F : AA co BB |] \ \ ==> F : (drop_set i AA) co (drop_set i BB)"; by (auto_tac (claset(), simpset() addsimps [constrains_def, drop_set_def])); by (force_tac (claset() addSIs [rinst [("x", "?ff(i := ?u)")] image_eqI], simpset()) 1); -qed "lift_prog_constrains_projection"; +qed "lift_prog_constrains_drop_set"; Goal "[| PLam I F : AA co BB; i: I |] \ \ ==> F i : (drop_set i AA) co (drop_set i BB)"; -by (rtac lift_prog_constrains_projection 1); +by (rtac lift_prog_constrains_drop_set 1); (*rotate this assumption to be last*) by (dres_inst_tac [("psi", "PLam I F : ?C")] asm_rl 1); by (asm_full_simp_tac (simpset() addsimps [PLam_def, constrains_JN]) 1); -qed "PLam_constrains_projection"; +qed "PLam_constrains_drop_set"; (** invariant **) -Goal "(lift_prog i F : invariant (lift_set i A)) = (F : invariant A)"; -by (auto_tac (claset(), - simpset() addsimps [invariant_def, lift_prog_stable_eq])); -qed "lift_prog_invariant_eq"; - Goal "[| F i : invariant A; i : I |] \ \ ==> PLam I F : invariant (lift_set i A)"; by (auto_tac (claset(), @@ -319,46 +110,7 @@ qed "const_PLam_invariant"; -(*** Substitution Axiom versions: Co, Stable ***) - -(*** Reachability ***) - -(** for lift_prog **) - -Goal "s : reachable F ==> f(i:=s) : reachable (lift_prog i F)"; -by (etac reachable.induct 1); -by (force_tac (claset() addIs [reachable.Acts, ext], - simpset()) 2); -by (force_tac (claset() addIs [reachable.Init], simpset()) 1); -qed "reachable_lift_progI"; - -Goal "f : reachable (lift_prog i F) ==> f i : reachable F"; -by (etac reachable.induct 1); -by Auto_tac; -by (ALLGOALS (blast_tac (claset() addIs reachable.intrs))); -qed "reachable_lift_progD"; - -Goal "reachable (lift_prog i F) = lift_set i (reachable F)"; -by Auto_tac; -by (etac reachable_lift_progD 1); -ren "f" 1; -by (dres_inst_tac [("f","f"),("i","i")] reachable_lift_progI 1); -by Auto_tac; -qed "reachable_lift_prog"; - -Goal "(lift_prog i F : (lift_set i A) Co (lift_set i B)) = \ -\ (F : A Co B)"; -by (simp_tac (simpset() addsimps [Constrains_def, reachable_lift_prog, - lift_set_Int RS sym, - lift_prog_constrains_eq]) 1); -qed "lift_prog_Constrains_eq"; - -Goal "(lift_prog i F : Stable (lift_set i A)) = (F : Stable A)"; -by (simp_tac (simpset() addsimps [Stable_def, lift_prog_Constrains_eq]) 1); -qed "lift_prog_Stable_eq"; - - -(** Reachability for PLam **) +(** Reachability **) Goal "[| f : reachable (PLam I F); i : I |] ==> f i : reachable (F i)"; by (etac reachable.induct 1); @@ -450,7 +202,7 @@ by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1); by (subgoal_tac "ALL i:I. f0 i : reachable (F i)" 1); by (blast_tac (claset() addIs [reachable.Init]) 2); -by (dtac PLam_constrains_projection 1); +by (dtac PLam_constrains_drop_set 1); by (assume_tac 1); by (asm_full_simp_tac (simpset() addsimps [drop_set_lift_set_Int2, @@ -478,7 +230,7 @@ \ i: I; finite I |] \ \ ==> F : A Co B"; by (full_simp_tac (simpset() addsimps [Constrains_eq_constrains]) 1); -by (dtac PLam_constrains_projection 1); +by (dtac PLam_constrains_drop_set 1); by (assume_tac 1); by (asm_full_simp_tac (simpset() addsimps [drop_set_INT, @@ -498,31 +250,18 @@ by (asm_simp_tac (simpset() addsimps [Stable_def, const_PLam_Constrains]) 1); qed "const_PLam_Stable"; +Goalw [Increasing_def] + "[| i: I; finite I |] \ +\ ==> ((plam x:I. F) : Increasing (f o sub i)) = (F : Increasing f)"; +by (subgoal_tac "ALL z. {s. z <= (f o sub i) s} = lift_set i {s. z <= f s}" 1); +by (asm_simp_tac (simpset() addsimps [lift_set_sub]) 2); +by (asm_full_simp_tac + (simpset() addsimps [finite_lessThan, const_PLam_Stable]) 1); +qed "const_PLam_Increasing"; (*** guarantees properties ***) -Goal "drop_prog i ((lift_prog i F) Join G) = F Join (drop_prog i G)"; -by (rtac program_equalityI 1); -by (simp_tac (simpset() addsimps [Acts_Join, image_Un, - image_compose RS sym, o_def]) 2); -by (simp_tac (simpset() addsimps [drop_set_lift_set_Int]) 1); -qed "drop_prog_lift_prog_Join"; - -Goal "(lift_prog i F) Join G = lift_prog i H \ -\ ==> H = F Join (drop_prog i G)"; -by (dres_inst_tac [("f", "drop_prog i")] arg_cong 1); -by (full_simp_tac (simpset() addsimps [drop_prog_lift_prog_Join]) 1); -by (etac sym 1); -qed "lift_prog_Join_eq_lift_prog_D"; - -Goal "F : X guar Y \ -\ ==> lift_prog i F : (lift_prog i `` X) guar (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"; - Goalw [PLam_def] "[| ALL i:I. F i : X guar Y |] \ \ ==> (PLam I F) : (INT i: I. lift_prog i `` X) \ diff -r 676027b1d770 -r 2bc63a44721b src/HOL/UNITY/PPROD.thy --- a/src/HOL/UNITY/PPROD.thy Fri Aug 06 17:28:45 1999 +0200 +++ b/src/HOL/UNITY/PPROD.thy Fri Aug 06 17:29:18 1999 +0200 @@ -3,36 +3,14 @@ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge -General products of programs (Pi operation), for replicating components. +Abstraction over replicated components (PLam) +General products of programs (Pi operation) *) -PPROD = Union + Comp + +PPROD = Lift_prog + constdefs - lift_set :: "['a, 'b set] => ('a => 'b) set" - "lift_set i A == {f. f i : A}" - - drop_set :: "['a, ('a=>'b) set] => 'b set" - "drop_set i A == (%f. f i) `` A" - - lift_act :: "['a, ('b*'b) set] => (('a=>'b) * ('a=>'b)) set" - "lift_act i act == {(f,f'). f(i:= f' i) = f' & (f i, f' i) : 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 (lift_set i (Init F), - lift_act i `` Acts F)" - - drop_prog :: "['a, ('a=>'b) program] => 'b program" - "drop_prog i F == - mk_program (drop_set i (Init F), - drop_act i `` (Acts F))" - - (*products of programs*) PLam :: ['a set, 'a => 'b program] => ('a => 'b) program "PLam I F == JN i:I. lift_prog i (F i)"