# HG changeset patch # User wenzelm # Date 918475376 -3600 # Node ID db63752140c770135a718343409d407c4b360262 # Parent f6335d319e9fe8301ba7318ddefd5948aed8c5cb updated (Stephan Merz); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Action.ML --- a/src/HOL/TLA/Action.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Action.ML Mon Feb 08 13:02:56 1999 +0100 @@ -6,351 +6,253 @@ Lemmas and tactics for TLA actions. *) -val act_rews = [pairSF_def RS eq_reflection,unl_before,unl_after,unchanged_def, - pr_con,pr_before,pr_lift,pr_lift2,pr_lift3,pr_all,pr_ex]; +(* The following assertion specializes "intI" for any world type + which is a pair, not just for "state * state". +*) +qed_goal "actionI" Action.thy "(!!s t. (s,t) |= A) ==> |- A" + (fn [prem] => [REPEAT (resolve_tac [prem,intI,prod_induct] 1)]); + +qed_goal "actionD" Action.thy "|- A ==> (s,t) |= A" + (fn [prem] => [rtac (prem RS intD) 1]); + +local + fun prover s = prove_goal Action.thy s + (fn _ => [rtac actionI 1, + rewrite_goals_tac (unl_after::intensional_rews), + rtac refl 1]) +in + val pr_rews = map (int_rewrite o prover) + [ "|- (#c)` = #c", + "|- f` = f", + "|- f` = f", + "|- f` = f", + "|- (! x. P x)` = (! x. (P x)`)", + "|- (? x. P x)` = (? x. (P x)`)" + ] +end; + +val act_rews = [unl_before,unl_after,unchanged_def] @ pr_rews; +Addsimps act_rews; val action_rews = act_rews @ intensional_rews; -qed_goal "actionI" Action.thy "(!!s t. ([[s,t]] |= A)) ==> A" - (fn [prem] => [REPEAT (resolve_tac [prem,intI,state2_ext] 1)]); - -qed_goal "actionD" Action.thy "A ==> ([[s,t]] |= A)" - (fn [prem] => [REPEAT (resolve_tac [prem,intD] 1)]); - - - (* ================ Functions to "unlift" action theorems into HOL rules ================ *) -(* Basic unlifting introduces a world parameter and applies basic rewrites, e.g. - A .= B gets ([[s,t]] |= A) = ([[s,t]] |= B) - A .-> B gets ([[s,t]] |= A) --> ([[s,t]] |= B) +(* The following functions are specialized versions of the corresponding + functions defined in Intensional.ML in that they introduce a + "world" parameter of the form (s,t) and apply additional rewrites. *) -fun action_unlift th = rewrite_rule action_rews (th RS actionD); - -(* A .-> B becomes A [[s,t]] ==> B [[s,t]] *) -fun action_mp th = zero_var_indexes ((action_unlift th) RS mp); +fun action_unlift th = + (rewrite_rule action_rews (th RS actionD)) + handle _ => int_unlift th; -(* A .-> B becomes [| A[[s,t]]; B[[s,t]] ==> R |] ==> R - so that it can be used as an elimination rule -*) -fun action_impE th = zero_var_indexes ((action_unlift th) RS impE); +(* Turn |- A = B into meta-level rewrite rule A == B *) +val action_rewrite = int_rewrite; -(* A .& B .-> C becomes [| A[[s,t]]; B[[s,t]] |] ==> C[[s,t]] *) -fun action_conjmp th = zero_var_indexes (conjI RS (action_mp th)); - -(* A .& B .-> C becomes [| A[[s,t]]; B[[s,t]]; (C[[s,t]] ==> R) |] ==> R *) -fun action_conjimpE th = zero_var_indexes (conjI RS (action_impE th)); - -(* Turn A .= B into meta-level rewrite rule A == B *) -fun action_rewrite th = (rewrite_rule action_rews (th RS inteq_reflection)); +fun action_use th = + case (concl_of th) of + Const _ $ (Const ("Intensional.Valid", _) $ _) => + ((flatten (action_unlift th)) handle _ => th) + | _ => th; (* ===================== Update simpset and classical prover ============================= *) -(* Make the simplifier use action_unlift rather than int_unlift +(*** +(* Make the simplifier use action_use rather than int_use when action simplifications are added. *) -fun maybe_unlift th = - (case concl_of th of - Const("Intensional.TrueInt",_) $ p - => (action_unlift th - handle _ => int_unlift th) - | _ => th); - -simpset_ref() := simpset() setmksimps ((mksimps mksimps_pairs) o maybe_unlift); - -(* make act_rews be always active -- intensional_rews has been added before *) -Addsimps act_rews; - -use "cladata.ML"; (* local version! *) - -(* ================================ action_simp_tac ================================== *) - -(* A dumb simplification tactic with just a little first-order logic: - should plug in only "very safe" rules that can be applied blindly. - Note that it applies whatever simplifications are currently active. -*) -fun action_simp_tac ss intros elims i = - (asm_full_simp_tac - (ss setloop ((resolve_tac (intros @ [refl,impI,conjI,actionI,allI])) - ORELSE' (eresolve_tac (elims @ [conjE,disjE,exE_prop])))) - i); -(* default version without additional plug-in rules *) -fun Action_simp_tac i = (action_simp_tac (simpset()) [] [] i); - - -(* ==================== Simplification of abstractions ==================== *) - -(* Somewhat obscure simplifications, rarely necessary to get rid - of abstractions that may be introduced by higher-order unification. -*) -qed_goal "pr_con_abs" Action.thy "(%w. c)` .= #c" - (fn _ => [rtac actionI 1, - rewrite_goals_tac (con_abs::action_rews), - rtac refl 1 - ]); - -qed_goal "pr_lift_abs" Action.thy "(%w. f(x w))` .= f[x`]" - (fn _ => [rtac actionI 1, - (* give all rewrites to the engine and it loops! *) - rewrite_goals_tac intensional_rews, - rewtac lift_abs, - rewtac pr_lift, - rewtac unl_lift, - rtac refl 1 - ]); - -qed_goal "pr_lift2_abs" Action.thy "(%w. f(x w) (y w))` .= f[x`,y`]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift2_abs, - rewtac pr_lift2, - rewtac unl_lift2, - rtac refl 1 - ]); - -qed_goal "pr_lift2_abs_con1" Action.thy "(%w. f x (y w))` .= f[#x, y`]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift2_abs_con1, - rewtac pr_lift2, - rewtac unl_lift2, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift2_abs_con2" Action.thy "(%w. f (x w) y)` .= f[x`, #y]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift2_abs_con2, - rewtac pr_lift2, - rewtac unl_lift2, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift3_abs" Action.thy "(%w. f(x w) (y w) (z w))` .= f[x`,y`,z`]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs, - rewtac pr_lift3, - rewtac unl_lift3, - rtac refl 1 - ]); +let + val ss = simpset_ref() + fun try_rewrite th = + (action_rewrite th) handle _ => (action_use th) handle _ => th +in + ss := !ss setmksimps ((mksimps mksimps_pairs) o try_rewrite) +end; +***) -qed_goal "pr_lift3_abs_con1" Action.thy "(%w. f x (y w) (z w))` .= f[#x, y`, z`]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs_con1, - rewtac pr_lift3, - rewtac unl_lift3, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift3_abs_con2" Action.thy "(%w. f (x w) y (z w))` .= f[x`, #y, z`]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs_con2, - rewtac pr_lift3, - rewtac unl_lift3, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift3_abs_con3" Action.thy "(%w. f (x w) (y w) z)` .= f[x`, y`, #z]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs_con3, - rewtac pr_lift3, - rewtac unl_lift3, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift3_abs_con12" Action.thy "(%w. f x y (z w))` .= f[#x, #y, z`]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs_con12, - rewtac pr_lift3, - rewtac unl_lift3, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift3_abs_con13" Action.thy "(%w. f x (y w) z)` .= f[#x, y`, #z]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs_con13, - rewtac pr_lift3, - rewtac unl_lift3, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -qed_goal "pr_lift3_abs_con23" Action.thy "(%w. f (x w) y z)` .= f[x`, #y, #z]" - (fn _ => [rtac actionI 1, - rewrite_goals_tac intensional_rews, - rewtac lift3_abs_con23, - rewtac pr_lift3, - rewtac unl_lift3, - rewtac pr_con, - rewtac unl_con, - rtac refl 1 - ]); - -(* We don't add these as default rewrite rules, because they are - rarely needed and may slow down automatic proofs. -*) -val pr_abs_rews = map (fn th => th RS inteq_reflection) - [pr_con_abs, - pr_lift_abs,pr_lift2_abs,pr_lift2_abs_con1,pr_lift2_abs_con2, - pr_lift3_abs,pr_lift3_abs_con1,pr_lift3_abs_con2,pr_lift3_abs_con3, - pr_lift3_abs_con12,pr_lift3_abs_con13,pr_lift3_abs_con23]; +AddSIs [actionI]; +AddDs [actionD]; (* =========================== square / angle brackets =========================== *) qed_goalw "idle_squareI" Action.thy [square_def] - "!!s t. ([[s,t]] |= unchanged v) ==> ([[s,t]] |= [A]_v)" - (fn _ => [ Auto_tac ]); + "!!s t. (s,t) |= unchanged v ==> (s,t) |= [A]_v" + (fn _ => [ Asm_full_simp_tac 1 ]); qed_goalw "busy_squareI" Action.thy [square_def] - "!!s t. ([[s,t]] |= A) ==> ([[s,t]] |= [A]_v)" - (fn _ => [ Auto_tac ]); + "!!s t. (s,t) |= A ==> (s,t) |= [A]_v" + (fn _ => [ Asm_simp_tac 1 ]); + +qed_goal "squareE" Action.thy + "[| (s,t) |= [A]_v; A (s,t) ==> B (s,t); v t = v s ==> B (s,t) |] ==> B (s,t)" + (fn prems => [cut_facts_tac prems 1, + rewrite_goals_tac (square_def::action_rews), + etac disjE 1, + REPEAT (eresolve_tac prems 1)]); + +qed_goalw "squareCI" Action.thy (square_def::action_rews) + "[| v t ~= v s ==> A (s,t) |] ==> (s,t) |= [A]_v" + (fn prems => [rtac disjCI 1, + eresolve_tac prems 1]); + +qed_goalw "angleI" Action.thy [angle_def] + "!!s t. [| A (s,t); v t ~= v s |] ==> (s,t) |= _v" + (fn _ => [ Asm_simp_tac 1 ]); -qed_goalw "square_simulation" Action.thy [square_def] - "[| unchanged f .& .~B .-> unchanged g; \ -\ A .& .~unchanged g .-> B \ -\ |] ==> [A]_f .-> [B]_g" - (fn [p1,p2] => [Auto_tac, - etac (action_conjimpE p2) 1, - etac swap 3, etac (action_conjimpE p1) 3, - ALLGOALS atac - ]); - +qed_goalw "angleE" Action.thy (angle_def::action_rews) + "[| (s,t) |= _v; [| A (s,t); v t ~= v s |] ==> R |] ==> R" + (fn prems => [cut_facts_tac prems 1, + etac conjE 1, + REPEAT (ares_tac prems 1)]); + +AddIs [angleI, squareCI]; +AddEs [angleE, squareE]; + +qed_goal "square_simulation" Action.thy + "!!f. [| |- unchanged f & ~B --> unchanged g; \ +\ |- A & ~unchanged g --> B \ +\ |] ==> |- [A]_f --> [B]_g" + (fn _ => [Clarsimp_tac 1, + etac squareE 1, + auto_tac (claset(), simpset() addsimps [square_def]) + ]); + qed_goalw "not_square" Action.thy [square_def,angle_def] - "(.~ [A]_v) .= <.~A>_v" + "|- (~ [A]_v) = <~A>_v" (fn _ => [ Auto_tac ]); qed_goalw "not_angle" Action.thy [square_def,angle_def] - "(.~ _v) .= [.~A]_v" + "|- (~ _v) = [~A]_v" (fn _ => [ Auto_tac ]); (* ============================== Facts about ENABLED ============================== *) -qed_goalw "enabledI" Action.thy [enabled_def] - "A [[s,t]] ==> (Enabled A) s" - (fn prems => [ REPEAT (resolve_tac (exI::prems) 1) ]); +qed_goal "enabledI" Action.thy + "|- A --> $Enabled A" + (fn _ => [ auto_tac (claset(), simpset() addsimps [enabled_def]) ]); qed_goalw "enabledE" Action.thy [enabled_def] - "[| (Enabled A) s; !!u. A[[s,u]] ==> PROP R |] ==> PROP R" + "[| s |= Enabled A; !!u. A (s,u) ==> Q |] ==> Q" (fn prems => [cut_facts_tac prems 1, - etac exE_prop 1, + etac exE 1, resolve_tac prems 1, atac 1 ]); qed_goal "notEnabledD" Action.thy - "!!G. ~ (Enabled G s) ==> ~ G [[s,t]]" - (fn _ => [ auto_tac (action_css addsimps2 [enabled_def]) ]); + "|- ~$Enabled G --> ~ G" + (fn _ => [ auto_tac (claset(), simpset() addsimps [enabled_def]) ]); (* Monotonicity *) qed_goal "enabled_mono" Action.thy - "[| (Enabled F) s; F .-> G |] ==> (Enabled G) s" + "[| s |= Enabled F; |- F --> G |] ==> s |= Enabled G" (fn [min,maj] => [rtac (min RS enabledE) 1, - rtac enabledI 1, - etac (action_mp maj) 1 + rtac (action_use enabledI) 1, + etac (action_use maj) 1 ]); (* stronger variant *) qed_goal "enabled_mono2" Action.thy - "[| (Enabled F) s; !!t. (F [[s,t]] ==> G[[s,t]] ) |] ==> (Enabled G) s" + "[| s |= Enabled F; !!t. F (s,t) ==> G (s,t) |] ==> s |= Enabled G" (fn [min,maj] => [rtac (min RS enabledE) 1, - rtac enabledI 1, + rtac (action_use enabledI) 1, etac maj 1 ]); qed_goal "enabled_disj1" Action.thy - "!!s. (Enabled F) s ==> (Enabled (F .| G)) s" - (fn _ => [etac enabled_mono 1, Auto_tac - ]); + "|- Enabled F --> Enabled (F | G)" + (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]); qed_goal "enabled_disj2" Action.thy - "!!s. (Enabled G) s ==> (Enabled (F .| G)) s" - (fn _ => [etac enabled_mono 1, Auto_tac - ]); + "|- Enabled G --> Enabled (F | G)" + (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]); qed_goal "enabled_conj1" Action.thy - "!!s. (Enabled (F .& G)) s ==> (Enabled F) s" - (fn _ => [etac enabled_mono 1, Auto_tac - ]); + "|- Enabled (F & G) --> Enabled F" + (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]); qed_goal "enabled_conj2" Action.thy - "!!s. (Enabled (F .& G)) s ==> (Enabled G) s" - (fn _ => [etac enabled_mono 1, Auto_tac - ]); + "|- Enabled (F & G) --> Enabled G" + (fn _ => [ auto_tac (claset() addSEs [enabled_mono], simpset()) ]); qed_goal "enabled_conjE" Action.thy - "[| (Enabled (F .& G)) s; [| (Enabled F) s; (Enabled G) s |] ==> PROP R |] ==> PROP R" + "[| s |= Enabled (F & G); [| s |= Enabled F; s |= Enabled G |] ==> Q |] ==> Q" (fn prems => [cut_facts_tac prems 1, resolve_tac prems 1, - etac enabled_conj1 1, etac enabled_conj2 1]); + etac (action_use enabled_conj1) 1, + etac (action_use enabled_conj2) 1 + ]); qed_goal "enabled_disjD" Action.thy - "!!s. (Enabled (F .| G)) s ==> ((Enabled F) s) | ((Enabled G) s)" - (fn _ => [etac enabledE 1, - auto_tac (action_css addSDs2 [notEnabledD] addSEs2 [enabledI]) - ]); + "|- Enabled (F | G) --> Enabled F | Enabled G" + (fn _ => [ auto_tac (claset(), simpset() addsimps [enabled_def]) ]); qed_goal "enabled_disj" Action.thy - "(Enabled (F .| G)) s = ( (Enabled F) s | (Enabled G) s )" - (fn _ => [rtac iffI 1, - etac enabled_disjD 1, - REPEAT (eresolve_tac [disjE,enabled_disj1,enabled_disj2] 1) + "|- Enabled (F | G) = (Enabled F | Enabled G)" + (fn _ => [Clarsimp_tac 1, + rtac iffI 1, + etac (action_use enabled_disjD) 1, + REPEAT (eresolve_tac (disjE::map action_use [enabled_disj1,enabled_disj2]) 1) ]); qed_goal "enabled_ex" Action.thy - "(Enabled (REX x. F x)) s = (EX x. (Enabled (F x)) s)" - (fn _ => [ auto_tac (action_css addsimps2 [enabled_def]) ]); - + "|- Enabled (? x. F x) = (? x. Enabled (F x))" + (fn _ => [ force_tac (claset(), simpset() addsimps [enabled_def]) 1 ]); + (* A rule that combines enabledI and baseE, but generates fewer possible instantiations *) qed_goal "base_enabled" Action.thy - "[| base_var(v); !!u. v u = c s ==> A [[s,u]] |] ==> Enabled A s" + "[| basevars vs; !!u. vs u = c s ==> A (s,u) |] ==> s |= Enabled A" (fn prems => [cut_facts_tac prems 1, - etac baseE 1, rtac enabledI 1, + etac baseE 1, rtac (action_use enabledI) 1, REPEAT (ares_tac prems 1)]); +(* ================================ action_simp_tac ================================== *) + +(* A dumb simplification-based tactic with just a little first-order logic: + should plug in only "very safe" rules that can be applied blindly. + Note that it applies whatever simplifications are currently active. +*) +fun action_simp_tac ss intros elims = + asm_full_simp_tac + (ss setloop ((resolve_tac ((map action_use intros) + @ [refl,impI,conjI,actionI,intI,allI])) + ORELSE' (eresolve_tac ((map action_use elims) + @ [conjE,disjE,exE])))); + +(* default version without additional plug-in rules *) +val Action_simp_tac = action_simp_tac (simpset()) [] []; + + + (* ---------------- enabled_tac: tactic to prove (Enabled A) -------------------- *) (* "Enabled A" can be proven as follows: - Assume that we know which state variables are "base variables"; - this should be expressed by a theorem of the form "base_var ". + this should be expressed by a theorem of the form "basevars (x,y,z,...)". - Resolve this theorem with baseE to introduce a constant for the value of the variables in the successor state, and resolve the goal with the result. - - E-resolve with PairVarE so that we have one constant per variable. - Resolve with enabledI and do some rewriting. - Solve for the unknowns using standard HOL reasoning. The following tactic combines these steps except the final one. *) - +(*** old version fun enabled_tac base_vars i = EVERY [(* apply actionI (plus rewriting) if the goal is of the form $(Enabled A), - do nothing if it is of the form (Enabled A) s *) - TRY ((rtac actionI i) THEN (SELECT_GOAL (rewrite_goals_tac action_rews) i)), - rtac (base_vars RS base_enabled) i, - REPEAT_DETERM (etac PairVarE i), + do nothing if it is of the form s |= Enabled A *) + TRY ((resolve_tac [actionI,intI] i) + THEN (SELECT_GOAL (rewrite_goals_tac action_rews) i)), + clarify_tac (claset() addSIs [base_vars RS base_enabled]) i, (SELECT_GOAL (rewrite_goals_tac action_rews) i) ]; +***) -(* Example of use: +fun enabled_tac base_vars = + clarsimp_tac (claset() addSIs [base_vars RS base_enabled], simpset()); -val [prem] = goal Action.thy "base_var ==> $x .-> $Enabled ($x .& (y$ .= #False))"; -by (REPEAT ((CHANGED (Action_simp_tac 1)) ORELSE (enabled_tac prem 1))); +(* Example: + +val [prem] = goal thy "basevars (x,y,z) ==> |- x --> Enabled ($x & (y$ = #False))"; +by (enabled_tac prem 1); +auto(); *) diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Action.thy --- a/src/HOL/TLA/Action.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Action.thy Mon Feb 08 13:02:56 1999 +0100 @@ -1,7 +1,7 @@ (* File: TLA/Action.thy Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Theory Name: Action Logic Image: HOL @@ -11,50 +11,65 @@ Action = Intensional + Stfun + +(** abstract syntax **) + types - state2 (* intention: pair of states *) - 'a trfct = "('a, state2) term" - action = "state2 form" + 'a trfun = "(state * state) => 'a" + action = bool trfun + +instance + "*" :: (world, world) world -arities - state2 :: world - consts - mkstate2 :: "[state,state] => state2" ("([[_,_]])") + (** abstract syntax **) + before :: 'a stfun => 'a trfun + after :: 'a stfun => 'a trfun + unch :: 'a stfun => action + + SqAct :: [action, 'a stfun] => action + AnAct :: [action, 'a stfun] => action + enabled :: action => stpred + +(** concrete syntax **) + +syntax + (* Syntax for writing action expressions in arbitrary contexts *) + "ACT" :: lift => 'a ("(ACT _)") - (* lift state variables to transition functions *) - before :: "'a stfun => 'a trfct" ("($_)" [100] 99) - after :: "'a stfun => 'a trfct" ("(_$)" [100] 99) - unchanged :: "'a stfun => action" + "_before" :: lift => lift ("($_)" [100] 99) + "_after" :: lift => lift ("(_$)" [100] 99) + "_unchanged" :: lift => lift ("(unchanged _)" [100] 99) + + (*** Priming: same as "after" ***) + "_prime" :: lift => lift ("(_`)" [100] 99) + + "_SqAct" :: [lift, lift] => lift ("([_]'_(_))" [0,1000] 99) + "_AnAct" :: [lift, lift] => lift ("(<_>'_(_))" [0,1000] 99) + "_Enabled" :: lift => lift ("(Enabled _)" [100] 100) - (* Priming *) - prime :: "'a trfct => 'a trfct" ("(_`)" [90] 89) - - SqAct :: "[action, 'a stfun] => action" ("([_]'_(_))" [0,60] 59) - AnAct :: "[action, 'a stfun] => action" ("(<_>'_(_))" [0,60] 59) - Enabled :: "action => stpred" +translations + "ACT A" => "(A::state*state => _)" + "_before" == "before" + "_after" => "_prime" + "_unchanged" == "unch" + "_prime" == "after" + "_SqAct" == "SqAct" + "_AnAct" == "AnAct" + "_Enabled" == "enabled" + "w |= [A]_v" <= "_SqAct A v w" + "w |= _v" <= "_AnAct A v w" + "s |= Enabled A" <= "_Enabled A s" + "w |= unchanged f" <= "_unchanged f w" rules - (* The following says that state2 is generated by mkstate2 *) - state2_ext "(!!s t. [[s,t]] |= (A::action)) ==> (st::state2) |= A" - - unl_before "($v) [[s,t]] == v s" - unl_after "(v$) [[s,t]] == v t" + unl_before "(ACT $v) (s,t) == v s" + unl_after "(ACT v`) (s,t) == v t" - pr_con "(#c)` == #c" - pr_before "($v)` == v$" - (* no corresponding rule for "after"! *) - pr_lift "(F[x])` == F[x`]" - pr_lift2 "(F[x,y])` == F[x`,y`]" - pr_lift3 "(F[x,y,z])` == F[x`,y`,z`]" - pr_all "(RALL x. P(x))` == (RALL x. P(x)`)" - pr_ex "(REX x. P(x))` == (REX x. P(x)`)" + unchanged_def "(s,t) |= unchanged v == (v t = v s)" + square_def "ACT [A]_v == ACT (A | unchanged v)" + angle_def "ACT _v == ACT (A & ~ unchanged v)" - unchanged_def "(unchanged v) [[s,t]] == (v t = v s)" - square_def "[A]_v == A .| unchanged v" - angle_def "_v == A .& .~ unchanged v" - - enabled_def "(Enabled A) s == EX u. A[[s,u]]" + enabled_def "s |= Enabled A == EX u. (s,u) |= A" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Buffer/Buffer.ML --- a/src/HOL/TLA/Buffer/Buffer.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Buffer/Buffer.ML Mon Feb 08 13:02:56 1999 +0100 @@ -8,31 +8,29 @@ (* ---------------------------- Data lemmas ---------------------------- *) -(* "xs ~= [] --> tl(xs @ ys) = (tl xs) @ ys" *) -Addsimps [tl_append2]; - +context List.thy; goal List.thy "xs ~= [] --> tl xs ~= xs"; by (auto_tac (claset(), simpset() addsimps [neq_Nil_conv])); qed_spec_mp "tl_not_self"; -Addsimps [tl_not_self]; +context Buffer.thy; -(* "!ys zs. (ys @ xs = zs @ xs) = (ys=zs)" has been subsumed *) +Addsimps [tl_not_self]; (* ---------------------------- Action lemmas ---------------------------- *) (* Dequeue is visible *) -Goal "_ .= Deq ic q oc"; +Goal "|- _(ic,q,oc) = Deq ic q oc"; by (auto_tac (claset(), simpset() addsimps [angle_def,Deq_def])); qed "Deq_visible"; (* Enabling condition for dequeue -- NOT NEEDED *) Goalw [temp_rewrite Deq_visible] - "!!q. base_var ==> $Enabled (_) .= ($q .~= .[])"; + "!!q. basevars (ic,q,oc) ==> |- Enabled (_(ic,q,oc)) = (q ~= #[])"; by (force_tac (claset() addSEs [base_enabled,enabledE], simpset() addsimps [Deq_def]) 1); qed "Deq_enabled"; (* For the left-to-right implication, we don't need the base variable stuff *) Goalw [temp_rewrite Deq_visible] - "$Enabled (_) .-> ($q .~= .[])"; + "|- Enabled (_(ic,q,oc)) --> (q ~= #[])"; by (auto_tac (claset() addSEs [enabledE], simpset() addsimps [Deq_def])); qed "Deq_enabledE"; diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Buffer/Buffer.thy --- a/src/HOL/TLA/Buffer/Buffer.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Buffer/Buffer.thy Mon Feb 08 13:02:56 1999 +0100 @@ -9,16 +9,11 @@ A simple FIFO buffer (synchronous communication, interleaving) *) -Buffer = TLA + List + +Buffer = TLA + consts - (* infix syntax for list operations *) - "IntNil" :: 'w::world => 'a list (".[]") - "IntCons" :: ['w::world => 'a, 'w => 'a list] => ('w => 'a list) ("(_ .#/ _)" [65,66] 65) - "IntApp" :: ['w::world => 'a list, 'w => 'a list] => ('w => 'a list) ("(_ .@/ _)" [65,66] 65) - (* actions *) - BInit :: "'a stfun => 'a list stfun => 'a stfun => action" + BInit :: "'a stfun => 'a list stfun => 'a stfun => stpred" Enq :: "'a stfun => 'a list stfun => 'a stfun => action" Deq :: "'a stfun => 'a list stfun => 'a stfun => action" Next :: "'a stfun => 'a list stfun => 'a stfun => action" @@ -27,30 +22,20 @@ IBuffer :: "'a stfun => 'a list stfun => 'a stfun => temporal" Buffer :: "'a stfun => 'a stfun => temporal" -syntax - "@listInt" :: args => ('a list, 'w) term (".[(_)]") - -translations - ".[]" == "con []" - "x .# xs" == "lift2 (op #) x xs" - "xs .@ ys" == "lift2 (op @) xs ys" - ".[ x, xs ]" == "x .# .[xs]" - ".[ x ]" == "x .# .[]" - rules - BInit_def "BInit ic q oc == $q .= .[]" - Enq_def "Enq ic q oc == (ic$ .~= $ic) - .& (q$ .= $q .@ .[ ic$ ]) - .& (oc$ .= $oc)" - Deq_def "Deq ic q oc == ($q .~= .[]) - .& (oc$ .= hd[ $q ]) - .& (q$ .= tl[ $q ]) - .& (ic$ .= $ic)" - Next_def "Next ic q oc == Enq ic q oc .| Deq ic q oc" - IBuffer_def "IBuffer ic q oc == Init (BInit ic q oc) - .& [][Next ic q oc]_ - .& WF(Deq ic q oc)_" - Buffer_def "Buffer ic oc == EEX q. IBuffer ic q oc" + BInit_def "BInit ic q oc == PRED q = #[]" + Enq_def "Enq ic q oc == ACT (ic$ ~= $ic) + & (q$ = $q @ [ ic$ ]) + & (oc$ = $oc)" + Deq_def "Deq ic q oc == ACT ($q ~= #[]) + & (oc$ = hd< $q >) + & (q$ = tl< $q >) + & (ic$ = $ic)" + Next_def "Next ic q oc == ACT (Enq ic q oc | Deq ic q oc)" + IBuffer_def "IBuffer ic q oc == TEMP Init (BInit ic q oc) + & [][Next ic q oc]_(ic,q,oc) + & WF(Deq ic q oc)_(ic,q,oc)" + Buffer_def "Buffer ic oc == TEMP (EEX q. IBuffer ic q oc)" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Buffer/DBuffer.ML --- a/src/HOL/TLA/Buffer/DBuffer.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Buffer/DBuffer.ML Mon Feb 08 13:02:56 1999 +0100 @@ -6,6 +6,7 @@ Double FIFO buffer implements simple FIFO buffer. *) + val db_css = (claset(), simpset() addsimps [qc_def]); Addsimps [qc_def]; @@ -14,15 +15,15 @@ (*** Proper initialization ***) -Goal "Init DBInit .-> Init (BInit inp qc out)"; +Goal "|- Init DBInit --> Init (BInit inp qc out)"; by (auto_tac (db_css addsimps2 [Init_def,DBInit_def,BInit_def])); qed "DBInit"; (*** Step simulation ***) -Goal "[DBNext]_ .-> [Next inp qc out]_"; +Goal "|- [DBNext]_(inp,mid,out,q1,q2) --> [Next inp qc out]_(inp,qc,out)"; by (rtac square_simulation 1); -by (Action_simp_tac 1); +by (Clarsimp_tac 1); by (action_simp_tac (simpset() addsimps hd_append::db_defs) [] [] 1); qed "DB_step_simulation"; @@ -30,25 +31,23 @@ (*** Simulation of fairness ***) (* Compute enabledness predicates for DBDeq and DBPass actions *) -Goal "_ .= DBDeq"; +Goal "|- _(inp,mid,out,q1,q2) = DBDeq"; by (auto_tac (db_css addsimps2 [angle_def,DBDeq_def,Deq_def])); qed "DBDeq_visible"; -Goal "$Enabled (_) .= ($q2 .~= .[])"; -by (rewtac (action_rewrite DBDeq_visible)); -by (cut_facts_tac [DB_base] 1); -by (old_auto_tac (db_css addSEs2 [base_enabled,enabledE] - addsimps2 [angle_def,DBDeq_def,Deq_def])); +Goalw [action_rewrite DBDeq_visible] + "|- Enabled (_(inp,mid,out,q1,q2)) = (q2 ~= #[])"; +by (force_tac (db_css addSIs2 [DB_base RS base_enabled] addSEs2 [enabledE] + addsimps2 [angle_def,DBDeq_def,Deq_def]) 1); qed "DBDeq_enabled"; -Goal "_ .= DBPass"; +Goal "|- _(inp,mid,out,q1,q2) = DBPass"; by (auto_tac (db_css addsimps2 [angle_def,DBPass_def,Deq_def])); qed "DBPass_visible"; -Goal "$Enabled (_) .= ($q1 .~= .[])"; -by (rewtac (action_rewrite DBPass_visible)); -by (cut_facts_tac [DB_base] 1); -by (force_tac (db_css addSEs2 [base_enabled,enabledE] +Goalw [action_rewrite DBPass_visible] + "|- Enabled (_(inp,mid,out,q1,q2)) = (q1 ~= #[])"; +by (force_tac (db_css addSIs2 [DB_base RS base_enabled] addSEs2 [enabledE] addsimps2 [angle_def,DBPass_def,Deq_def]) 1); qed "DBPass_enabled"; @@ -58,8 +57,8 @@ which is in turn reduced to the two leadsto conditions (1) DBuffer => (Enabled (Deq inp qc out) ~> q2 ~= []) (2) DBuffer => (q2 ~= [] ~> DBDeq) - and the fact that DBDeq implies _ - (and therefore DBDeq ~> _ trivially holds). + and the fact that DBDeq implies _(inp,qc,out) + (and therefore DBDeq ~> _(inp,qc,out) trivially holds). Condition (1) is reduced to (1a) DBuffer => (qc ~= [] /\ q2 = [] ~> q2 ~= []) @@ -73,47 +72,47 @@ *) (* Condition (1a) *) -Goal - "[][DBNext]_ .& WF(DBPass)_ \ -\ .-> ($qc .~= .[] .& $q2 .= .[] ~> $q2 .~= .[])"; +Goal "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2) \ +\ --> (qc ~= #[] & q2 = #[] ~> q2 ~= #[])"; by (rtac WF1 1); -by (action_simp_tac (simpset() addsimps square_def::db_defs) [] [] 1); -by (action_simp_tac (simpset() addsimps [angle_def,DBPass_def]) [] [] 1); -by (action_simp_tac (simpset() addsimps [DBPass_enabled]) [] [] 1); +by (force_tac (db_css addsimps2 db_defs) 1); +by (force_tac (db_css addsimps2 [angle_def,DBPass_def]) 1); +by (force_tac (db_css addsimps2 [DBPass_enabled]) 1); qed "DBFair_1a"; (* Condition (1) *) -Goal - "[][DBNext]_ .& WF(DBPass)_ \ -\ .-> ($Enabled (_) ~> $q2 .~= .[])"; -by (auto_tac (temp_css addSIs2 [leadsto_classical] addSEs2 [temp_conjimpE DBFair_1a])); -by (force_tac (temp_css addsimps2 [leadsto,Init_def] addDs2 [STL2bD] - addSDs2 [action_mp Deq_enabledE] addSEs2 [STL4E]) 1); +Goal "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2) \ +\ --> (Enabled (_(inp,qc,out)) ~> q2 ~= #[])"; +by (Clarsimp_tac 1); +by (rtac (temp_use leadsto_classical) 1); +by (rtac ((temp_use DBFair_1a) RS (temp_use LatticeTransitivity)) 1); +by (TRYALL atac); +by (rtac (temp_use ImplLeadsto_gen) 1); +by (force_tac (db_css addSIs2 [necT] addSDs2 [STL2_gen, Deq_enabledE] + addsimps2 Init_defs) 1); qed "DBFair_1"; (* Condition (2) *) -Goal - "[][DBNext]_ .& WF(DBDeq)_ \ -\ .-> ($q2 .~= .[] ~> DBDeq)"; +Goal "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBDeq)_(inp,mid,out,q1,q2) \ +\ --> (q2 ~= #[] ~> DBDeq)"; by (rtac WF_leadsto 1); -by (action_simp_tac (simpset() addsimps [DBDeq_visible,DBDeq_enabled]) [] [] 1); -by (action_simp_tac (simpset() addsimps [angle_def]) [] [] 1); -by (action_simp_tac (simpset() addsimps square_def::db_defs) [tempI] [Stable] 1); +by (force_tac (db_css addsimps2 [DBDeq_enabled]) 1); +by (force_tac (db_css addsimps2 [angle_def]) 1); +by (force_tac (db_css addsimps2 db_defs addSEs2 [Stable]) 1); qed "DBFair_2"; (* High-level fairness *) -Goal - "[][DBNext]_ .& WF(DBPass)_ \ -\ .& WF(DBDeq)_ \ -\ .-> WF(Deq inp qc out)_"; -by (auto_tac (db_css addSIs2 [leadsto_WF])); -by (auto_tac (db_css addSIs2 [(temp_mp DBFair_1) RSN(2,LatticeTransitivity)])); -by (auto_tac (db_css addSIs2 [(temp_mp DBFair_2) RSN(2,LatticeTransitivity)])); -by (auto_tac (db_css addSIs2 [ImplLeadsto] addSEs2 [STL4E] +Goal "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2) \ +\ & WF(DBDeq)_(inp,mid,out,q1,q2) \ +\ --> WF(Deq inp qc out)_(inp,qc,out)"; +by (auto_tac (temp_css addSIs2 [leadsto_WF, + (temp_use DBFair_1) RSN(2,(temp_use LatticeTransitivity)), + (temp_use DBFair_2) RSN(2,(temp_use LatticeTransitivity))])); +by (auto_tac (db_css addSIs2 [ImplLeadsto_simple] addsimps2 [angle_def,DBDeq_def,Deq_def,hd_append])); qed "DBFair"; (*** Main theorem ***) -Goalw [DBuffer_def,Buffer_def,IBuffer_def] "DBuffer .-> Buffer inp out"; -by (ALLGOALS (force_tac (db_css addSIs2 (map temp_mp [eexI,DBInit,DB_step_simulation RS STL4,DBFair])))); +Goalw [DBuffer_def,Buffer_def,IBuffer_def] "|- DBuffer --> Buffer inp out"; +by (force_tac (temp_css addSIs2 [eexI,DBInit,DB_step_simulation RS STL4,DBFair]) 1); qed "DBuffer_impl_Buffer"; diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Buffer/DBuffer.thy --- a/src/HOL/TLA/Buffer/DBuffer.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Buffer/DBuffer.thy Mon Feb 08 13:02:56 1999 +0100 @@ -16,25 +16,26 @@ inp, mid, out :: nat stfun q1, q2, qc :: nat list stfun - DBInit, DBEnq, DBDeq, DBPass, DBNext :: action - DBuffer :: temporal + DBInit :: stpred + DBEnq, DBDeq, DBPass, DBNext :: action + DBuffer :: temporal rules - DB_base "base_var " + DB_base "basevars (inp,mid,out,q1,q2)" (* the concatenation of the two buffers *) - qc_def "$qc .= $q2 .@ $q1" + qc_def "PRED qc == PRED (q2 @ q1)" - DBInit_def "DBInit == BInit inp q1 mid .& BInit mid q2 out" - DBEnq_def "DBEnq == Enq inp q1 mid .& unchanged " - DBDeq_def "DBDeq == Deq mid q2 out .& unchanged " - DBPass_def "DBPass == Deq inp q1 mid - .& (q2$ .= $q2 .@ .[ mid$ ]) - .& (out$ .= $out)" - DBNext_def "DBNext == DBEnq .| DBDeq .| DBPass" - DBuffer_def "DBuffer == Init(DBInit) - .& [][DBNext]_ - .& WF(DBDeq)_ - .& WF(DBPass)_" + DBInit_def "DBInit == PRED (BInit inp q1 mid & BInit mid q2 out)" + DBEnq_def "DBEnq == ACT Enq inp q1 mid & unchanged (q2,out)" + DBDeq_def "DBDeq == ACT Deq mid q2 out & unchanged (inp,q1)" + DBPass_def "DBPass == ACT Deq inp q1 mid + & (q2$ = $q2 @ [ mid$ ]) + & (out$ = $out)" + DBNext_def "DBNext == ACT (DBEnq | DBDeq | DBPass)" + DBuffer_def "DBuffer == TEMP Init DBInit + & [][DBNext]_(inp,mid,out,q1,q2) + & WF(DBDeq)_(inp,mid,out,q1,q2) + & WF(DBPass)_(inp,mid,out,q1,q2)" end \ No newline at end of file diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Buffer/index.html --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/TLA/Buffer/index.html Mon Feb 08 13:02:56 1999 +0100 @@ -0,0 +1,8 @@ +buffer +

buffer

+The name of every theory is linked to its theory file
+\/
stands for subtheories (child theories)
+/\ stands for supertheories (parent theories)

+Back to the index of ex +


\//\ Buffer
+\//\ DBuffer
diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Inc/Inc.ML --- a/src/HOL/TLA/Inc/Inc.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Inc/Inc.ML Mon Feb 08 13:02:56 1999 +0100 @@ -14,65 +14,61 @@ (*** Invariant proof for Psi: "manual" proof proves individual lemmas ***) -qed_goal "PsiInv_Init" Inc.thy "InitPsi .-> PsiInv" +qed_goal "PsiInv_Init" Inc.thy "|- InitPsi --> PsiInv" (fn _ => [ auto_tac (Inc_css addsimps2 InitPsi_def::PsiInv_defs) ]); -qed_goal "PsiInv_alpha1" Inc.thy "alpha1 .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_alpha1" Inc.thy "|- alpha1 & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 alpha1_def::PsiInv_defs) ]); -qed_goal "PsiInv_alpha2" Inc.thy "alpha2 .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_alpha2" Inc.thy "|- alpha2 & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 alpha2_def::PsiInv_defs) ]); -qed_goal "PsiInv_beta1" Inc.thy "beta1 .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_beta1" Inc.thy "|- beta1 & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 beta1_def::PsiInv_defs) ]); -qed_goal "PsiInv_beta2" Inc.thy "beta2 .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_beta2" Inc.thy "|- beta2 & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 beta2_def::PsiInv_defs) ]); -qed_goal "PsiInv_gamma1" Inc.thy "gamma1 .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_gamma1" Inc.thy "|- gamma1 & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 gamma1_def::PsiInv_defs) ]); -qed_goal "PsiInv_gamma2" Inc.thy "gamma2 .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_gamma2" Inc.thy "|- gamma2 & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 gamma2_def::PsiInv_defs) ]); -qed_goal "PsiInv_stutter" Inc.thy "unchanged .& PsiInv .-> PsiInv`" +qed_goal "PsiInv_stutter" Inc.thy "|- unchanged (x,y,sem,pc1,pc2) & $PsiInv --> PsiInv`" (fn _ => [ auto_tac (Inc_css addsimps2 PsiInv_defs) ]); -qed_goal "PsiInv" Inc.thy "Psi .-> []PsiInv" (K [ +qed_goal "PsiInv" Inc.thy "|- Psi --> []PsiInv" (K [ inv_tac (Inc_css addsimps2 [Psi_def]) 1, - SELECT_GOAL (auto_tac (Inc_css addSIs2 [action_mp PsiInv_Init] - addsimps2 [Init_def])) 1, - force_tac (Inc_css addSEs2 (map action_conjimpE - [PsiInv_alpha1,PsiInv_alpha2,PsiInv_beta1, - PsiInv_beta2,PsiInv_gamma1,PsiInv_gamma2]) - addIs2 [action_mp PsiInv_stutter] - addsimps2 [square_def,N1_def, N2_def]) 1]); - - + force_tac (Inc_css addsimps2 [PsiInv_Init, Init_def]) 1, + auto_tac (Inc_css addIs2 + [PsiInv_alpha1,PsiInv_alpha2,PsiInv_beta1, + PsiInv_beta2,PsiInv_gamma1,PsiInv_gamma2,PsiInv_stutter] + addsimps2 [square_def,N1_def, N2_def]) ]); (* Automatic proof works too, but it make take a while on a slow machine. - More substantial examples require manual guidance anyway. + More realistic examples require user guidance anyway. -Goal "Psi .-> []PsiInv"; -by (auto_inv_tac (simpset() addsimps PsiInv_defs @ Psi_defs @ pcount.simps) 1); +Goal "|- Psi --> []PsiInv"; +by (auto_inv_tac (simpset() addsimps PsiInv_defs @ Psi_defs) 1); *) (**** Step simulation ****) -qed_goal "Init_sim" Inc.thy "Psi .-> Init(InitPhi)" +qed_goal "Init_sim" Inc.thy "|- Psi --> Init InitPhi" (fn _ => [ auto_tac (Inc_css addsimps2 [InitPhi_def,Psi_def,InitPsi_def,Init_def]) ]); -qed_goal "Step_sim" Inc.thy "Psi .-> [][M1 .| M2]_" +qed_goal "Step_sim" Inc.thy "|- Psi --> [][M1 | M2]_(x,y)" (fn _ => [auto_tac (Inc_css addsimps2 [square_def,M1_def,M2_def] @ Psi_defs - addSEs2 [STL4E]) + addSEs2 [STL4E]) ]); (**** Proof of fairness ****) (* The goal is to prove Fair_M1 far below, which asserts - Psi .-> WF(M1)_ + |- Psi --> WF(M1)_(x,y) (the other fairness condition is symmetrical). The strategy is to use WF2 (with beta1 as the helpful action). Proving its @@ -87,97 +83,94 @@ *) qed_goal "Stuck_at_b" Inc.thy - "[][(N1 .| N2) .& .~ beta1]_ .-> stable($pc1 .= #b)" - (fn _ => [rtac StableL 1, - auto_tac (Inc_css addsimps2 square_def::Psi_defs) - ]); + "|- [][(N1 | N2) & ~ beta1]_(x,y,sem,pc1,pc2) --> stable(pc1 = #b)" + (fn _ => [ auto_tac (Inc_css addSEs2 [Stable,squareE] addsimps2 Psi_defs) ]); qed_goal "N1_enabled_at_g" Inc.thy - "($pc1 .= #g) .-> $(Enabled (_))" - (fn _ => [Action_simp_tac 1, + "|- pc1 = #g --> Enabled (_(x,y,sem,pc1,pc2))" + (fn _ => [Clarsimp_tac 1, res_inst_tac [("F","gamma1")] enabled_mono 1, enabled_tac Inc_base 1, auto_tac (Inc_css addsimps2 [angle_def,gamma1_def,N1_def]) ]); qed_goal "g1_leadsto_a1" Inc.thy - "[][(N1 .| N2) .& .~ beta1]_ .& SF(N1)_ .& []#True \ -\ .-> ($pc1 .= #g ~> $pc1 .= #a)" + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N1)_(x,y,sem,pc1,pc2) & []#True \ +\ --> (pc1 = #g ~> pc1 = #a)" (fn _ => [rtac SF1 1, - (* the first two subgoals are simple action formulas and succumb to the - auto_tac; don't expand N1 in the third subgoal *) - SELECT_GOAL (auto_tac (Inc_css addsimps2 [square_def] @ Psi_defs)) 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [angle_def] @ Psi_defs)) 1, - (* reduce []A .-> <>Enabled B to A .-> Enabled B *) - auto_tac (Inc_css addSIs2 [InitDmdD, action_mp N1_enabled_at_g] - addSDs2 [STL2bD] + action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1, + action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1, + (* reduce |- []A --> <>Enabled B to |- A --> Enabled B *) + auto_tac (Inc_css addSIs2 [InitDmd_gen, N1_enabled_at_g] + addSDs2 [STL2_gen] addsimps2 [Init_def]) ]); (* symmetrical for N2, and similar for beta2 *) qed_goal "N2_enabled_at_g" Inc.thy - "($pc2 .= #g) .-> $(Enabled (_))" - (fn _ => [Action_simp_tac 1, + "|- pc2 = #g --> Enabled (_(x,y,sem,pc1,pc2))" + (fn _ => [Clarsimp_tac 1, res_inst_tac [("F","gamma2")] enabled_mono 1, enabled_tac Inc_base 1, auto_tac (Inc_css addsimps2 [angle_def,gamma2_def,N2_def]) ]); qed_goal "g2_leadsto_a2" Inc.thy - "[][(N1 .| N2) .& .~ beta1]_ .& SF(N2)_ .& []#True \ -\ .-> ($pc2 .= #g ~> $pc2 .= #a)" + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \ +\ --> (pc2 = #g ~> pc2 = #a)" (fn _ => [rtac SF1 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [square_def] @ Psi_defs)) 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [angle_def] @ Psi_defs)) 1, - auto_tac (Inc_css addSIs2 [InitDmdD, action_mp N2_enabled_at_g] - addSDs2 [STL2bD] + action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1, + action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1, + auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_g] + addSDs2 [STL2_gen] addsimps2 [Init_def]) ]); qed_goal "N2_enabled_at_b" Inc.thy - "($pc2 .= #b) .-> $(Enabled (_))" - (fn _ => [Action_simp_tac 1, + "|- pc2 = #b --> Enabled (_(x,y,sem,pc1,pc2))" + (fn _ => [Clarsimp_tac 1, res_inst_tac [("F","beta2")] enabled_mono 1, enabled_tac Inc_base 1, auto_tac (Inc_css addsimps2 [angle_def,beta2_def,N2_def]) ]); qed_goal "b2_leadsto_g2" Inc.thy - "[][(N1 .| N2) .& .~ beta1]_ .& SF(N2)_ .& []#True \ -\ .-> ($pc2 .= #b ~> $pc2 .= #g)" + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \ +\ --> (pc2 = #b ~> pc2 = #g)" (fn _ => [rtac SF1 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [square_def] @ Psi_defs)) 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [angle_def] @ Psi_defs)) 1, - auto_tac (Inc_css addSIs2 [InitDmdD, action_mp N2_enabled_at_b] - addSDs2 [STL2bD] + action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1, + action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1, + auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_b] + addSDs2 [STL2_gen] addsimps2 [Init_def]) ]); (* Combine above lemmas: the second component will eventually reach pc2 = a *) qed_goal "N2_leadsto_a" Inc.thy - "[][(N1 .| N2) .& .~ beta1]_ .& SF(N2)_ .& []#True \ -\ .-> (($pc2 .= #a .| $pc2 .= #b .| $pc2 .= #g) ~> $pc2 .= #a)" + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \ +\ --> (pc2 = #a | pc2 = #b | pc2 = #g ~> pc2 = #a)" (fn _ => [auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro]), - rtac (LatticeReflexivity RS tempD) 1, - rtac LatticeTransitivity 1, - auto_tac (Inc_css addSIs2 (map temp_mp [b2_leadsto_g2,g2_leadsto_a2])) + rtac (temp_use LatticeReflexivity) 1, + rtac (temp_use LatticeTransitivity) 1, + auto_tac (Inc_css addSIs2 [b2_leadsto_g2,g2_leadsto_a2]) ]); -(* A variant that gets rid of the disjunction, thanks to induction over data types *) +(* Get rid of complete disjunction on the left-hand side of ~> above. *) qed_goal "N2_live" Inc.thy - "[][(N1 .| N2) .& .~ beta1]_ .& SF(N2)_ \ -\ .-> <>($pc2 .= #a)" - (fn _ => [auto_tac (Inc_css addSIs2 [(temp_mp N2_leadsto_a) RSN(2,leadsto_init)]), - rewrite_goals_tac (Init_def::action_rews), - exhaust_tac "pc2 (fst_st sigma)" 1, + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) \ +\ --> <>(pc2 = #a)" + (fn _ => [auto_tac (Inc_css addsimps2 Init_defs + addSIs2 [(temp_use N2_leadsto_a) + RSN(2, (temp_use leadsto_init))]), + exhaust_tac "pc2 (st1 sigma)" 1, Auto_tac ]); (* Now prove that the first component will eventually reach pc1 = b from pc1 = a *) qed_goal "N1_enabled_at_both_a" Inc.thy - "$pc2 .= #a .& (PsiInv .& $pc1 .= #a) .-> $(Enabled (_))" - (fn _ => [Action_simp_tac 1, + "|- pc2 = #a & (PsiInv & pc1 = #a) --> Enabled (_(x,y,sem,pc1,pc2))" + (fn _ => [Clarsimp_tac 1, res_inst_tac [("F","alpha1")] enabled_mono 1, enabled_tac Inc_base 1, auto_tac (Inc_css addIs2 [sym] @@ -185,43 +178,44 @@ ]); qed_goal "a1_leadsto_b1" Inc.thy - "[](PsiInv .& [(N1 .| N2) .& .~ beta1]_) \ -\ .& SF(N1)_ .& [] SF(N2)_ \ -\ .-> ($pc1 .= #a ~> $pc1 .= #b)" + "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2)) \ +\ & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2) \ +\ --> (pc1 = #a ~> pc1 = #b)" (fn _ => [rtac SF1 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [square_def] @ Psi_defs)) 1, - SELECT_GOAL (auto_tac (Inc_css addsimps2 [angle_def] @ Psi_defs)) 1, - auto_tac (Inc_css addSIs2 [N1_enabled_at_both_a RS (temp_mp DmdImpl)]), - auto_tac (Inc_css addSIs2 [temp_mp BoxDmdT2, temp_mp N2_live] + action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1, + action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1, + clarsimp_tac (Inc_css addSIs2 [N1_enabled_at_both_a RS (temp_use DmdImpl)]) 1, + auto_tac (Inc_css addSIs2 [BoxDmd2_simple, N2_live] addsimps2 split_box_conj::more_temp_simps) ]); (* Combine the leadsto properties for N1: it will arrive at pc1 = b *) qed_goal "N1_leadsto_b" Inc.thy - "[](PsiInv .& [(N1 .| N2) .& .~ beta1]_) \ -\ .& SF(N1)_ .& [] SF(N2)_ \ -\ .-> (($pc1 .= #b .| $pc1 .= #g .| $pc1 .= #a) ~> $pc1 .= #b)" + "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2)) \ +\ & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2) \ +\ --> (pc1 = #b | pc1 = #g | pc1 = #a ~> pc1 = #b)" (fn _ => [auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro]), - rtac (LatticeReflexivity RS tempD) 1, - rtac LatticeTransitivity 1, - auto_tac (Inc_css addSIs2 (map temp_mp [a1_leadsto_b1,g1_leadsto_a1]) + rtac (temp_use LatticeReflexivity) 1, + rtac (temp_use LatticeTransitivity) 1, + auto_tac (Inc_css addSIs2 [a1_leadsto_b1,g1_leadsto_a1] addsimps2 [split_box_conj]) ]); qed_goal "N1_live" Inc.thy - "[](PsiInv .& [(N1 .| N2) .& .~ beta1]_) \ -\ .& SF(N1)_ .& [] SF(N2)_ \ -\ .-> <>($pc1 .= #b)" - (fn _ => [auto_tac (Inc_css addSIs2 [(temp_mp N1_leadsto_b) RSN(2,leadsto_init)]), - rewrite_goals_tac (Init_def::action_rews), - exhaust_tac "pc1 (fst_st sigma)" 1, + "|- []($PsiInv & [(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2)) \ +\ & SF(N1)_(x,y,sem,pc1,pc2) & [] SF(N2)_(x,y,sem,pc1,pc2) \ +\ --> <>(pc1 = #b)" + (fn _ => [auto_tac (Inc_css addsimps2 Init_defs + addSIs2 [(temp_use N1_leadsto_b) + RSN(2, temp_use leadsto_init)]), + exhaust_tac "pc1 (st1 sigma)" 1, Auto_tac ]); qed_goal "N1_enabled_at_b" Inc.thy - "($pc1 .= #b) .-> $(Enabled (_))" - (fn _ => [Action_simp_tac 1, + "|- pc1 = #b --> Enabled (_(x,y,sem,pc1,pc2))" + (fn _ => [Clarsimp_tac 1, res_inst_tac [("F","beta1")] enabled_mono 1, enabled_tac Inc_base 1, auto_tac (Inc_css addsimps2 [angle_def,beta1_def,N1_def]) @@ -229,23 +223,21 @@ (* Now assemble the bits and pieces to prove that Psi is fair. *) -goal Inc.thy "[](PsiInv .& [(N1 .| N2)]_) \ -\ .& SF(N1)_ .& [] SF(N2)_ \ -\ .-> SF(M1)_"; -by (res_inst_tac [("B","beta1"),("P","$pc1 .= #b")] SF2 1); +qed_goal "Fair_M1_lemma" Inc.thy + "|- []($PsiInv & [(N1 | N2)]_(x,y,sem,pc1,pc2)) \ +\ & SF(N1)_(x,y,sem,pc1,pc2) & []SF(N2)_(x,y,sem,pc1,pc2) \ +\ --> SF(M1)_(x,y)" + (fn _ => [ res_inst_tac [("B","beta1"),("P","PRED pc1 = #b")] SF2 1, + (* action premises *) + force_tac (Inc_css addsimps2 [angle_def,M1_def,beta1_def]) 1, + force_tac (Inc_css addsimps2 angle_def::Psi_defs) 1, + force_tac (Inc_css addSEs2 [N1_enabled_at_b]) 1, + (* temporal premise: use previous lemmas and simple TL *) + force_tac (Inc_css addSIs2 [DmdStable, N1_live,Stuck_at_b] + addEs2 [STL4E] addsimps2 [square_def]) 1 + ]); -(* the action premises are simple *) - by (force_tac (Inc_css addsimps2 [angle_def,M1_def,beta1_def]) 1); - by (force_tac (Inc_css addsimps2 angle_def::Psi_defs) 1); - by (force_tac (Inc_css addSEs2 [action_mp N1_enabled_at_b]) 1); -(* temporal premise: use previous lemmas and simple TL *) -by (force_tac (Inc_css addSIs2 DmdStable::(map temp_mp [N1_live,Stuck_at_b]) - addEs2 [STL4E] addsimps2 [square_def]) 1); -qed "Fair_M1_lemma"; - -qed_goal "Fair_M1" Inc.thy "Psi .-> WF(M1)_" - (fn _ => [auto_tac (Inc_css addSIs2 SFImplWF::(map temp_mp [Fair_M1_lemma, PsiInv]) - addsimps2 [split_box_conj]), - auto_tac (Inc_css addsimps2 Psi_def::more_temp_simps) +qed_goal "Fair_M1" Inc.thy "|- Psi --> WF(M1)_(x,y)" + (fn _ => [auto_tac (Inc_css addSIs2 [SFImplWF, Fair_M1_lemma, PsiInv] + addsimps2 [Psi_def,split_box_conj]@more_temp_simps) ]); - diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Inc/Inc.thy --- a/src/HOL/TLA/Inc/Inc.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Inc/Inc.thy Mon Feb 08 13:02:56 1999 +0100 @@ -9,59 +9,62 @@ Lamport's "increment" example. *) -Inc = TLA + Nat + Pcount + +Inc = TLA + Nat + + +(* program counter as an enumeration type *) +datatype pcount = a | b | g consts (* program variables *) - x,y,sem :: "nat stfun" - pc1,pc2 :: "pcount stfun" + x,y,sem :: nat stfun + pc1,pc2 :: pcount stfun (* names of actions and predicates *) - M1,M2,N1,N2 :: "action" - alpha1,alpha2,beta1,beta2,gamma1,gamma2 :: "action" - InitPhi, InitPsi :: "action" - PsiInv,PsiInv1,PsiInv2,PsiInv3 :: "action" + M1,M2,N1,N2 :: action + alpha1,alpha2,beta1,beta2,gamma1,gamma2 :: action + InitPhi, InitPsi :: stpred + PsiInv,PsiInv1,PsiInv2,PsiInv3 :: stpred (* temporal formulas *) - Phi, Psi :: "temporal" + Phi, Psi :: temporal rules (* the "base" variables, required to compute enabledness predicates *) - Inc_base "base_var " + Inc_base "basevars (x, y, sem, pc1, pc2)" (* definitions for high-level program *) - InitPhi_def "InitPhi == ($x .= # 0) .& ($y .= # 0)" - M1_def "M1 == (x$ .= Suc[$x]) .& (y$ .= $y)" - M2_def "M2 == (y$ .= Suc[$y]) .& (x$ .= $x)" - Phi_def "Phi == Init(InitPhi) .& [][M1 .| M2]_ .& \ -\ WF(M1)_ .& WF(M2)_" + InitPhi_def "InitPhi == PRED x = # 0 & y = # 0" + M1_def "M1 == ACT x` = Suc<$x> & y` = $y" + M2_def "M2 == ACT y` = Suc<$y> & x` = $x" + Phi_def "Phi == TEMP Init InitPhi & [][M1 | M2]_(x,y) + & WF(M1)_(x,y) & WF(M2)_(x,y)" (* definitions for low-level program *) - InitPsi_def "InitPsi == ($pc1 .= #a) .& ($pc2 .= #a) .& \ -\ ($x .= # 0) .& ($y .= # 0) .& ($sem .= Suc[# 0])" - alpha1_def "alpha1 == ($pc1 .= #a) .& (pc1$ .= #b) .& ($sem .= Suc[sem$]) .& \ -\ unchanged()" - alpha2_def "alpha2 == ($pc2 .= #a) .& (pc2$ .= #b) .& ($sem .= Suc[sem$]) .& \ -\ unchanged()" - beta1_def "beta1 == ($pc1 .= #b) .& (pc1$ .= #g) .& (x$ .= Suc[$x]) .& \ -\ unchanged()" - beta2_def "beta2 == ($pc2 .= #b) .& (pc2$ .= #g) .& (y$ .= Suc[$y]) .& \ -\ unchanged()" - gamma1_def "gamma1 == ($pc1 .= #g) .& (pc1$ .= #a) .& (sem$ .= Suc[$sem]) .& \ -\ unchanged()" - gamma2_def "gamma2 == ($pc2 .= #g) .& (pc2$ .= #a) .& (sem$ .= Suc[$sem]) .& \ -\ unchanged()" - N1_def "N1 == alpha1 .| beta1 .| gamma1" - N2_def "N2 == alpha2 .| beta2 .| gamma2" - Psi_def "Psi == Init(InitPsi) \ -\ .& [][N1 .| N2]_ \ -\ .& SF(N1)_ \ -\ .& SF(N2)_" + InitPsi_def "InitPsi == PRED pc1 = #a & pc2 = #a + & x = # 0 & y = # 0 & sem = # 1" + alpha1_def "alpha1 == ACT $pc1 = #a & pc1$ = #b & $sem = Suc + & unchanged(x,y,pc2)" + alpha2_def "alpha2 == ACT $pc2 = #a & pc2$ = #b & $sem = Suc + & unchanged(x,y,pc1)" + beta1_def "beta1 == ACT $pc1 = #b & pc1$ = #g & x$ = Suc<$x> + & unchanged(y,sem,pc2)" + beta2_def "beta2 == ACT $pc2 = #b & pc2$ = #g & y$ = Suc<$y> + & unchanged(x,sem,pc1)" + gamma1_def "gamma1 == ACT $pc1 = #g & pc1$ = #a & sem$ = Suc<$sem> + & unchanged(x,y,pc2)" + gamma2_def "gamma2 == ACT $pc2 = #g & pc2$ = #a & sem$ = Suc<$sem> + & unchanged(x,y,pc1)" + N1_def "N1 == ACT (alpha1 | beta1 | gamma1)" + N2_def "N2 == ACT (alpha2 | beta2 | gamma2)" + Psi_def "Psi == TEMP Init InitPsi + & [][N1 | N2]_(x,y,sem,pc1,pc2) + & SF(N1)_(x,y,sem,pc1,pc2) + & SF(N2)_(x,y,sem,pc1,pc2)" - PsiInv1_def "PsiInv1 == ($sem .= Suc[# 0]) .& ($pc1 .= #a) .& ($pc2 .= #a)" - PsiInv2_def "PsiInv2 == ($sem .= # 0) .& ($pc1 .= #a) .& ($pc2 .= #b .| $pc2 .= #g)" - PsiInv3_def "PsiInv3 == ($sem .= # 0) .& ($pc2 .= #a) .& ($pc1 .= #b .| $pc1 .= #g)" - PsiInv_def "PsiInv == PsiInv1 .| PsiInv2 .| PsiInv3" + PsiInv1_def "PsiInv1 == PRED sem = # 1 & pc1 = #a & pc2 = #a" + PsiInv2_def "PsiInv2 == PRED sem = # 0 & pc1 = #a & (pc2 = #b | pc2 = #g)" + PsiInv3_def "PsiInv3 == PRED sem = # 0 & pc2 = #a & (pc1 = #b | pc1 = #g)" + PsiInv_def "PsiInv == PRED (PsiInv1 | PsiInv2 | PsiInv3)" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Init.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/TLA/Init.ML Mon Feb 08 13:02:56 1999 +0100 @@ -0,0 +1,43 @@ +local + fun prover s = prove_goal Init.thy s + (K [force_tac (claset(), simpset() addsimps [Init_def]) 1]) +in + val const_simps = map (int_rewrite o prover) + [ "|- (Init #True) = #True", + "|- (Init #False) = #False"] + val Init_simps = map (int_rewrite o prover) + [ "|- (Init ~F) = (~ Init F)", + "|- (Init (P --> Q)) = (Init P --> Init Q)", + "|- (Init (P & Q)) = (Init P & Init Q)", + "|- (Init (P | Q)) = (Init P | Init Q)", + "|- (Init (P = Q)) = ((Init P) = (Init Q))", + "|- (Init (!x. F x)) = (!x. (Init F x))", + "|- (Init (? x. F x)) = (? x. (Init F x))", + "|- (Init (?! x. F x)) = (?! x. (Init F x))" + ] +end; + +Addsimps const_simps; + +Goal "|- (Init $P) = (Init P)"; +by (force_tac (claset(), simpset() addsimps [Init_def,fw_act_def,fw_stp_def]) 1); +qed "Init_stp_act"; +val Init_simps = (int_rewrite Init_stp_act)::Init_simps; +bind_thm("Init_stp_act_rev", symmetric(int_rewrite Init_stp_act)); + +Goal "|- (Init F) = F"; +by (force_tac (claset(), simpset() addsimps [Init_def,fw_temp_def]) 1); +qed "Init_temp"; +val Init_simps = (int_rewrite Init_temp)::Init_simps; + +(* Trivial instances of the definitions that avoid introducing lambda expressions. *) +Goalw [Init_def,fw_stp_def] "(sigma |= Init P) = P (st1 sigma)"; +by (rtac refl 1); +qed "Init_stp"; + +Goalw [Init_def,fw_act_def] "(sigma |= Init A) = A (st1 sigma, st2 sigma)"; +by (rtac refl 1); +qed "Init_act"; + +val Init_defs = [Init_stp, Init_act, int_use Init_temp]; + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Init.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/TLA/Init.thy Mon Feb 08 13:02:56 1999 +0100 @@ -0,0 +1,46 @@ +(* + File: TLA/Init.thy + Author: Stephan Merz + Copyright: 1998 University of Munich + + Theory Name: Init + Logic Image: HOL + +Introduces type of temporal formulas. Defines interface between +temporal formulas and its "subformulas" (state predicates and actions). +*) + +Init = Action + + +types + behavior + temporal = behavior form + +arities + behavior :: term + +instance + behavior :: world + +consts + Initial :: ('w::world => bool) => temporal + first_world :: behavior => ('w::world) + st1, st2 :: behavior => state + +syntax + TEMP :: lift => 'a ("(TEMP _)") + "_Init" :: lift => lift ("(Init _)"[40] 50) + +translations + "TEMP F" => "(F::behavior => _)" + "_Init" == "Initial" + "sigma |= Init F" <= "_Init F sigma" + +defs + Init_def "sigma |= Init F == (first_world sigma) |= F" + fw_temp_def "first_world == %sigma. sigma" + fw_stp_def "first_world == st1" + fw_act_def "first_world == %sigma. (st1 sigma, st2 sigma)" +end + +ML diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/IntLemmas.ML --- a/src/HOL/TLA/IntLemmas.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/IntLemmas.ML Mon Feb 08 13:02:56 1999 +0100 @@ -1,7 +1,7 @@ (* File: IntLemmas.ML Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Lemmas and tactics for "intensional" logics. @@ -12,32 +12,29 @@ qed_goal "substW" Intensional.thy - "[| x .= y; w |= (P::[('v::world) => 'a, 'w::world] => bool)(x) |] ==> w |= P(y)" + "[| |- x = y; w |= P(x) |] ==> w |= P(y)" (fn [prem1,prem2] => [rtac (rewrite_rule ([prem1] RL [inteq_reflection]) prem2) 1]); (* Lift HOL rules to intensional reasoning *) -qed_goal "reflW" Intensional.thy "x .= x" - (fn _ => [ rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 ]); +qed_goal "reflW" Intensional.thy "|- x = x" + (fn _ => [Simp_tac 1]); - -qed_goal "symW" Intensional.thy "s .= t ==> t .= s" +qed_goal "symW" Intensional.thy "|- s = t ==> |- t = s" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, dtac intD 1, rewrite_goals_tac intensional_rews, etac sym 1 ]); -qed_goal "not_symW" Intensional.thy "s .~= t ==> t .~= s" +qed_goal "not_symW" Intensional.thy "|- s ~= t ==> |- t ~= s" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, dtac intD 1, rewrite_goals_tac intensional_rews, etac not_sym 1 ]); qed_goal "transW" Intensional.thy - "[| r .= s; s .= t |] ==> r .= t" + "[| |- r = s; |- s = t |] ==> |- r = t" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, REPEAT (dtac intD 1), rewrite_goals_tac intensional_rews, @@ -45,36 +42,35 @@ atac 1 ]); qed_goal "box_equalsW" Intensional.thy - "[| a .= b; a .= c; b .= d |] ==> c .= d" + "[| |- a = b; |- a = c; |- b = d |] ==> |- c = d" (fn prems => [ (rtac transW 1), (rtac transW 1), (rtac symW 1), (REPEAT (resolve_tac prems 1)) ]); +(* NB: Antecedent is a standard HOL (non-intensional) formula. *) qed_goal "fun_congW" Intensional.thy - "(f::('a => 'b)) = g ==> f[x] .= g[x]" + "f = g ==> |- f = g" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, rewrite_goals_tac intensional_rews, etac fun_cong 1 ]); qed_goal "fun_cong2W" Intensional.thy - "(f::(['a,'b] => 'c)) = g ==> f[x,y] .= g[x,y]" + "f = g ==> |- f = g" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, - rewrite_goals_tac intensional_rews, - asm_full_simp_tac HOL_ss 1 ]); + Asm_full_simp_tac 1 ]); qed_goal "fun_cong3W" Intensional.thy - "(f::(['a,'b,'c] => 'd)) = g ==> f[x,y,z] .= g[x,y,z]" + "f = g ==> |- f = g" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, - rewrite_goals_tac intensional_rews, - asm_full_simp_tac HOL_ss 1 ]); + Asm_full_simp_tac 1 ]); -qed_goal "arg_congW" Intensional.thy "x .= y ==> (f::'a=>'b)[x] .= f[y]" +qed_goal "arg_congW" Intensional.thy "|- x = y ==> |- f = f" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, dtac intD 1, @@ -82,7 +78,7 @@ etac arg_cong 1 ]); qed_goal "arg_cong2W" Intensional.thy - "[| u .= v; x .= y |] ==> (f::['a,'b]=>'c)[u,x] .= f[v,y]" + "[| |- u = v; |- x = y |] ==> |- f = f" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, REPEAT (dtac intD 1), @@ -91,7 +87,7 @@ rtac refl 1 ]); qed_goal "arg_cong3W" Intensional.thy - "[| r .= s; u .= v; x .= y |] ==> (f::['a,'b,'c]=>'d)[r,u,x] .= f[s,v,y]" + "[| |- r = s; |- u = v; |- x = y |] ==> |- f = f" (fn prems => [ cut_facts_tac prems 1, rtac intI 1, REPEAT (dtac intD 1), @@ -100,7 +96,7 @@ rtac refl 1 ]); qed_goal "congW" Intensional.thy - "[| (f::'a=>'b) = g; x .= y |] ==> f[x] .= g[y]" + "[| f = g; |- x = y |] ==> |- f = g" (fn prems => [ rtac box_equalsW 1, rtac reflW 3, rtac arg_congW 1, @@ -110,7 +106,7 @@ resolve_tac prems 1 ]); qed_goal "cong2W" Intensional.thy - "[| (f::['a,'b]=>'c) = g; u .= v; x .= y |] ==> f[u,x] .= g[v,y]" + "[| f = g; |- u = v; |- x = y |] ==> |- f = g" (fn prems => [ rtac box_equalsW 1, rtac reflW 3, rtac arg_cong2W 1, @@ -120,7 +116,7 @@ resolve_tac prems 1 ]); qed_goal "cong3W" Intensional.thy - "[| (f::['a,'b,'c]=>'d) = g; r .= s; u .= v; x .= y |] ==> (f[r,u,x]) .= (g[s,v,y])" + "[| f = g; |- r = s; |- u = v; |- x = y |] ==> |- f = g" (fn prems => [ rtac box_equalsW 1, rtac reflW 3, rtac arg_cong3W 1, @@ -133,48 +129,38 @@ (** Lifted equivalence **) (* Note the object-level implication in the hypothesis. Meta-level implication - would not be correct! *) + would be incorrect! *) qed_goal "iffIW" Intensional.thy - "[| A .-> B; B .-> A |] ==> A .= B" + "[| |- A --> B; |- B --> A |] ==> |- A = B" (fn prems => [ cut_facts_tac prems 1, - rtac intI 1, - REPEAT (dtac intD 1), - rewrite_goals_tac intensional_rews, - (fast_tac prop_cs 1) ]); + rewrite_goals_tac (Valid_def::intensional_rews), + Blast_tac 1 ]); qed_goal "iffD2W" Intensional.thy - "[| (P::('w::world) form) .= Q; w |= Q |] ==> w |= P" - (fn prems => - [cut_facts_tac prems 1, - dtac intD 1, - rewrite_goals_tac intensional_rews, - fast_tac prop_cs 1 ]); + "[| |- P = Q; w |= Q |] ==> w |= P" + (fn prems => [ cut_facts_tac prems 1, + rewrite_goals_tac (Valid_def::intensional_rews), + Blast_tac 1 ]); val iffD1W = symW RS iffD2W; (** #True **) -qed_goal "TrueIW" Intensional.thy "#True" - (fn _ => [rtac intI 1, rewrite_goals_tac intensional_rews, rtac TrueI 1]); - - -qed_goal "eqTrueIW" Intensional.thy "(P::('w::world) form) ==> P .= #True" +qed_goal "eqTrueIW" Intensional.thy "|- P ==> |- P = #True" (fn prems => [cut_facts_tac prems 1, rtac intI 1, dtac intD 1, - rewrite_goals_tac intensional_rews, - asm_full_simp_tac HOL_ss 1] ); + Asm_full_simp_tac 1]); -qed_goal "eqTrueEW" Intensional.thy "P .= #True ==> (P::('w::world) form)" +qed_goal "eqTrueEW" Intensional.thy "|- P = #True ==> |- P" (fn prems => [cut_facts_tac prems 1, rtac intI 1, dtac intD 1, - rewrite_goals_tac intensional_rews, - asm_full_simp_tac HOL_ss 1] ); + Asm_full_simp_tac 1]); (** #False **) -qed_goal "FalseEW" Intensional.thy "#False ==> P::('w::world) form" +qed_goal "FalseEW" Intensional.thy "|- #False ==> |- P" (fn prems => [cut_facts_tac prems 1, rtac intI 1, dtac intD 1, @@ -182,23 +168,20 @@ etac FalseE 1]); qed_goal "False_neq_TrueW" Intensional.thy - "(#False::('w::world) form) .= #True ==> P::('w::world) form" + "|- #False = #True ==> |- P" (fn [prem] => [rtac (prem RS eqTrueEW RS FalseEW) 1]); (** Negation **) (* Again use object-level implication *) -qed_goal "notIW" Intensional.thy "(P .-> #False) ==> .~P" +qed_goal "notIW" Intensional.thy "|- P --> #False ==> |- ~P" (fn prems => [cut_facts_tac prems 1, - rtac intI 1, - dtac intD 1, - rewrite_goals_tac intensional_rews, - fast_tac prop_cs 1]); - + rewrite_goals_tac (Valid_def::intensional_rews), + Blast_tac 1]); qed_goal "notEWV" Intensional.thy - "[| .~P; P::('w::world) form |] ==> R::('w::world) form" + "[| |- ~P; |- P |] ==> |- R" (fn prems => [cut_facts_tac prems 1, rtac intI 1, REPEAT (dtac intD 1), @@ -210,7 +193,7 @@ are allowed to be (intensional) formulas of different types! *) qed_goal "notEW" Intensional.thy - "[| w |= .~P; w |= P |] ==> R::('w::world) form" + "[| w |= ~P; w |= P |] ==> |- R" (fn prems => [cut_facts_tac prems 1, rtac intI 1, rewrite_goals_tac intensional_rews, @@ -218,14 +201,14 @@ (** Implication **) -qed_goal "impIW" Intensional.thy "(!!w. (w |= A) ==> (w |= B)) ==> A .-> B" +qed_goal "impIW" Intensional.thy "(!!w. (w |= A) ==> (w |= B)) ==> |- A --> B" (fn [prem] => [ rtac intI 1, rewrite_goals_tac intensional_rews, rtac impI 1, etac prem 1 ]); -qed_goal "mpW" Intensional.thy "[| A .-> B; w |= A |] ==> w |= B" +qed_goal "mpW" Intensional.thy "[| |- A --> B; w |= A |] ==> w |= B" (fn prems => [ cut_facts_tac prems 1, dtac intD 1, rewrite_goals_tac intensional_rews, @@ -233,124 +216,111 @@ atac 1 ]); qed_goal "impEW" Intensional.thy - "[| A .-> B; w |= A; w |= B ==> w |= C |] ==> w |= (C::('w::world) form)" + "[| |- A --> B; w |= A; w |= B ==> w |= C |] ==> w |= C" (fn prems => [ (REPEAT (resolve_tac (prems@[mpW]) 1)) ]); -qed_goal "rev_mpW" Intensional.thy "[| w |= P; P .-> Q |] ==> w |= Q" +qed_goal "rev_mpW" Intensional.thy "[| w |= P; |- P --> Q |] ==> w |= Q" (fn prems => [ (REPEAT (resolve_tac (prems@[mpW]) 1)) ]); -qed_goal "contraposW" Intensional.thy "[| w |= .~Q; P .-> Q |] ==> w |= .~P" - (fn [major,minor] => [rewrite_goals_tac intensional_rews, - rtac contrapos 1, - rtac (rewrite_rule intensional_rews major) 1, +qed_goalw "contraposW" Intensional.thy intensional_rews + "[| w |= ~Q; |- P --> Q |] ==> w |= ~P" + (fn [major,minor] => [rtac (major RS contrapos) 1, etac rev_mpW 1, rtac minor 1]); qed_goal "iffEW" Intensional.thy - "[| (P::('w::world) form) .= Q; [| P .-> Q; Q .-> P |] ==> R::('w::world) form |] ==> R" + "[| |- P = Q; [| |- P --> Q; |- Q --> P |] ==> R |] ==> R" (fn [p1,p2] => [REPEAT(ares_tac([p1 RS iffD2W, p1 RS iffD1W, p2, impIW])1)]); (** Conjunction **) -qed_goal "conjIW" Intensional.thy "[| w |= P; w |= Q |] ==> w |= P .& Q" - (fn prems => [rewrite_goals_tac intensional_rews, - REPEAT (resolve_tac ([conjI]@prems) 1)]); +qed_goalw "conjIW" Intensional.thy intensional_rews "[| w |= P; w |= Q |] ==> w |= P & Q" + (fn prems => [REPEAT (resolve_tac ([conjI]@prems) 1)]); -qed_goal "conjunct1W" Intensional.thy "(w |= P .& Q) ==> w |= P" +qed_goal "conjunct1W" Intensional.thy "(w |= P & Q) ==> w |= P" (fn prems => [cut_facts_tac prems 1, rewrite_goals_tac intensional_rews, etac conjunct1 1]); -qed_goal "conjunct2W" Intensional.thy "(w |= P .& Q) ==> w |= Q" +qed_goal "conjunct2W" Intensional.thy "(w |= P & Q) ==> w |= Q" (fn prems => [cut_facts_tac prems 1, rewrite_goals_tac intensional_rews, etac conjunct2 1]); qed_goal "conjEW" Intensional.thy - "[| w |= P .& Q; [| w |= P; w |= Q |] ==> w |= R |] ==> w |= (R::('w::world) form)" + "[| w |= P & Q; [| w |= P; w |= Q |] ==> w |= R |] ==> w |= R" (fn prems => [cut_facts_tac prems 1, resolve_tac prems 1, etac conjunct1W 1, etac conjunct2W 1]); (** Disjunction **) -qed_goal "disjI1W" Intensional.thy "w |= P ==> w |= P .| Q" - (fn [prem] => [rewrite_goals_tac intensional_rews, - rtac disjI1 1, - rtac prem 1]); +qed_goalw "disjI1W" Intensional.thy intensional_rews "w |= P ==> w |= P | Q" + (fn [prem] => [REPEAT (resolve_tac [disjI1,prem] 1)]); -qed_goal "disjI2W" Intensional.thy "w |= Q ==> w |= P .| Q" - (fn [prem] => [rewrite_goals_tac intensional_rews, - rtac disjI2 1, - rtac prem 1]); +qed_goalw "disjI2W" Intensional.thy intensional_rews "w |= Q ==> w |= P | Q" + (fn [prem] => [REPEAT (resolve_tac [disjI2,prem] 1)]); qed_goal "disjEW" Intensional.thy - "[| w |= P .| Q; P .-> R; Q .-> R |] ==> w |= R" + "[| w |= P | Q; |- P --> R; |- Q --> R |] ==> w |= R" (fn prems => [cut_facts_tac prems 1, REPEAT (dtac intD 1), rewrite_goals_tac intensional_rews, - fast_tac prop_cs 1]); + Blast_tac 1]); (** Classical propositional logic **) -qed_goal "classicalW" Intensional.thy "(.~P .-> P) ==> P::('w::world)form" - (fn prems => [cut_facts_tac prems 1, - rtac intI 1, - dtac intD 1, - rewrite_goals_tac intensional_rews, - fast_tac prop_cs 1]); +qed_goalw "classicalW" Intensional.thy (Valid_def::intensional_rews) + "!!P. |- ~P --> P ==> |- P" + (fn prems => [Blast_tac 1]); -qed_goal "notnotDW" Intensional.thy ".~.~P ==> P::('w::world) form" - (fn prems => [cut_facts_tac prems 1, - rtac intI 1, +qed_goal "notnotDW" Intensional.thy "!!P. |- ~~P ==> |- P" + (fn prems => [rtac intI 1, dtac intD 1, rewrite_goals_tac intensional_rews, etac notnotD 1]); -qed_goal "disjCIW" Intensional.thy "(w |= .~Q .-> P) ==> (w |= P.|Q)" - (fn prems => [cut_facts_tac prems 1, - rewrite_goals_tac intensional_rews, - fast_tac prop_cs 1]); +qed_goal "disjCIW" Intensional.thy "!!P Q. (w |= ~Q --> P) ==> (w |= P|Q)" + (fn prems => [rewrite_goals_tac intensional_rews, + Blast_tac 1]); qed_goal "impCEW" Intensional.thy - "[| P.->Q; (w |= .~P) ==> (w |= R); (w |= Q) ==> (w |= R) |] ==> w |= (R::('w::world) form)" + "[| |- P --> Q; (w |= ~P) ==> (w |= R); (w |= Q) ==> (w |= R) |] ==> w |= R" (fn [a1,a2,a3] => [rtac (excluded_middle RS disjE) 1, etac (rewrite_rule intensional_rews a2) 1, rtac a3 1, etac (a1 RS mpW) 1]); -(* The following generates too many parse trees... - -qed_goal "iffCEW" Intensional.thy - "[| P .= Q; \ +qed_goalw "iffCEW" Intensional.thy intensional_rews + "[| |- P = Q; \ \ [| (w |= P); (w |= Q) |] ==> (w |= R); \ -\ [| (w |= .~P); (w |= .~Q) |] ==> (w |= R) \ -\ |] ==> w |= (R::('w::world) form)" - -*) +\ [| (w |= ~P); (w |= ~Q) |] ==> (w |= R) \ +\ |] ==> w |= R" + (fn [a1,a2,a3] => + [rtac iffCE 1, + etac a2 2, atac 2, + etac a3 2, atac 2, + rtac (int_unlift a1) 1]); qed_goal "case_split_thmW" Intensional.thy - "[| P .-> Q; .~P .-> Q |] ==> Q::('w::world) form" - (fn prems => [cut_facts_tac prems 1, - rtac intI 1, - REPEAT (dtac intD 1), - rewrite_goals_tac intensional_rews, - fast_tac prop_cs 1]); + "!!P. [| |- P --> Q; |- ~P --> Q |] ==> |- Q" + (fn _ => [rewrite_goals_tac (Valid_def::intensional_rews), + Blast_tac 1]); fun case_tacW a = res_inst_tac [("P",a)] case_split_thmW; (** Rigid quantifiers **) -qed_goal "allIW" Intensional.thy "(!!x. P(x)) ==> RALL x. P(x)" +qed_goal "allIW" Intensional.thy "(!!x. |- P x) ==> |- ! x. P(x)" (fn [prem] => [rtac intI 1, rewrite_goals_tac intensional_rews, rtac allI 1, - rtac (prem RS intE) 1]); + rtac (prem RS intD) 1]); -qed_goal "specW" Intensional.thy "(RALL x. P(x)) ==> P(x)" +qed_goal "specW" Intensional.thy "|- ! x. P x ==> |- P x" (fn prems => [cut_facts_tac prems 1, rtac intI 1, dtac intD 1, @@ -359,24 +329,24 @@ qed_goal "allEW" Intensional.thy - "[| RALL x. P(x); P(x) ==> R |] ==> R::('w::world) form" + "[| |- ! x. P x; |- P x ==> |- R |] ==> |- R" (fn major::prems=> [ (REPEAT (resolve_tac (prems @ [major RS specW]) 1)) ]); qed_goal "all_dupEW" Intensional.thy - "[| RALL x. P(x); [| P(x); RALL x. P(x) |] ==> R |] ==> R::('w::world) form" + "[| |- ! x. P x; [| |- P x; |- ! x. P x |] ==> |- R |] ==> |- R" (fn prems => [ (REPEAT (resolve_tac (prems @ (prems RL [specW])) 1)) ]); -qed_goal "exIW" Intensional.thy "P(x) ==> REX x. P(x)" +qed_goal "exIW" Intensional.thy "|- P x ==> |- ? x. P x" (fn [prem] => [rtac intI 1, rewrite_goals_tac intensional_rews, rtac exI 1, rtac (prem RS intD) 1]); qed_goal "exEW" Intensional.thy - "[| w |= REX x. P(x); !!x. P(x) .-> Q |] ==> w |= Q" + "[| w |= ? x. P x; !!x. |- P x --> Q |] ==> w |= Q" (fn [major,minor] => [rtac exE 1, rtac (rewrite_rule intensional_rews major) 1, etac rev_mpW 1, @@ -385,8 +355,7 @@ (** Classical quantifier reasoning **) qed_goal "exCIW" Intensional.thy - "(w |= (RALL x. .~P(x)) .-> P(a)) ==> w |= REX x. P(x)" - (fn prems => [cut_facts_tac prems 1, - rewrite_goals_tac intensional_rews, - fast_tac HOL_cs 1]); + "!!P. w |= (! x. ~P x) --> P a ==> w |= ? x. P x" + (fn prems => [rewrite_goals_tac intensional_rews, + Blast_tac 1]); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Intensional.ML --- a/src/HOL/TLA/Intensional.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Intensional.ML Mon Feb 08 13:02:56 1999 +0100 @@ -1,204 +1,136 @@ (* File: Intensional.ML Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Lemmas and tactics for "intensional" logics. *) -val intensional_rews = [unl_con,unl_lift,unl_lift2,unl_lift3,unl_Rall,unl_Rex]; +val intensional_rews = [unl_con,unl_lift,unl_lift2,unl_lift3,unl_Rall,unl_Rex,unl_Rex1]; + +qed_goalw "inteq_reflection" Intensional.thy [Valid_def,unl_lift2] + "|- x=y ==> (x==y)" + (fn [prem] => [rtac eq_reflection 1, rtac ext 1, rtac (prem RS spec) 1 ]); -(** Lift usual HOL simplifications to "intensional" level. - Convert s .= t into rewrites s == t, so we can use the standard - simplifier. -**) +qed_goalw "intI" Intensional.thy [Valid_def] "(!!w. w |= A) ==> |- A" + (fn [prem] => [REPEAT (resolve_tac [allI,prem] 1)]); + +qed_goalw "intD" Intensional.thy [Valid_def] "|- A ==> w |= A" + (fn [prem] => [rtac (prem RS spec) 1]); + + +(** Lift usual HOL simplifications to "intensional" level. **) local fun prover s = (prove_goal Intensional.thy s - (fn _ => [rewrite_goals_tac (int_valid::intensional_rews), - blast_tac HOL_cs 1])) RS inteq_reflection; + (fn _ => [rewrite_goals_tac (Valid_def::intensional_rews), + blast_tac HOL_cs 1])) RS inteq_reflection in val int_simps = map prover - [ "(x.=x) .= #True", - "(.~#True) .= #False", "(.~#False) .= #True", "(.~ .~ P) .= P", - "((.~P) .= P) .= #False", "(P .= (.~P)) .= #False", - "(P .~= Q) .= (P .= (.~Q))", - "(#True.=P) .= P", "(P.=#True) .= P", - "(#True .-> P) .= P", "(#False .-> P) .= #True", - "(P .-> #True) .= #True", "(P .-> P) .= #True", - "(P .-> #False) .= (.~P)", "(P .-> .~P) .= (.~P)", - "(P .& #True) .= P", "(#True .& P) .= P", - "(P .& #False) .= #False", "(#False .& P) .= #False", - "(P .& P) .= P", "(P .& .~P) .= #False", "(.~P .& P) .= #False", - "(P .| #True) .= #True", "(#True .| P) .= #True", - "(P .| #False) .= P", "(#False .| P) .= P", - "(P .| P) .= P", "(P .| .~P) .= #True", "(.~P .| P) .= #True", - "(RALL x. P) .= P", "(REX x. P) .= P", - "(.~Q .-> .~P) .= (P .-> Q)", - "(P.|Q .-> R) .= ((P.->R).&(Q.->R))" ]; - + [ "|- (x=x) = #True", + "|- (~#True) = #False", "|- (~#False) = #True", "|- (~~ P) = P", + "|- ((~P) = P) = #False", "|- (P = (~P)) = #False", + "|- (P ~= Q) = (P = (~Q))", + "|- (#True=P) = P", "|- (P=#True) = P", + "|- (#True --> P) = P", "|- (#False --> P) = #True", + "|- (P --> #True) = #True", "|- (P --> P) = #True", + "|- (P --> #False) = (~P)", "|- (P --> ~P) = (~P)", + "|- (P & #True) = P", "|- (#True & P) = P", + "|- (P & #False) = #False", "|- (#False & P) = #False", + "|- (P & P) = P", "|- (P & ~P) = #False", "|- (~P & P) = #False", + "|- (P | #True) = #True", "|- (#True | P) = #True", + "|- (P | #False) = P", "|- (#False | P) = P", + "|- (P | P) = P", "|- (P | ~P) = #True", "|- (~P | P) = #True", + "|- (! x. P) = P", "|- (? x. P) = P", + "|- (~Q --> ~P) = (P --> Q)", + "|- (P|Q --> R) = ((P-->R)&(Q-->R))" ]; end; -Addsimps (intensional_rews @ int_simps); +qed_goal "TrueW" Intensional.thy "|- #True" + (fn _ => [simp_tac (simpset() addsimps [Valid_def,unl_con]) 1]); -(* Derive introduction and destruction rules from definition of - intensional validity. -*) -qed_goal "intI" Intensional.thy "(!!w. w |= A) ==> A" - (fn prems => [rewtac int_valid, - resolve_tac prems 1 - ]); +Addsimps (TrueW::intensional_rews); +Addsimps int_simps; +AddSIs [intI]; +AddDs [intD]; -qed_goalw "intD" Intensional.thy [int_valid] "A ==> w |= A" - (fn [prem] => [ rtac (forall_elim_var 0 prem) 1 ]); (* ======== Functions to "unlift" intensional implications into HOL rules ====== *) (* Basic unlifting introduces a parameter "w" and applies basic rewrites, e.g. - F .= G gets (w |= F) = (w |= G) - F .-> G gets (w |= F) --> (w |= G) -*) -fun int_unlift th = rewrite_rule intensional_rews (th RS intD); - -(* F .-> G becomes w |= F ==> w |= G *) -fun int_mp th = zero_var_indexes ((int_unlift th) RS mp); - -(* F .-> G becomes [| w |= F; w |= G ==> R |] ==> R - so that it can be used as an elimination rule -*) -fun int_impE th = zero_var_indexes ((int_unlift th) RS impE); - -(* F .& G .-> H becomes [| w |= F; w |= G |] ==> w |= H *) -fun int_conjmp th = zero_var_indexes (conjI RS (int_mp th)); - -(* F .& G .-> H becomes [| w |= F; w |= G; (w |= H ==> R) |] ==> R *) -fun int_conjimpE th = zero_var_indexes (conjI RS (int_impE th)); - -(* Turn F .= G into meta-level rewrite rule F == G *) -fun int_rewrite th = (rewrite_rule intensional_rews (th RS inteq_reflection)); - -(* Make the simplifier accept "intensional" goals by first unlifting them. - This is the standard way of proving "intensional" theorems; apply - int_rewrite (or action_rewrite, temp_rewrite) to convert "x .= y" into "x == y" - if you want to rewrite without unlifting. -*) -fun maybe_unlift th = - (case concl_of th of - Const("Intensional.TrueInt",_) $ p => int_unlift th - | _ => th); - -simpset_ref() := simpset() setmksimps ((mksimps mksimps_pairs) o maybe_unlift); - - -(* ==================== Rewrites for abstractions ==================== *) - -(* The following are occasionally useful. Don't add them to the default - simpset, or it will loop! Alternatively, we could replace the "unl_XXX" - rules by definitions of lifting via lambda abstraction, but then proof - states would have lots of lambdas, and would be hard to read. + |- F = G becomes F w = G w + |- F --> G becomes F w --> G w *) -qed_goal "con_abs" Intensional.thy "(%w. c) == #c" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); +fun int_unlift th = + rewrite_rule intensional_rews ((th RS intD) handle _ => th); -qed_goal "lift_abs" Intensional.thy "(%w. f(x w)) == (f[x])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); +(* Turn |- F = G into meta-level rewrite rule F == G *) +fun int_rewrite th = + zero_var_indexes (rewrite_rule intensional_rews (th RS inteq_reflection)); -qed_goal "lift2_abs" Intensional.thy "(%w. f(x w) (y w)) == (f[x,y])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); +(* flattening turns "-->" into "==>" and eliminates conjunctions in the + antecedent. For example, + + P & Q --> (R | S --> T) becomes [| P; Q; R | S |] ==> T -qed_goal "lift2_abs_con1" Intensional.thy "(%w. f x (y w)) == (f[#x,y])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); - -qed_goal "lift2_abs_con2" Intensional.thy "(%w. f(x w) y) == (f[x,#y])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); - -qed_goal "lift3_abs" Intensional.thy "(%w. f(x w) (y w) (z w)) == (f[x,y,z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); + Flattening can be useful with "intensional" lemmas (after unlifting). + Naive resolution with mp and conjI may run away because of higher-order + unification, therefore the code is a little awkward. +*) +fun flatten t = + let + (* analogous to RS, but using matching instead of resolution *) + fun matchres tha i thb = + case Seq.chop (2, biresolution true [(false,tha)] i thb) of + ([th],_) => th + | ([],_) => raise THM("matchres: no match", i, [tha,thb]) + | _ => raise THM("matchres: multiple unifiers", i, [tha,thb]) -qed_goal "lift3_abs_con1" Intensional.thy "(%w. f x (y w) (z w)) == (f[#x,y,z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); + (* match tha with some premise of thb *) + fun matchsome tha thb = + let fun hmatch 0 = raise THM("matchsome: no match", 0, [tha,thb]) + | hmatch n = (matchres tha n thb) handle _ => hmatch (n-1) + in hmatch (nprems_of thb) end -qed_goal "lift3_abs_con2" Intensional.thy "(%w. f (x w) y (z w)) == (f[x,#y,z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); - -qed_goal "lift3_abs_con3" Intensional.thy "(%w. f (x w) (y w) z) == (f[x,y,#z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); + fun hflatten t = + case (concl_of t) of + Const _ $ (Const ("op -->", _) $ _ $ _) => hflatten (t RS mp) + | _ => (hflatten (matchsome conjI t)) handle _ => zero_var_indexes t + in + hflatten t +end; -qed_goal "lift3_abs_con12" Intensional.thy "(%w. f x y (z w)) == (f[#x,#y,z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); +fun int_use th = + case (concl_of th) of + Const _ $ (Const ("Intensional.Valid", _) $ _) => + ((flatten (int_unlift th)) handle _ => th) + | _ => th; -qed_goal "lift3_abs_con13" Intensional.thy "(%w. f x (y w) z) == (f[#x,y,#z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); +(*** +(* Make the simplifier accept "intensional" goals by either turning them into + a meta-equality or by unlifting them. +*) -qed_goal "lift3_abs_con23" Intensional.thy "(%w. f (x w) y z) == (f[x,#y,#z])" - (fn _ => [rtac inteq_reflection 1, - rtac intI 1, - rewrite_goals_tac intensional_rews, - rtac refl 1 - ]); +let + val ss = simpset_ref() + fun try_rewrite th = (int_rewrite th) handle _ => (int_use th) handle _ => th +in + ss := !ss setmksimps ((mksimps mksimps_pairs) o try_rewrite) +end; +***) (* ========================================================================= *) -qed_goal "Not_rall" Intensional.thy - "(.~ (RALL x. F(x))) .= (REX x. .~ F(x))" - (fn _ => [rtac intI 1, - rewrite_goals_tac intensional_rews, - fast_tac HOL_cs 1 - ]); +qed_goal "Not_Rall" Intensional.thy + "|- (~(! x. F x)) = (? x. ~F x)" + (fn _ => [simp_tac (simpset() addsimps [Valid_def]) 1]); -qed_goal "Not_rex" Intensional.thy - "(.~ (REX x. F(x))) .= (RALL x. .~ F(x))" - (fn _ => [rtac intI 1, - rewrite_goals_tac intensional_rews, - fast_tac HOL_cs 1 - ]); +qed_goal "Not_Rex" Intensional.thy + "|- (~ (? x. F x)) = (! x. ~ F x)" + (fn _ => [simp_tac (simpset() addsimps [Valid_def]) 1]); (* IntLemmas.ML contains a collection of further lemmas about "intensional" logic. These are not loaded by default because they are not required for the diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Intensional.thy --- a/src/HOL/TLA/Intensional.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Intensional.thy Mon Feb 08 13:02:56 1999 +0100 @@ -1,7 +1,7 @@ (* File: TLA/Intensional.thy Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Theory Name: Intensional Logic Image: HOL @@ -10,95 +10,168 @@ on top of HOL, with lifting of constants and functions. *) -Intensional = Prod + +Intensional = Main + -classes - world < logic (* Type class of "possible worlds". Concrete types - will be provided by children theories. *) +axclass + world < term + +(** abstract syntax **) types - ('a,'w) term = "'w => 'a" (* Intention: 'w::world *) - 'w form = "'w => bool" + ('w,'a) expr = 'w => 'a (* intention: 'w::world, 'a::term *) + 'w form = ('w, bool) expr consts - TrueInt :: "('w::world form) => prop" ("(_)" 5) - - (* Holds at *) - holdsAt :: "['w::world, 'w form] => bool" ("(_ |= _)" [100,9] 8) - - (* Lifting base functions to "intensional" level *) - con :: "'a => ('w::world => 'a)" ("(#_)" [100] 99) - lift :: "['a => 'b, 'w::world => 'a] => ('w => 'b)" ("(_[_])") - lift2 :: "['a => ('b => 'c), 'w::world => 'a, 'w => 'b] => ('w => 'c)" ("(_[_,/ _])") - lift3 :: "['a => 'b => 'c => 'd, 'w::world => 'a, 'w => 'b, 'w => 'c] => ('w => 'd)" ("(_[_,/ _,/ _])") + Valid :: ('w::world) form => bool + const :: 'a => ('w::world, 'a) expr + lift :: ['a => 'b, ('w::world, 'a) expr] => ('w,'b) expr + lift2 :: ['a => 'b => 'c, ('w::world,'a) expr, ('w,'b) expr] => ('w,'c) expr + lift3 :: ['a => 'b => 'c => 'd, ('w::world,'a) expr, ('w,'b) expr, ('w,'c) expr] => ('w,'d) expr - (* Lifted infix functions *) - IntEqu :: "['w::world => 'a, 'w => 'a] => 'w form" ("(_ .=/ _)" [50,51] 50) - IntNeq :: "['w::world => 'a, 'w => 'a] => 'w form" ("(_ .~=/ _)" [50,51] 50) - NotInt :: "('w::world) form => 'w form" ("(.~ _)" [40] 40) - AndInt :: "[('w::world) form, 'w form] => 'w form" ("(_ .&/ _)" [36,35] 35) - OrInt :: "[('w::world) form, 'w form] => 'w form" ("(_ .|/ _)" [31,30] 30) - ImpInt :: "[('w::world) form, 'w form] => 'w form" ("(_ .->/ _)" [26,25] 25) - IfInt :: "[('w::world) form, ('a,'w) term, ('a,'w) term] => ('a,'w) term" ("(.if (_)/ .then (_)/ .else (_))" 10) - PlusInt :: "[('w::world) => ('a::plus), 'w => 'a] => ('w => 'a)" ("(_ .+/ _)" [66,65] 65) - MinusInt :: "[('w::world) => ('a::minus), 'w => 'a] => ('w => 'a)" ("(_ .-/ _)" [66,65] 65) - TimesInt :: "[('w::world) => ('a::times), 'w => 'a] => ('w => 'a)" ("(_ .*/ _)" [71,70] 70) + (* "Rigid" quantification (logic level) *) + RAll :: "('a => ('w::world) form) => 'w form" (binder "Rall " 10) + REx :: "('a => ('w::world) form) => 'w form" (binder "Rex " 10) + REx1 :: "('a => ('w::world) form) => 'w form" (binder "Rex! " 10) - LessInt :: "['w::world => 'a::ord, 'w => 'a] => 'w form" ("(_/ .< _)" [50, 51] 50) - LeqInt :: "['w::world => 'a::ord, 'w => 'a] => 'w form" ("(_/ .<= _)" [50, 51] 50) +(** concrete syntax **) - (* lifted set membership *) - memInt :: "[('a,'w::world) term, ('a set,'w) term] => 'w form" ("(_/ .: _)" [50, 51] 50) - - (* "Rigid" quantification *) - RAll :: "('a => 'w::world form) => 'w form" (binder "RALL " 10) - REx :: "('a => 'w::world form) => 'w form" (binder "REX " 10) +nonterminals + lift + liftargs syntax - "@tupleInt" :: "args => ('a * 'b, 'w) term" ("(1{[_]})") + "" :: id => lift ("_") + "" :: longid => lift ("_") + "" :: var => lift ("_") + "_applC" :: [lift, cargs] => lift ("(1_/ _)" [1000, 1000] 999) + "" :: lift => lift ("'(_')") + "_lambda" :: [idts, 'a] => lift ("(3%_./ _)" [0, 3] 3) + "_constrain" :: [lift, type] => lift ("(_::_)" [4, 0] 3) + "" :: lift => liftargs ("_") + "_liftargs" :: [lift, liftargs] => liftargs ("_,/ _") + "_Valid" :: lift => bool ("(|- _)" 5) + "_holdsAt" :: ['a, lift] => bool ("(_ |= _)" [100,10] 10) + + (* Syntax for lifted expressions outside the scope of |- or |= *) + "LIFT" :: lift => 'a ("LIFT _") + + (* generic syntax for lifted constants and functions *) + "_const" :: 'a => lift ("(#_)" [1000] 999) + "_lift" :: ['a, lift] => lift ("(_<_>)" [1000] 999) + "_lift2" :: ['a, lift, lift] => lift ("(_<_,/ _>)" [1000] 999) + "_lift3" :: ['a, lift, lift, lift] => lift ("(_<_,/ _,/ _>)" [1000] 999) + + (* concrete syntax for common infix functions: reuse same symbol *) + "_liftEqu" :: [lift, lift] => lift ("(_ =/ _)" [50,51] 50) + "_liftNeq" :: [lift, lift] => lift ("(_ ~=/ _)" [50,51] 50) + "_liftNot" :: lift => lift ("(~ _)" [40] 40) + "_liftAnd" :: [lift, lift] => lift ("(_ &/ _)" [36,35] 35) + "_liftOr" :: [lift, lift] => lift ("(_ |/ _)" [31,30] 30) + "_liftImp" :: [lift, lift] => lift ("(_ -->/ _)" [26,25] 25) + "_liftIf" :: [lift, lift, lift] => lift ("(if (_)/ then (_)/ else (_))" 10) + "_liftPlus" :: [lift, lift] => lift ("(_ +/ _)" [66,65] 65) + "_liftMinus" :: [lift, lift] => lift ("(_ -/ _)" [66,65] 65) + "_liftTimes" :: [lift, lift] => lift ("(_ */ _)" [71,70] 70) + "_liftDiv" :: [lift, lift] => lift ("(_ div _)" [71,70] 70) + "_liftMod" :: [lift, lift] => lift ("(_ mod _)" [71,70] 70) + "_liftLess" :: [lift, lift] => lift ("(_/ < _)" [50, 51] 50) + "_liftLeq" :: [lift, lift] => lift ("(_/ <= _)" [50, 51] 50) + "_liftMem" :: [lift, lift] => lift ("(_/ : _)" [50, 51] 50) + "_liftNotMem" :: [lift, lift] => lift ("(_/ ~: _)" [50, 51] 50) + "_liftFinset" :: liftargs => lift ("{(_)}") + (** TODO: syntax for lifted collection / comprehension **) + "_liftPair" :: [lift,liftargs] => lift ("(1'(_,/ _'))") + (* infix syntax for list operations *) + "_liftCons" :: [lift, lift] => lift ("(_ #/ _)" [65,66] 65) + "_liftApp" :: [lift, lift] => lift ("(_ @/ _)" [65,66] 65) + "_liftList" :: liftargs => lift ("[(_)]") + + (* Rigid quantification (syntax level) *) + "_RAll" :: [idts, lift] => lift ("(3! _./ _)" [0, 10] 10) + "_REx" :: [idts, lift] => lift ("(3? _./ _)" [0, 10] 10) + "_REx1" :: [idts, lift] => lift ("(3?! _./ _)" [0, 10] 10) + "_ARAll" :: [idts, lift] => lift ("(3ALL _./ _)" [0, 10] 10) + "_AREx" :: [idts, lift] => lift ("(3EX _./ _)" [0, 10] 10) + "_AREx1" :: [idts, lift] => lift ("(3EX! _./ _)" [0, 10] 10) translations + "_const" == "const" + "_lift" == "lift" + "_lift2" == "lift2" + "_lift3" == "lift3" + "_Valid" == "Valid" + "_RAll x A" == "Rall x. A" + "_REx x A" == "Rex x. A" + "_REx1 x A" == "Rex! x. A" + "_ARAll" => "_RAll" + "_AREx" => "_REx" + "_AREx1" => "_REx1" - "{[x,y,z]}" == "{[x, {[y,z]} ]}" - "{[x,y]}" == "Pair [x, y]" - "{[x]}" => "x" + "w |= A" => "A w" + "LIFT A" => "A::_=>_" - "u .= v" == "op =[u,v]" - "u .~= v" == ".~(u .= v)" - ".~ A" == "Not[A]" - "A .& B" == "op &[A,B]" - "A .| B" == "op |[A,B]" - "A .-> B" == "op -->[A,B]" - ".if A .then u .else v" == "If[A,u,v]" - "u .+ v" == "op +[u,v]" - "u .- v" == "op -[u,v]" - "u .* v" == "op *[u,v]" + "_liftEqu" == "_lift2 (op =)" + "_liftNeq u v" == "_liftNot (_liftEqu u v)" + "_liftNot" == "_lift Not" + "_liftAnd" == "_lift2 (op &)" + "_liftOr" == "_lift2 (op | )" + "_liftImp" == "_lift2 (op -->)" + "_liftIf" == "_lift3 If" + "_liftPlus" == "_lift2 (op +)" + "_liftMinus" == "_lift2 (op -)" + "_liftTimes" == "_lift2 (op *)" + "_liftDiv" == "_lift2 (op div)" + "_liftMod" == "_lift2 (op mod)" + "_liftLess" == "_lift2 (op <)" + "_liftLeq" == "_lift2 (op <=)" + "_liftMem" == "_lift2 (op :)" + "_liftNotMem x xs" == "_liftNot (_liftMem x xs)" + "_liftFinset (_liftargs x xs)" == "_lift2 insert x (_liftFinset xs)" + "_liftFinset x" == "_lift2 insert x (_const {})" + "_liftPair x (_liftargs y z)" == "_liftPair x (_liftPair y z)" + "_liftPair" == "_lift2 Pair" + "_liftCons" == "lift2 (op #)" + "_liftApp" == "lift2 (op @)" + "_liftList (_liftargs x xs)" == "_liftCons x (_liftList xs)" + "_liftList x" == "_liftCons x (_const [])" - "a .< b" == "op < [a,b]" - "a .<= b" == "op <= [a,b]" - "a .: A" == "op :[a,A]" + - "holdsAt w (lift f x)" == "lift f x w" - "holdsAt w (lift2 f x y)" == "lift2 f x y w" - "holdsAt w (lift3 f x y z)" == "lift3 f x y z w" - - "w |= A" => "A(w)" + "w |= ~A" <= "_liftNot A w" + "w |= A & B" <= "_liftAnd A B w" + "w |= A | B" <= "_liftOr A B w" + "w |= A --> B" <= "_liftImp A B w" + "w |= u = v" <= "_liftEqu u v w" + "w |= ! x. A" <= "_RAll x A w" + "w |= ? x. A" <= "_REx x A w" + "w |= ?! x. A" <= "_REx1 x A w" syntax (symbols) - holdsAt :: "['w::world, 'w form] => bool" ("(_ \\ _)" [100,9] 8) - + "_Valid" :: lift => bool ("(\\ _)" 5) + "_holdsAt" :: ['a, lift] => bool ("(_ \\ _)" [100,10] 10) + "_liftNeq" :: [lift, lift] => lift (infixl "\\" 50) + "_liftNot" :: lift => lift ("\\ _" [40] 40) + "_liftAnd" :: [lift, lift] => lift (infixr "\\" 35) + "_liftOr" :: [lift, lift] => lift (infixr "\\" 30) + "_liftImp" :: [lift, lift] => lift (infixr "\\\\" 25) + "_RAll" :: [idts, lift] => lift ("(3\\_./ _)" [0, 10] 10) + "_REx" :: [idts, lift] => lift ("(3\\_./ _)" [0, 10] 10) + "_REx1" :: [idts, lift] => lift ("(3\\!_./ _)" [0, 10] 10) + "_liftLeq" :: [lift, lift] => lift ("(_/ \\ _)" [50, 51] 50) + "_liftMem" :: [lift, lift] => lift ("(_/ \\ _)" [50, 51] 50) + "_liftNotMem" :: [lift, lift] => lift ("(_/ \\ _)" [50, 51] 50) rules - inteq_reflection "(x .= y) ==> (x == y)" + Valid_def "|- A == ALL w. w |= A" - int_valid "TrueInt(A) == (!! w. w |= A)" + unl_con "LIFT #c w == c" + unl_lift "LIFT f w == f (x w)" + unl_lift2 "LIFT f w == f (x w) (y w)" + unl_lift3 "LIFT f w == f (x w) (y w) (z w)" - unl_con "(#c) w == c" (* constants *) - unl_lift "(f[x]) w == f(x w)" - unl_lift2 "(f[x,y]) w == f (x w) (y w)" - unl_lift3 "(f[x, y, z]) w == f (x w) (y w) (z w)" + unl_Rall "w |= ! x. A x == ! x. (w |= A x)" + unl_Rex "w |= ? x. A x == ? x. (w |= A x)" + unl_Rex1 "w |= ?! x. A x == ?! x. (w |= A x)" +end - unl_Rall "(RALL x. A(x)) w == ALL x. (w |= A(x))" - unl_Rex "(REX x. A(x)) w == EX x. (w |= A(x))" - -end +ML diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MIlive.ML --- a/src/HOL/TLA/Memory/MIlive.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MIlive.ML Mon Feb 08 13:02:56 1999 +0100 @@ -14,10 +14,10 @@ (* ------------------------------ State S1 ------------------------------ *) qed_goal "S1_successors" MemoryImplementation.thy - "$(S1 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .-> $(S1 rmhist p)` .| $(S2 rmhist p)`" + "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> (S1 rmhist p)` | (S2 rmhist p)`" (fn _ => [split_idle_tac [] 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_1]) + auto_tac (MI_css addSDs2 [Step1_2_1]) ]); (* Show that the implementation can satisfy the high-level fairness requirements @@ -25,61 +25,61 @@ *) qed_goal "S1_RNextdisabled" MemoryImplementation.thy - "$(S1 rmhist p) .-> \ -\ .~$(Enabled (_))" + "|- S1 rmhist p --> \ +\ ~Enabled (_(rtrner memCh!p, resbar rmhist!p))" (fn _ => [action_simp_tac (simpset() addsimps [angle_def,S_def,S1_def]) - [notI] [enabledE,MemoryidleE] 1, - auto_tac MI_fast_css + [notI] [enabledE,temp_elim Memoryidle] 1, + Force_tac 1 ]); qed_goal "S1_Returndisabled" MemoryImplementation.thy - "$(S1 rmhist p) .-> \ -\ .~$(Enabled (_))" + "|- S1 rmhist p --> \ +\ ~Enabled (_(rtrner memCh!p, resbar rmhist!p))" (fn _ => [action_simp_tac (simpset() addsimps [angle_def,MemReturn_def,Return_def,S_def,S1_def]) [notI] [enabledE] 1 ]); qed_goal "RNext_fair" MemoryImplementation.thy - "!!sigma. (sigma |= []<>($(S1 rmhist p))) \ -\ ==> (sigma |= WF(RNext memCh mem (resbar rmhist) p)_)" - (fn _ => [auto_tac (MI_css addsimps2 [temp_rewrite WF_alt] + "|- []<>S1 rmhist p \ +\ --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [auto_tac (MI_css addsimps2 [WF_alt] addSIs2 [S1_RNextdisabled] addSEs2 [STL4E,DmdImplE]) ]); qed_goal "Return_fair" MemoryImplementation.thy - "!!sigma. (sigma |= []<>($(S1 rmhist p))) \ -\ ==> (sigma |= WF(MemReturn memCh (resbar rmhist) p)_)" - (fn _ => [auto_tac (MI_css addsimps2 [temp_rewrite WF_alt] + "|- []<>S1 rmhist p \ +\ --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [auto_tac (MI_css addsimps2 [WF_alt] addSIs2 [S1_Returndisabled] addSEs2 [STL4E,DmdImplE]) ]); (* ------------------------------ State S2 ------------------------------ *) qed_goal "S2_successors" MemoryImplementation.thy - "$(S2 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(S2 rmhist p)` .| $(S3 rmhist p)`" + "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> (S2 rmhist p)` | (S3 rmhist p)`" (fn _ => [split_idle_tac [] 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_2]) + auto_tac (MI_css addSDs2 [Step1_2_2]) ]); qed_goal "S2MClkFwd_successors" MemoryImplementation.thy - "$(S2 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .& _(c p) \ -\ .-> $(S3 rmhist p)`" - (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSEs2 [action_conjimpE Step1_2_2]) ]); + "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ +\ & _(c p) \ +\ --> (S3 rmhist p)`" + (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_2]) ]); qed_goal "S2MClkFwd_enabled" MemoryImplementation.thy - "$(S2 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(Enabled (_(c p)))" - (fn _ => [cut_facts_tac [MI_base] 1, - auto_tac (MI_css addsimps2 [c_def,base_pair] - addSIs2 [MClkFwd_ch_enabled,action_mp MClkFwd_enabled]), - ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S2_def]) [] []) + "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> $Enabled (_(c p))" + (fn _ => [auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkFwd_ch_enabled,MClkFwd_enabled]), + cut_facts_tac [MI_base] 1, + blast_tac (claset() addDs [base_pair]) 1, + ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S2_def])) ]); qed_goal "S2_live" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_) .& WF(MClkFwd memCh crCh cst p)_(c p) \ -\ .-> ($(S2 rmhist p) ~> $(S3 rmhist p))" + "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) & WF(MClkFwd memCh crCh cst p)_(c p) \ +\ --> (S2 rmhist p ~> S3 rmhist p)" (fn _ => [REPEAT (resolve_tac [WF1,S2_successors, S2MClkFwd_successors,S2MClkFwd_enabled] 1) ]); @@ -88,185 +88,165 @@ (* ------------------------------ State S3 ------------------------------ *) qed_goal "S3_successors" MemoryImplementation.thy - "$(S3 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(S3 rmhist p)` .| ($(S4 rmhist p) .| $(S6 rmhist p))`" + "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> (S3 rmhist p)` | (S4 rmhist p | S6 rmhist p)`" (fn _ => [split_idle_tac [] 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_3]) + auto_tac (MI_css addSDs2 [Step1_2_3]) ]); qed_goal "S3RPC_successors" MemoryImplementation.thy - "$(S3 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .& _(r p) \ -\ .-> ($(S4 rmhist p) .| $(S6 rmhist p))`" - (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSEs2 [action_conjimpE Step1_2_3]) ]); + "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ +\ & _(r p) \ +\ --> (S4 rmhist p | S6 rmhist p)`" + (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_3]) ]); qed_goal "S3RPC_enabled" MemoryImplementation.thy - "$(S3 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(Enabled (_(r p)))" - (fn _ => [cut_facts_tac [MI_base] 1, - auto_tac (MI_css addsimps2 [r_def,base_pair] - addSIs2 [RPCFail_Next_enabled,action_mp RPCFail_enabled]), - ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S3_def]) [] []) + "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> $Enabled (_(r p))" + (fn _ => [auto_tac (MI_css addsimps2 [r_def] + addSIs2 [RPCFail_Next_enabled,RPCFail_enabled]), + cut_facts_tac [MI_base] 1, + blast_tac (claset() addDs [base_pair]) 1, + ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S3_def])) ]); qed_goal "S3_live" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_) \ -\ .& WF(RPCNext crCh rmCh rst p)_(r p) \ -\ .-> ($(S3 rmhist p) ~> ($(S4 rmhist p) .| $(S6 rmhist p)))" + "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) & WF(RPCNext crCh rmCh rst p)_(r p) \ +\ --> (S3 rmhist p ~> S4 rmhist p | S6 rmhist p)" (fn _ => [REPEAT (resolve_tac [WF1,S3_successors,S3RPC_successors,S3RPC_enabled] 1)]); (* ------------- State S4 -------------------------------------------------- *) qed_goal "S4_successors" MemoryImplementation.thy - "$(S4 rmhist p) .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .-> $(S4 rmhist p)` .| $(S5 rmhist p)`" + "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ & (!l. $MemInv mm l) \ +\ --> (S4 rmhist p)` | (S5 rmhist p)`" (fn _ => [split_idle_tac [] 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_4]) + auto_tac (MI_css addSDs2 [Step1_2_4]) ]); (* ------------- State S4a: S4 /\ (ires p = NotAResult) ------------------------------ *) qed_goal "S4a_successors" MemoryImplementation.thy - "($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .-> ($(S4 rmhist p) .& ($(ires@p) .= #NotAResult))` \ -\ .| (($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p))`" + "|- $(S4 rmhist p & ires!p = #NotAResult) \ +\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l) \ +\ --> (S4 rmhist p & ires!p = #NotAResult)` \ +\ | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)`" (fn _ => [split_idle_tac [m_def] 1, - auto_tac (MI_css addsimps2 [m_def] addSEs2 [action_conjimpE Step1_2_4]) + auto_tac (MI_css addSDs2 [Step1_2_4]) ]); qed_goal "S4aRNext_successors" MemoryImplementation.thy - "($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .& _(m p) \ -\ .-> (($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p))`" + "|- ($(S4 rmhist p & ires!p = #NotAResult) \ +\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \ +\ & _(m p) \ +\ --> ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)`" (fn _ => [auto_tac (MI_css addsimps2 [angle_def] - addSEs2 [action_conjimpE Step1_2_4, - action_conjimpE ReadResult, action_impE WriteResult]) + addSDs2 [Step1_2_4, ReadResult, WriteResult]) ]); qed_goal "S4aRNext_enabled" MemoryImplementation.thy - "($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .-> $(Enabled (_(m p)))" - (fn _ => [auto_tac (MI_css addsimps2 [m_def] addSIs2 [action_mp RNext_enabled]), - ALLGOALS (cut_facts_tac [MI_base]), - auto_tac (MI_css addsimps2 [base_pair]), - (* it's faster to expand S4 only where necessary *) - action_simp_tac (simpset() addsimps [S_def,S4_def]) [] [] 1 + "|- $(S4 rmhist p & ires!p = #NotAResult) \ +\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l) \ +\ --> $Enabled (_(m p))" + (fn _ => [auto_tac (MI_css addsimps2 [m_def] addSIs2 [RNext_enabled]), + cut_facts_tac [MI_base] 1, + blast_tac (claset() addDs [base_pair]) 1, + asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1 ]); qed_goal "S4a_live" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_ .& (RALL l. $(MemInv mem l))) \ -\ .& WF(RNext rmCh mem ires p)_(m p) \ -\ .-> (($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ ~> ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p))" - (fn _ => [rtac WF1 1, - ALLGOALS (action_simp_tac (simpset()) - (map ((rewrite_rule [slice_def]) o action_mp) - [S4a_successors,S4aRNext_successors,S4aRNext_enabled]) - []) - ]); + "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \ +\ & WF(RNext rmCh mm ires p)_(m p) \ +\ --> (S4 rmhist p & ires!p = #NotAResult \ +\ ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)" + (K [REPEAT (resolve_tac [WF1, S4a_successors, S4aRNext_successors, S4aRNext_enabled] 1)]); (* ------------- State S4b: S4 /\ (ires p # NotAResult) ------------------------------ *) qed_goal "S4b_successors" MemoryImplementation.thy - "($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) \ -\ .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .-> ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult))` .| $(S5 rmhist p)`" + "|- $(S4 rmhist p & ires!p ~= #NotAResult) \ +\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l) \ +\ --> (S4 rmhist p & ires!p ~= #NotAResult)` | (S5 rmhist p)`" (fn _ => [split_idle_tac [m_def] 1, - auto_tac (MI_css addSEs2 (action_impE WriteResult - :: map action_conjimpE [Step1_2_4,ReadResult])) + auto_tac (MI_css addSDs2 [WriteResult,Step1_2_4,ReadResult]) ]); qed_goal "S4bReturn_successors" MemoryImplementation.thy - "($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) \ -\ .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .& _(m p) \ -\ .-> ($(S5 rmhist p))`" - (fn _ => [auto_tac (MI_css addsimps2 [angle_def] - addSEs2 [action_conjimpE Step1_2_4] - addEs2 [action_conjimpE ReturnNotReadWrite]) + "|- ($(S4 rmhist p & ires!p ~= #NotAResult) \ +\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \ +\ & _(m p) \ +\ --> (S5 rmhist p)`" + (fn _ => [force_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_4] + addDs2 [ReturnNotReadWrite]) 1 ]); qed_goal "S4bReturn_enabled" MemoryImplementation.thy - "($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) \ -\ .& (ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l))) \ -\ .-> $(Enabled (_(m p)))" - (fn _ => [cut_facts_tac [MI_base] 1, - auto_tac (MI_css addsimps2 [m_def,base_pair] - addSIs2 [action_mp MemReturn_enabled]), - ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S4_def]) [] []) + "|- $(S4 rmhist p & ires!p ~= #NotAResult) \ +\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l) \ +\ --> $Enabled (_(m p))" + (fn _ => [auto_tac (MI_css addsimps2 [m_def] addSIs2 [MemReturn_enabled]), + cut_facts_tac [MI_base] 1, + blast_tac (claset() addDs [base_pair]) 1, + asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1 ]); qed_goal "S4b_live" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_ .& (RALL l. $(MemInv mem l))) \ -\ .& WF(MemReturn rmCh ires p)_(m p) \ -\ .-> (($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) ~> $(S5 rmhist p))" - (fn _ => [rtac WF1 1, - ALLGOALS (action_simp_tac (simpset()) - (map ((rewrite_rule [slice_def]) o action_mp) - [S4b_successors,S4bReturn_successors,S4bReturn_enabled]) - [allE]) - ]); + "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (!l. $MemInv mm l)) \ +\ & WF(MemReturn rmCh ires p)_(m p) \ +\ --> (S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p)" + (K [REPEAT (resolve_tac [WF1, S4b_successors,S4bReturn_successors, S4bReturn_enabled] 1)]); (* ------------------------------ State S5 ------------------------------ *) qed_goal "S5_successors" MemoryImplementation.thy - "$(S5 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(S5 rmhist p)` .| $(S6 rmhist p)`" + "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> (S5 rmhist p)` | (S6 rmhist p)`" (fn _ => [split_idle_tac [] 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_5]) + auto_tac (MI_css addSDs2 [Step1_2_5]) ]); qed_goal "S5RPC_successors" MemoryImplementation.thy - "$(S5 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .& _(r p) \ -\ .-> $(S6 rmhist p)`" - (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSEs2 [action_conjimpE Step1_2_5]) ]); + "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ +\ & _(r p) \ +\ --> (S6 rmhist p)`" + (fn _ => [ auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_5]) ]); qed_goal "S5RPC_enabled" MemoryImplementation.thy - "$(S5 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(Enabled (_(r p)))" - (fn _ => [cut_facts_tac [MI_base] 1, - auto_tac (MI_css addsimps2 [r_def,base_pair] - addSIs2 [RPCFail_Next_enabled,action_mp RPCFail_enabled]), - ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S5_def]) [] []) + "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> $Enabled (_(r p))" + (fn _ => [auto_tac (MI_css addsimps2 [r_def] + addSIs2 [RPCFail_Next_enabled, RPCFail_enabled]), + cut_facts_tac [MI_base] 1, + blast_tac (claset() addDs [base_pair]) 1, + ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S5_def])) ]); qed_goal "S5_live" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_) \ -\ .& WF(RPCNext crCh rmCh rst p)_(r p) \ -\ .-> ($(S5 rmhist p) ~> $(S6 rmhist p))" + "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ +\ & WF(RPCNext crCh rmCh rst p)_(r p) \ +\ --> (S5 rmhist p ~> S6 rmhist p)" (fn _ => [REPEAT (resolve_tac [WF1,S5_successors,S5RPC_successors,S5RPC_enabled] 1)]); (* ------------------------------ State S6 ------------------------------ *) qed_goal "S6_successors" MemoryImplementation.thy - "$(S6 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .-> $(S1 rmhist p)` .| $(S3 rmhist p)` .| $(S6 rmhist p)`" + "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> (S1 rmhist p)` | (S3 rmhist p)` | (S6 rmhist p)`" (fn _ => [split_idle_tac [] 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_6]) + auto_tac (MI_css addSDs2 [Step1_2_6]) ]); qed_goal "S6MClkReply_successors" MemoryImplementation.thy - "$(S6 rmhist p) .& (ImpNext p .& [HNext rmhist p]_) \ -\ .& _(c p) \ -\ .-> $(S1 rmhist p)`" - (fn _ => [auto_tac (MI_css addsimps2 [angle_def] addSEs2 [action_conjimpE Step1_2_6, - action_impE MClkReplyNotRetry]) + "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ +\ & _(c p) \ +\ --> (S1 rmhist p)`" + (fn _ => [auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_6, MClkReplyNotRetry]) ]); qed_goal "MClkReplyS6" MemoryImplementation.thy - "$(ImpInv rmhist p) .& _(c p) .-> $(S6 rmhist p)" + "|- $ImpInv rmhist p & _(c p) --> $S6 rmhist p" (fn _ => [action_simp_tac (simpset() addsimps [angle_def,MClkReply_def,Return_def, @@ -275,108 +255,111 @@ ]); qed_goal "S6MClkReply_enabled" MemoryImplementation.thy - "$(S6 rmhist p) .-> $(Enabled (_(c p)))" - (fn _ => [cut_facts_tac [MI_base] 1, - auto_tac (MI_css addsimps2 [c_def,base_pair] - addSIs2 [action_mp MClkReply_enabled]), + "|- S6 rmhist p --> Enabled (_(c p))" + (fn _ => [auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkReply_enabled]), + cut_facts_tac [MI_base] 1, + blast_tac (claset() addDs [base_pair]) 1, ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S6_def]) [] []) ]); qed_goal "S6_live" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_ .& $(ImpInv rmhist p)) \ -\ .& SF(MClkReply memCh crCh cst p)_(c p) .& []<>($(S6 rmhist p)) \ -\ .-> []<>($(S1 rmhist p))" - (fn _ => [Auto_tac, + "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & $(ImpInv rmhist p)) \ +\ & SF(MClkReply memCh crCh cst p)_(c p) & []<>(S6 rmhist p) \ +\ --> []<>(S1 rmhist p)" + (fn _ => [Clarsimp_tac 1, subgoal_tac "sigma |= []<>(_(c p))" 1, - eres_inst_tac [("P","_(c p)")] - EnsuresInfinite 1, atac 1, + etac InfiniteEnsures 1, atac 1, action_simp_tac (simpset()) [] - (map action_conjimpE [MClkReplyS6,S6MClkReply_successors]) 1, + (map temp_elim [MClkReplyS6,S6MClkReply_successors]) 1, auto_tac (MI_css addsimps2 [SF_def]), etac swap 1, - auto_tac (MI_css addSIs2 [action_mp S6MClkReply_enabled] - addSEs2 [STL4E,DmdImplE]) + auto_tac (MI_css addSIs2 [S6MClkReply_enabled] addSEs2 [STL4E, DmdImplE]) ]); (* ------------------------------ complex leadsto properties ------------------------------ *) qed_goal "S5S6LeadstoS6" MemoryImplementation.thy - "!!sigma. (sigma |= $(S5 rmhist p) ~> $(S6 rmhist p)) \ -\ ==> (sigma |= ($(S5 rmhist p) .| $(S6 rmhist p)) ~> $(S6 rmhist p))" - (fn _ => [auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro, - temp_unlift LatticeReflexivity]) + "!!sigma. sigma |= S5 rmhist p ~> S6 rmhist p \ +\ ==> sigma |= (S5 rmhist p | S6 rmhist p) ~> S6 rmhist p" + (fn _ => [auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro, LatticeReflexivity]) ]); qed_goal "S4bS5S6LeadstoS6" MemoryImplementation.thy - "!!sigma. [| (sigma |= ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) ~> $(S5 rmhist p)); \ -\ (sigma |= $(S5 rmhist p) ~> $(S6 rmhist p)) |] \ -\ ==> (sigma |= (($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p) .| $(S6 rmhist p)) ~> $(S6 rmhist p))" + "!!sigma. [| sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p; \ +\ sigma |= S5 rmhist p ~> S6 rmhist p |] \ +\ ==> sigma |= (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p \ +\ ~> S6 rmhist p" (fn _ => [auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro,S5S6LeadstoS6] addIs2 [LatticeTransitivity]) ]); qed_goal "S4S5S6LeadstoS6" MemoryImplementation.thy - "!!sigma. [| (sigma |= ($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ ~> ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p)); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) ~> $(S5 rmhist p)); \ -\ (sigma |= $(S5 rmhist p) ~> $(S6 rmhist p)) |] \ -\ ==> (sigma |= ($(S4 rmhist p) .| $(S5 rmhist p) .| $(S6 rmhist p)) ~> $(S6 rmhist p))" - (fn _ => [subgoal_tac "sigma |= (($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) .| ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p) .| $(S6 rmhist p)) ~> $(S6 rmhist p)" 1, - eres_inst_tac [("G", "($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) .| ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p) .| $(S6 rmhist p)")] LatticeTransitivity 1, - SELECT_GOAL (auto_tac (MI_css addSIs2 [ImplLeadsto, temp_unlift necT])) 1, - rtac LatticeDisjunctionIntro 1, - etac LatticeTransitivity 1, - etac LatticeTriangle 1, atac 1, + "!!sigma. [| sigma |= S4 rmhist p & ires!p = #NotAResult \ +\ ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p; \ +\ sigma |= S5 rmhist p ~> S6 rmhist p |] \ +\ ==> sigma |= S4 rmhist p | S5 rmhist p | S6 rmhist p ~> S6 rmhist p" + (fn _ => [subgoal_tac "sigma |= (S4 rmhist p & ires!p = #NotAResult) | (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p ~> S6 rmhist p" 1, + eres_inst_tac [("G", "PRED ((S4 rmhist p & ires!p = #NotAResult) | (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p)")] (temp_use LatticeTransitivity) 1, + force_tac (MI_css addsimps2 Init_defs addSIs2 [ImplLeadsto_gen, necT]) 1, + rtac (temp_use LatticeDisjunctionIntro) 1, + etac (temp_use LatticeTransitivity) 1, + etac (temp_use LatticeTriangle2) 1, atac 1, auto_tac (MI_css addSIs2 [S4bS5S6LeadstoS6]) ]); qed_goal "S3S4S5S6LeadstoS6" MemoryImplementation.thy - "!!sigma. [| (sigma |= $(S3 rmhist p) ~> ($(S4 rmhist p) .| $(S6 rmhist p))); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ ~> ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p)); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) ~> $(S5 rmhist p)); \ -\ (sigma |= $(S5 rmhist p) ~> $(S6 rmhist p)) |] \ -\ ==> (sigma |= ($(S3 rmhist p) .| $(S4 rmhist p) .| $(S5 rmhist p) .| $(S6 rmhist p)) ~> $(S6 rmhist p))" - (fn _ => [rtac LatticeDisjunctionIntro 1, - rtac LatticeTriangle 1, atac 2, - rtac (S4S5S6LeadstoS6 RS LatticeTransitivity) 1, - auto_tac (MI_css addSIs2 [S4S5S6LeadstoS6,temp_unlift necT] - addIs2 [ImplLeadsto]) + "!!sigma. [| sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p = #NotAResult \ +\ ~> (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p; \ +\ sigma |= S5 rmhist p ~> S6 rmhist p |] \ +\ ==> sigma |= S3 rmhist p | S4 rmhist p | S5 rmhist p | S6 rmhist p ~> S6 rmhist p" + (fn _ => [rtac (temp_use LatticeDisjunctionIntro) 1, + etac (temp_use LatticeTriangle2) 1, + rtac (S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1, + auto_tac (MI_css addSIs2 [S4S5S6LeadstoS6,necT] + addIs2 [ImplLeadsto_gen] addsimps2 Init_defs) ]); qed_goal "S2S3S4S5S6LeadstoS6" MemoryImplementation.thy - "!!sigma. [| (sigma |= $(S2 rmhist p) ~> $(S3 rmhist p)); \ -\ (sigma |= $(S3 rmhist p) ~> ($(S4 rmhist p) .| $(S6 rmhist p))); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ ~> ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p)); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) ~> $(S5 rmhist p)); \ -\ (sigma |= $(S5 rmhist p) ~> $(S6 rmhist p)) |] \ -\ ==> (sigma |= ($(S2 rmhist p) .| $(S3 rmhist p) .| $(S4 rmhist p) .| $(S5 rmhist p) .| $(S6 rmhist p)) ~> $(S6 rmhist p))" - (fn _ => [rtac LatticeDisjunctionIntro 1, - rtac LatticeTransitivity 1, atac 2, - rtac (S3S4S5S6LeadstoS6 RS LatticeTransitivity) 1, - auto_tac (MI_css addSIs2 [S3S4S5S6LeadstoS6,temp_unlift necT] - addIs2 [ImplLeadsto]) + "!!sigma. [| sigma |= S2 rmhist p ~> S3 rmhist p; \ +\ sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p = #NotAResult \ +\ ~> S4 rmhist p & ires!p ~= #NotAResult | S5 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p; \ +\ sigma |= S5 rmhist p ~> S6 rmhist p |] \ +\ ==> sigma |= S2 rmhist p | S3 rmhist p | S4 rmhist p | S5 rmhist p | S6 rmhist p \ +\ ~> S6 rmhist p" + (fn _ => [rtac (temp_use LatticeDisjunctionIntro) 1, + rtac (temp_use LatticeTransitivity) 1, atac 2, + rtac (S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1, + auto_tac (MI_css addSIs2 [S3S4S5S6LeadstoS6,necT] + addIs2 [ImplLeadsto_gen] addsimps2 Init_defs) ]); qed_goal "NotS1LeadstoS6" MemoryImplementation.thy - "!!sigma. [| (sigma |= []($(ImpInv rmhist p))); \ -\ (sigma |= $(S2 rmhist p) ~> $(S3 rmhist p)); \ -\ (sigma |= $(S3 rmhist p) ~> ($(S4 rmhist p) .| $(S6 rmhist p))); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .= #NotAResult)) \ -\ ~> ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) .| $(S5 rmhist p)); \ -\ (sigma |= ($(S4 rmhist p) .& ($(ires@p) .~= #NotAResult)) ~> $(S5 rmhist p)); \ -\ (sigma |= $(S5 rmhist p) ~> $(S6 rmhist p)) |] \ -\ ==> (sigma |= .~$(S1 rmhist p) ~> $(S6 rmhist p))" - (fn _ => [rtac (S2S3S4S5S6LeadstoS6 RS LatticeTransitivity) 1, - auto_tac (MI_css addsimps2 [ImpInv_def] addIs2 [ImplLeadsto] addSEs2 [STL4E]) + "!!sigma. [| sigma |= []ImpInv rmhist p; \ +\ sigma |= S2 rmhist p ~> S3 rmhist p; \ +\ sigma |= S3 rmhist p ~> S4 rmhist p | S6 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p = #NotAResult \ +\ ~> S4 rmhist p & ires!p ~= #NotAResult | S5 rmhist p; \ +\ sigma |= S4 rmhist p & ires!p ~= #NotAResult ~> S5 rmhist p; \ +\ sigma |= S5 rmhist p ~> S6 rmhist p |] \ +\ ==> sigma |= ~S1 rmhist p ~> S6 rmhist p" + (fn _ => [rtac (S2S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1, + TRYALL atac, + etac (temp_use INV_leadsto) 1, + rtac (temp_use ImplLeadsto_gen) 1, + rtac (temp_use necT) 1, + auto_tac (MI_css addsimps2 ImpInv_def::Init_defs addSIs2 [necT]) ]); qed_goal "S1Infinite" MemoryImplementation.thy - "!!sigma. [| (sigma |= .~$(S1 rmhist p) ~> $(S6 rmhist p)); \ -\ (sigma |= []<>($(S6 rmhist p)) .-> []<>($(S1 rmhist p))) |] \ -\ ==> (sigma |= []<>($(S1 rmhist p)))" + "!!sigma. [| sigma |= ~S1 rmhist p ~> S6 rmhist p; \ +\ sigma |= []<>S6 rmhist p --> []<>S1 rmhist p |] \ +\ ==> sigma |= []<>S1 rmhist p" (fn _ => [rtac classical 1, - asm_full_simp_tac (simpset() addsimps [NotBox, temp_rewrite NotDmd]) 1, - auto_tac (MI_css addSEs2 [mp,leadsto_infinite] addSDs2 [temp_mp DBImplBDAct]) + asm_full_simp_tac (simpset() addsimps [temp_use NotBox, NotDmd]) 1, + auto_tac (MI_css addSEs2 [mp,leadsto_infinite] addSDs2 [DBImplBD]) ]); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MIsafe.ML --- a/src/HOL/TLA/Memory/MIsafe.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MIsafe.ML Mon Feb 08 13:02:56 1999 +0100 @@ -10,19 +10,15 @@ (* RPCFailure notin MemVals U {OK,BadArg} *) -qed_goal "MVOKBAnotRF" MemoryImplementation.thy +qed_goalw "MVOKBAnotRF" MemoryImplementation.thy [MVOKBA_def] "!!x. MVOKBA x ==> x ~= RPCFailure" - (fn _ => [ auto_tac (HOL_css addsimps2 (RP_simps @ [MVOKBA_def])) ]); -bind_thm("MVOKBAnotRFE", make_elim MVOKBAnotRF); + (fn _ => [ Auto_tac ]); (* NotAResult notin MemVals U {OK,BadArg,RPCFailure} *) -qed_goal "MVOKBARFnotNR" MemoryImplementation.thy +qed_goalw "MVOKBARFnotNR" MemoryImplementation.thy [MVOKBARF_def] "!!x. MVOKBARF x ==> x ~= NotAResult" - (fn _ => [ auto_tac (HOL_css addsimps2 (RP_simps @ [MVOKBARF_def]) - addSEs2 [MemValNotAResultE]) - ]); -bind_thm("MVOKBARFnotNRE", make_elim MVOKBARFnotNR); + (fn _ => [ Auto_tac ]); (* ========================= Si's are mutually exclusive ==================================== *) (* Si and Sj are mutually exclusive for i # j. This helps to simplify the big @@ -33,240 +29,186 @@ (* --- not used --- qed_goal "S1_excl" MemoryImplementation.thy - "$(S1 rmhist p) .-> $(S1 rmhist p) .& .~$(S2 rmhist p) .& .~$(S3 rmhist p) .& \ -\ .~$(S4 rmhist p) .& .~$(S5 rmhist p) .& .~$(S6 rmhist p)" + "|- S1 rmhist p --> S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p & \ +\ ~S4 rmhist p & ~S5 rmhist p & ~S6 rmhist p" (fn _ => [ auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, S3_def, S4_def, S5_def, S6_def]) ]); *) qed_goal "S2_excl" MemoryImplementation.thy - "$(S2 rmhist p) .-> $(S2 rmhist p) .& .~$(S1 rmhist p)" + "|- S2 rmhist p --> S2 rmhist p & ~S1 rmhist p" (fn _ => [ auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def]) ]); -bind_thm("S2_exclE", action_impE S2_excl); qed_goal "S3_excl" MemoryImplementation.thy - "$(S3 rmhist p) .-> $(S3 rmhist p) .& .~$(S1 rmhist p) .& .~$(S2 rmhist p)" + "|- S3 rmhist p --> S3 rmhist p & ~S1 rmhist p & ~S2 rmhist p" (fn _ => [ auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, S3_def]) ]); -bind_thm("S3_exclE", action_impE S3_excl); qed_goal "S4_excl" MemoryImplementation.thy - "$(S4 rmhist p) .-> $(S4 rmhist p) .& .~$(S1 rmhist p) .& .~$(S2 rmhist p) .& .~$(S3 rmhist p)" + "|- S4 rmhist p --> S4 rmhist p & ~S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p" (fn _ => [ auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def]) ]); -bind_thm("S4_exclE", action_impE S4_excl); qed_goal "S5_excl" MemoryImplementation.thy - "$(S5 rmhist p) .-> $(S5 rmhist p) .& .~$(S1 rmhist p) .& .~$(S2 rmhist p) .& \ -\ .~$(S3 rmhist p) .& .~$(S4 rmhist p)" + "|- S5 rmhist p --> S5 rmhist p & ~S1 rmhist p & ~S2 rmhist p \ +\ & ~S3 rmhist p & ~S4 rmhist p" (fn _ => [ auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def]) ]); -bind_thm("S5_exclE", action_impE S5_excl); qed_goal "S6_excl" MemoryImplementation.thy - "$(S6 rmhist p) .-> $(S6 rmhist p) .& .~$(S1 rmhist p) .& .~$(S2 rmhist p) .& \ -\ .~$(S3 rmhist p) .& .~$(S4 rmhist p) .& .~$(S5 rmhist p)" + "|- S6 rmhist p --> S6 rmhist p & ~S1 rmhist p & ~S2 rmhist p \ +\ & ~S3 rmhist p & ~S4 rmhist p & ~S5 rmhist p" (fn _ => [ auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]) ]); -bind_thm("S6_exclE", action_impE S6_excl); (* ==================== Lemmas about the environment ============================== *) qed_goal "Envbusy" MemoryImplementation.thy - "$(Calling memCh p) .-> .~ ENext p" + "|- $(Calling memCh p) --> ~ENext p" (fn _ => [ auto_tac (MI_css addsimps2 [ENext_def,Call_def]) ]); (* ==================== Lemmas about the implementation's states ==================== *) (* The following series of lemmas are used in establishing the implementation's next-state relation (Step 1.2 of the proof in the paper). For each state Si, we - establish which component actions are possible and their results. + determine which component actions are possible and what state they result in. *) (* ------------------------------ State S1 ---------------------------------------- *) qed_goal "S1Env" MemoryImplementation.thy - "(ENext p) .& $(S1 rmhist p) .& unchanged .-> (S2 rmhist p)$" - (fn _ => [auto_tac (MI_css - addsimps2 [ENext_def,Call_def,c_def,r_def,m_def, - caller_def,rtrner_def,MVNROKBA_def, - S_def,S1_def,S2_def,Calling_def]) + "|- ENext p & $(S1 rmhist p) & unchanged (c p, r p, m p, rmhist!p) --> (S2 rmhist p)$" + (fn _ => [force_tac (MI_css + addsimps2 [ENext_def,Call_def,c_def,r_def,m_def, + caller_def,rtrner_def,MVNROKBA_def, + S_def,S1_def,S2_def,Calling_def]) 1 ]); -bind_thm("S1EnvE", action_conjimpE S1Env); qed_goal "S1ClerkUnch" MemoryImplementation.thy - "[MClkNext memCh crCh cst p]_(c p) .& $(S1 rmhist p) .-> unchanged (c p)" - (fn _ => [auto_tac (MI_fast_css addSEs2 [action_conjimpE MClkidle] - addsimps2 [square_def,S_def,S1_def]) - ]); -bind_thm("S1ClerkUnchE", action_conjimpE S1ClerkUnch); + "|- [MClkNext memCh crCh cst p]_(c p) & $(S1 rmhist p) --> unchanged (c p)" + (fn _ => [auto_tac (MI_fast_css addSDs2 [MClkidle] addsimps2 [S_def,S1_def]) ]); qed_goal "S1RPCUnch" MemoryImplementation.thy - "[RPCNext crCh rmCh rst p]_(r p) .& $(S1 rmhist p) .-> unchanged (r p)" - (fn _ => [auto_tac (MI_fast_css addSEs2 [action_impE RPCidle] - addsimps2 [square_def,S_def,S1_def]) - ]); -bind_thm("S1RPCUnchE", action_conjimpE S1RPCUnch); + "|- [RPCNext crCh rmCh rst p]_(r p) & $(S1 rmhist p) --> unchanged (r p)" + (fn _ => [auto_tac (MI_fast_css addSDs2 [RPCidle] addsimps2 [S_def,S1_def]) ]); qed_goal "S1MemUnch" MemoryImplementation.thy - "[RNext rmCh mem ires p]_(m p) .& $(S1 rmhist p) .-> unchanged (m p)" - (fn _ => [auto_tac (MI_fast_css addSEs2 [action_impE Memoryidle] - addsimps2 [square_def,S_def,S1_def]) - ]); -bind_thm("S1MemUnchE", action_conjimpE S1MemUnch); + "|- [RNext rmCh mm ires p]_(m p) & $(S1 rmhist p) --> unchanged (m p)" + (fn _ => [auto_tac (MI_fast_css addSDs2 [Memoryidle] addsimps2 [S_def,S1_def]) ]); qed_goal "S1Hist" MemoryImplementation.thy - "[HNext rmhist p]_ .& $(S1 rmhist p) .-> unchanged (rmhist@p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,HNext_def,MemReturn_def, - RPCFail_def,MClkReply_def,Return_def, - S_def,S1_def]) + "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S1 rmhist p) --> unchanged (rmhist!p)" + (fn _ => [action_simp_tac (simpset() addsimps [HNext_def, S_def, S1_def, MemReturn_def, + RPCFail_def,MClkReply_def,Return_def]) + [] [squareE] 1 ]); -bind_thm("S1HistE", action_conjimpE S1Hist); (* ------------------------------ State S2 ---------------------------------------- *) qed_goal "S2EnvUnch" MemoryImplementation.thy - "[ENext p]_(e p) .& $(S2 rmhist p) .-> unchanged (e p)" - (fn _ => [auto_tac (MI_fast_css addSEs2 [action_impE Envbusy] - addsimps2 [square_def,S_def,S2_def]) - ]); -bind_thm("S2EnvUnchE", action_conjimpE S2EnvUnch); + "|- [ENext p]_(e p) & $(S2 rmhist p) --> unchanged (e p)" + (fn _ => [auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S2_def]) ]); qed_goal "S2Clerk" MemoryImplementation.thy - "MClkNext memCh crCh cst p .& $(S2 rmhist p) .-> MClkFwd memCh crCh cst p" - (fn _ => [auto_tac (MI_fast_css addsimps2 [MClkNext_def,MClkRetry_def,MClkReply_def, - S_def,S2_def]) + "|- MClkNext memCh crCh cst p & $(S2 rmhist p) --> MClkFwd memCh crCh cst p" + (fn _ => [auto_tac (MI_css addsimps2 [MClkNext_def,MClkRetry_def,MClkReply_def, + S_def,S2_def]) ]); -bind_thm("S2ClerkE", action_conjimpE S2Clerk); -(* The dumb action_simp_tac wins 15 : 129 over auto_tac *) qed_goal "S2Forward" MemoryImplementation.thy - "$(S2 rmhist p) .& (MClkFwd memCh crCh cst p) .& unchanged \ -\ .-> (S3 rmhist p)$" + "|- $(S2 rmhist p) & MClkFwd memCh crCh cst p & unchanged (e p, r p, m p, rmhist!p) \ +\ --> (S3 rmhist p)$" (fn _ => [action_simp_tac (simpset() addsimps [MClkFwd_def,Call_def,e_def,r_def,m_def,caller_def,rtrner_def, S_def,S2_def,S3_def,Calling_def]) [] [] 1 ]); -bind_thm("S2ForwardE", action_conjimpE S2Forward); qed_goal "S2RPCUnch" MemoryImplementation.thy - "[RPCNext crCh rmCh rst p]_(r p) .& $(S2 rmhist p) .-> unchanged (r p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S2_def] - addSEs2 [action_impE RPCidle]) - ]); -bind_thm("S2RPCUnchE", action_conjimpE S2RPCUnch); + "|- [RPCNext crCh rmCh rst p]_(r p) & $(S2 rmhist p) --> unchanged (r p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [RPCidle]) ]); qed_goal "S2MemUnch" MemoryImplementation.thy - "[RNext rmCh mem ires p]_(m p) .& $(S2 rmhist p) .-> unchanged (m p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S2_def] - addSEs2 [action_impE Memoryidle]) - ]); -bind_thm("S2MemUnchE", action_conjimpE S2MemUnch); + "|- [RNext rmCh mm ires p]_(m p) & $(S2 rmhist p) --> unchanged (m p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [Memoryidle]) ]); qed_goal "S2Hist" MemoryImplementation.thy - "[HNext rmhist p]_ .& $(S2 rmhist p) .-> unchanged (rmhist@p)" + "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S2 rmhist p) --> unchanged (rmhist!p)" (fn _ => [auto_tac (MI_fast_css - addsimps2 [square_def,HNext_def,MemReturn_def, + addsimps2 [HNext_def,MemReturn_def, RPCFail_def,MClkReply_def,Return_def,S_def,S2_def]) ]); -bind_thm("S2HistE", action_conjimpE S2Hist); (* ------------------------------ State S3 ---------------------------------------- *) qed_goal "S3EnvUnch" MemoryImplementation.thy - "[ENext p]_(e p) .& $(S3 rmhist p) .-> unchanged (e p)" - (fn _ => [auto_tac (MI_fast_css addSEs2 [action_impE Envbusy] - addsimps2 [square_def,S_def,S3_def]) - ]); -bind_thm("S3EnvUnchE", action_conjimpE S3EnvUnch); + "|- [ENext p]_(e p) & $(S3 rmhist p) --> unchanged (e p)" + (fn _ => [auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S3_def]) ]); qed_goal "S3ClerkUnch" MemoryImplementation.thy - "[MClkNext memCh crCh cst p]_(c p) .& $(S3 rmhist p) .-> unchanged (c p)" - (fn _ => [auto_tac (MI_fast_css addSEs2 [action_impE MClkbusy] - addsimps2 [square_def,S_def,S3_def]) - ]); -bind_thm("S3ClerkUnchE", action_conjimpE S3ClerkUnch); + "|- [MClkNext memCh crCh cst p]_(c p) & $(S3 rmhist p) --> unchanged (c p)" + (fn _ => [auto_tac (MI_css addSDs2 [MClkbusy] addsimps2 [square_def,S_def,S3_def]) ]); qed_goal "S3LegalRcvArg" MemoryImplementation.thy - "$(S3 rmhist p) .-> IsLegalRcvArg[ arg[$(crCh@p)] ]" - (fn _ => [action_simp_tac - (simpset() addsimps [IsLegalRcvArg_def,MClkRelayArg_def,S_def,S3_def]) - [exI] [] 1 - ]); + "|- S3 rmhist p --> IsLegalRcvArg>" + (fn _ => [auto_tac (MI_css addsimps2 [IsLegalRcvArg_def,MClkRelayArg_def,S_def,S3_def]) ]); qed_goal "S3RPC" MemoryImplementation.thy - "(RPCNext crCh rmCh rst p) .& $(S3 rmhist p) \ -\ .-> (RPCFwd crCh rmCh rst p) .| (RPCFail crCh rmCh rst p)" - (fn _ => [auto_tac MI_css, - etac ((rewrite_rule action_rews (S3LegalRcvArg RS actionD)) RS impdupE) 1, + "|- RPCNext crCh rmCh rst p & $(S3 rmhist p) \ +\ --> RPCFwd crCh rmCh rst p | RPCFail crCh rmCh rst p" + (fn _ => [Clarsimp_tac 1, + forward_tac [action_use S3LegalRcvArg] 1, auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCReply_def,S_def,S3_def]) ]); -bind_thm("S3RPCE", action_conjimpE S3RPC); qed_goal "S3Forward" MemoryImplementation.thy - "(RPCFwd crCh rmCh rst p) .& HNext rmhist p .& $(S3 rmhist p) .& unchanged \ -\ .-> (S4 rmhist p)$ .& unchanged (rmhist@p)" + "|- RPCFwd crCh rmCh rst p & HNext rmhist p & $(S3 rmhist p) & unchanged (e p, c p, m p) \ +\ --> (S4 rmhist p)$ & unchanged (rmhist!p)" (fn _ => [action_simp_tac (simpset() addsimps [RPCFwd_def,HNext_def,MemReturn_def,RPCFail_def,MClkReply_def, Return_def,Call_def,e_def,c_def,m_def,caller_def,rtrner_def, S_def,S3_def,S4_def,Calling_def]) [] [] 1 ]); -bind_thm("S3ForwardE", action_conjimpE S3Forward); qed_goal "S3Fail" MemoryImplementation.thy - "(RPCFail crCh rmCh rst p) .& $(S3 rmhist p) .& HNext rmhist p .& unchanged \ -\ .-> (S6 rmhist p)$" + "|- RPCFail crCh rmCh rst p & $(S3 rmhist p) & HNext rmhist p & unchanged (e p, c p, m p) \ +\ --> (S6 rmhist p)$" (fn _ => [action_simp_tac (simpset() addsimps [HNext_def,RPCFail_def,Return_def,e_def,c_def,m_def, caller_def,rtrner_def,MVOKBARF_def, S_def,S3_def,S6_def,Calling_def]) [] [] 1 ]); -bind_thm("S3FailE", action_conjimpE S3Fail); qed_goal "S3MemUnch" MemoryImplementation.thy - "[RNext rmCh mem ires p]_(m p) .& $(S3 rmhist p) .-> unchanged (m p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S3_def] - addSEs2 [action_impE Memoryidle]) - ]); -bind_thm("S3MemUnchE", action_conjimpE S3MemUnch); + "|- [RNext rmCh mm ires p]_(m p) & $(S3 rmhist p) --> unchanged (m p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S3_def] addSDs2 [Memoryidle]) ]); qed_goal "S3Hist" MemoryImplementation.thy - "HNext rmhist p .& $(S3 rmhist p) .& unchanged (r p) .-> unchanged (rmhist@p)" - (fn _ => [auto_tac (MI_fast_css + "|- HNext rmhist p & $(S3 rmhist p) & unchanged (r p) --> unchanged (rmhist!p)" + (fn _ => [auto_tac (MI_css addsimps2 [HNext_def,MemReturn_def,RPCFail_def,MClkReply_def, Return_def,r_def,rtrner_def,S_def,S3_def,Calling_def]) ]); -bind_thm("S3HistE", action_conjimpE S3Hist); (* ------------------------------ State S4 ---------------------------------------- *) qed_goal "S4EnvUnch" MemoryImplementation.thy - "[ENext p]_(e p) .& $(S4 rmhist p) .-> unchanged (e p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S4_def] - addSEs2 [action_impE Envbusy]) - ]); -bind_thm("S4EnvUnchE", action_conjimpE S4EnvUnch); + "|- [ENext p]_(e p) & $(S4 rmhist p) --> unchanged (e p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [Envbusy]) ]); qed_goal "S4ClerkUnch" MemoryImplementation.thy - "[MClkNext memCh crCh cst p]_(c p) .& $(S4 rmhist p) .-> unchanged (c p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S4_def] - addSEs2 [action_impE MClkbusy]) - ]); -bind_thm("S4ClerkUnchE", action_conjimpE S4ClerkUnch); + "|- [MClkNext memCh crCh cst p]_(c p) & $(S4 rmhist p) --> unchanged (c p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [MClkbusy]) ]); qed_goal "S4RPCUnch" MemoryImplementation.thy - "[RPCNext crCh rmCh rst p]_(r p) .& $(S4 rmhist p) .-> unchanged (r p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S4_def] - addSEs2 [action_conjimpE RPCbusy]) - ]); -bind_thm("S4RPCUnchE", action_conjimpE S4RPCUnch); + "|- [RPCNext crCh rmCh rst p]_(r p) & $(S4 rmhist p) --> unchanged (r p)" + (fn _ => [auto_tac (MI_fast_css addsimps2 [S_def,S4_def] addSDs2 [RPCbusy]) ]); qed_goal "S4ReadInner" MemoryImplementation.thy - "(ReadInner rmCh mem ires p l) .& $(S4 rmhist p) .& unchanged \ -\ .& (HNext rmhist p) .& $(MemInv mem l) \ -\ .-> (S4 rmhist p)$ .& unchanged (rmhist@p)" + "|- ReadInner rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p) \ +\ & HNext rmhist p & $(MemInv mm l) \ +\ --> (S4 rmhist p)$ & unchanged (rmhist!p)" (fn _ => [action_simp_tac (simpset() addsimps [ReadInner_def,GoodRead_def, BadRead_def,HNext_def, MemReturn_def, RPCFail_def,MClkReply_def,Return_def, @@ -276,17 +218,15 @@ ]); qed_goal "S4Read" MemoryImplementation.thy - "(Read rmCh mem ires p) .& $(S4 rmhist p) .& unchanged \ -\ .& (HNext rmhist p) .& (RALL l. $(MemInv mem l)) \ -\ .-> (S4 rmhist p)$ .& unchanged (rmhist@p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [Read_def] - addSEs2 [action_conjimpE S4ReadInner]) - ]); -bind_thm("S4ReadE", action_conjimpE S4Read); + "|- Read rmCh mm ires p & $(S4 rmhist p) & unchanged (e p, c p, r p) \ +\ & HNext rmhist p & (!l. $MemInv mm l) \ +\ --> (S4 rmhist p)$ & unchanged (rmhist!p)" + (fn _ => [auto_tac (MI_css addsimps2 [Read_def] addSDs2 [S4ReadInner]) ]); qed_goal "S4WriteInner" MemoryImplementation.thy - "(WriteInner rmCh mem ires p l v) .& $(S4 rmhist p) .& unchanged .& (HNext rmhist p) \ -\ .-> (S4 rmhist p)$ .& unchanged (rmhist@p)" + "|- WriteInner rmCh mm ires p l v & $(S4 rmhist p) & unchanged (e p, c p, r p) \ +\ & HNext rmhist p \ +\ --> (S4 rmhist p)$ & unchanged (rmhist!p)" (fn _ => [action_simp_tac (simpset() addsimps [WriteInner_def,GoodWrite_def, BadWrite_def,HNext_def, MemReturn_def,RPCFail_def,MClkReply_def,Return_def, @@ -296,64 +236,53 @@ ]); qed_goal "S4Write" MemoryImplementation.thy - "(Write rmCh mem ires p l) .& $(S4 rmhist p) .& unchanged .& (HNext rmhist p) \ -\ .-> (S4 rmhist p)$ .& unchanged (rmhist@p)" - (fn _ => [ auto_tac (MI_css addsimps2 [Write_def] addSEs2 [action_conjimpE S4WriteInner]) ]); -bind_thm("S4WriteE", action_conjimpE S4Write); + "|- Write rmCh mm ires p l & $(S4 rmhist p) & unchanged (e p, c p, r p) & (HNext rmhist p) \ +\ --> (S4 rmhist p)$ & unchanged (rmhist!p)" + (fn _ => [ auto_tac (MI_css addsimps2 [Write_def] addSDs2 [S4WriteInner]) ]); qed_goal "WriteS4" MemoryImplementation.thy - "$(ImpInv rmhist p) .& (Write rmCh mem ires p l) .-> $(S4 rmhist p)" - (fn _ => [auto_tac (MI_fast_css + "|- $ImpInv rmhist p & Write rmCh mm ires p l --> $S4 rmhist p" + (fn _ => [auto_tac (MI_css addsimps2 [Write_def,WriteInner_def,ImpInv_def,WrRequest_def, S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def]) ]); -bind_thm("WriteS4E", action_conjimpE WriteS4); qed_goal "S4Return" MemoryImplementation.thy - "(MemReturn rmCh ires p) .& $(S4 rmhist p) .& unchanged .& (HNext rmhist p) \ -\ .-> (S5 rmhist p)$" - (fn _ => [auto_tac (MI_fast_css + "|- MemReturn rmCh ires p & $S4 rmhist p & unchanged (e p, c p, r p) & HNext rmhist p \ +\ --> (S5 rmhist p)$" + (fn _ => [auto_tac (MI_css addsimps2 [HNext_def,MemReturn_def,Return_def,e_def,c_def,r_def, rtrner_def,caller_def,MVNROKBA_def,MVOKBA_def, S_def,S4_def,S5_def,Calling_def]) ]); -bind_thm("S4ReturnE", action_conjimpE S4Return); qed_goal "S4Hist" MemoryImplementation.thy - "(HNext rmhist p) .& $(S4 rmhist p) .& (m p)$ .= $(m p) .-> (rmhist@p)$ .= $(rmhist@p)" - (fn _ => [auto_tac (MI_fast_css + "|- HNext rmhist p & $S4 rmhist p & (m p)$ = $(m p) --> (rmhist!p)$ = $(rmhist!p)" + (fn _ => [auto_tac (MI_css addsimps2 [HNext_def,MemReturn_def,RPCFail_def,MClkReply_def, Return_def,m_def,rtrner_def,S_def,S4_def,Calling_def]) ]); -bind_thm("S4HistE", action_conjimpE S4Hist); (* ------------------------------ State S5 ---------------------------------------- *) qed_goal "S5EnvUnch" MemoryImplementation.thy - "[ENext p]_(e p) .& $(S5 rmhist p) .-> unchanged (e p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S5_def] - addSEs2 [action_impE Envbusy]) - ]); -bind_thm("S5EnvUnchE", action_conjimpE S5EnvUnch); + "|- [ENext p]_(e p) & $(S5 rmhist p) --> unchanged (e p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Envbusy]) ]); qed_goal "S5ClerkUnch" MemoryImplementation.thy - "[MClkNext memCh crCh cst p]_(c p) .& $(S5 rmhist p) .-> unchanged (c p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S5_def] - addSEs2 [action_impE MClkbusy]) - ]); -bind_thm("S5ClerkUnchE", action_conjimpE S5ClerkUnch); + "|- [MClkNext memCh crCh cst p]_(c p) & $(S5 rmhist p) --> unchanged (c p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [MClkbusy]) ]); qed_goal "S5RPC" MemoryImplementation.thy - "(RPCNext crCh rmCh rst p) .& $(S5 rmhist p) \ -\ .-> (RPCReply crCh rmCh rst p) .| (RPCFail crCh rmCh rst p)" - (fn _ => [auto_tac (MI_fast_css + "|- RPCNext crCh rmCh rst p & $(S5 rmhist p) \ +\ --> RPCReply crCh rmCh rst p | RPCFail crCh rmCh rst p" + (fn _ => [auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCFwd_def,S_def,S5_def]) ]); -bind_thm("S5RPCE", action_conjimpE S5RPC); qed_goal "S5Reply" MemoryImplementation.thy - "(RPCReply crCh rmCh rst p) .& $(S5 rmhist p) .& unchanged \ -\ .-> (S6 rmhist p)$" + "|- RPCReply crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) \ +\ --> (S6 rmhist p)$" (fn _ => [action_simp_tac (simpset() addsimps [RPCReply_def,Return_def,e_def,c_def,m_def, @@ -361,11 +290,10 @@ S_def,S5_def,S6_def,Calling_def]) [] [] 1 ]); -bind_thm("S5ReplyE", action_conjimpE S5Reply); qed_goal "S5Fail" MemoryImplementation.thy - "(RPCFail crCh rmCh rst p) .& $(S5 rmhist p) .& unchanged \ -\ .-> (S6 rmhist p)$" + "|- RPCFail crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) \ +\ --> (S6 rmhist p)$" (fn _ => [action_simp_tac (simpset() addsimps [RPCFail_def,Return_def,e_def,c_def,m_def, @@ -373,77 +301,60 @@ S_def,S5_def,S6_def,Calling_def]) [] [] 1 ]); -bind_thm("S5FailE", action_conjimpE S5Fail); qed_goal "S5MemUnch" MemoryImplementation.thy - "[RNext rmCh mem ires p]_(m p) .& $(S5 rmhist p) .-> unchanged (m p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S5_def] - addSEs2 [action_impE Memoryidle]) - ]); -bind_thm("S5MemUnchE", action_conjimpE S5MemUnch); + "|- [RNext rmCh mm ires p]_(m p) & $(S5 rmhist p) --> unchanged (m p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Memoryidle]) ]); qed_goal "S5Hist" MemoryImplementation.thy - "[HNext rmhist p]_ .& $(S5 rmhist p) .-> (rmhist@p)$ .= $(rmhist@p)" + "|- [HNext rmhist p]_(c p, r p, m p, rmhist!p) & $(S5 rmhist p) --> (rmhist!p)$ = $(rmhist!p)" (fn _ => [auto_tac (MI_fast_css - addsimps2 [square_def,HNext_def,MemReturn_def, + addsimps2 [HNext_def,MemReturn_def, RPCFail_def,MClkReply_def,Return_def,S_def,S5_def]) ]); -bind_thm("S5HistE", action_conjimpE S5Hist); (* ------------------------------ State S6 ---------------------------------------- *) qed_goal "S6EnvUnch" MemoryImplementation.thy - "[ENext p]_(e p) .& $(S6 rmhist p) .-> unchanged (e p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S6_def] - addSEs2 [action_impE Envbusy]) - ]); -bind_thm("S6EnvUnchE", action_conjimpE S6EnvUnch); + "|- [ENext p]_(e p) & $(S6 rmhist p) --> unchanged (e p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Envbusy]) ]); qed_goal "S6Clerk" MemoryImplementation.thy - "(MClkNext memCh crCh cst p) .& $(S6 rmhist p) \ -\ .-> (MClkRetry memCh crCh cst p) .| (MClkReply memCh crCh cst p)" - (fn _ => [ auto_tac (MI_fast_css addsimps2 [MClkNext_def,MClkFwd_def,S_def,S6_def]) ]); -bind_thm("S6ClerkE", action_conjimpE S6Clerk); + "|- MClkNext memCh crCh cst p & $(S6 rmhist p) \ +\ --> MClkRetry memCh crCh cst p | MClkReply memCh crCh cst p" + (fn _ => [ auto_tac (MI_css addsimps2 [MClkNext_def,MClkFwd_def,S_def,S6_def]) ]); qed_goal "S6Retry" MemoryImplementation.thy - "(MClkRetry memCh crCh cst p) .& (HNext rmhist p) .& $(S6 rmhist p) .& unchanged \ -\ .-> (S3 rmhist p)$ .& unchanged (rmhist@p)" + "|- MClkRetry memCh crCh cst p & HNext rmhist p & $S6 rmhist p & unchanged (e p,r p,m p) \ +\ --> (S3 rmhist p)$ & unchanged (rmhist!p)" (fn _ => [action_simp_tac (simpset() addsimps [HNext_def,MClkReply_def,MClkRetry_def,Call_def, Return_def,e_def,r_def,m_def,caller_def,rtrner_def, S_def,S6_def,S3_def,Calling_def]) [] [] 1]); -bind_thm("S6RetryE", action_conjimpE S6Retry); qed_goal "S6Reply" MemoryImplementation.thy - "(MClkReply memCh crCh cst p) .& (HNext rmhist p) .& $(S6 rmhist p) .& unchanged \ -\ .-> (S1 rmhist p)$" + "|- MClkReply memCh crCh cst p & HNext rmhist p & $S6 rmhist p & unchanged (e p,r p,m p) \ +\ --> (S1 rmhist p)$" (fn _ => [action_simp_tac (simpset() addsimps [HNext_def,MemReturn_def,RPCFail_def,Return_def, MClkReply_def,e_def,r_def,m_def,caller_def,rtrner_def, S_def,S6_def,S1_def,Calling_def]) [] [] 1 ]); -bind_thm("S6ReplyE", action_conjimpE S6Reply); qed_goal "S6RPCUnch" MemoryImplementation.thy - "[RPCNext crCh rmCh rst p]_(r p) .& $(S6 rmhist p) .-> unchanged (r p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S6_def] - addSEs2 [action_impE RPCidle]) - ]); -bind_thm("S6RPCUnchE", action_conjimpE S6RPCUnch); + "|- [RPCNext crCh rmCh rst p]_(r p) & $S6 rmhist p --> unchanged (r p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [RPCidle]) ]); qed_goal "S6MemUnch" MemoryImplementation.thy - "[RNext rmCh mem ires p]_(m p) .& $(S6 rmhist p) .-> unchanged (m p)" - (fn _ => [auto_tac (MI_fast_css addsimps2 [square_def,S_def,S6_def] - addSEs2 [action_impE Memoryidle]) - ]); -bind_thm("S6MemUnchE", action_conjimpE S6MemUnch); + "|- [RNext rmCh mm ires p]_(m p) & $(S6 rmhist p) --> unchanged (m p)" + (fn _ => [auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Memoryidle]) ]); qed_goal "S6Hist" MemoryImplementation.thy - "(HNext rmhist p) .& $(S6 rmhist p) .& (c p)$ .= $(c p) .-> (rmhist@p)$ .= $(rmhist@p)" - (fn _ => [auto_tac (MI_fast_css + "|- HNext rmhist p & $S6 rmhist p & (c p)$ = $(c p) --> (rmhist!p)$ = $(rmhist!p)" + (fn _ => [auto_tac (MI_css addsimps2 [HNext_def,MClkReply_def,Return_def,c_def,rtrner_def, S_def,S6_def,Calling_def]) ]); -bind_thm("S6HistE", action_conjimpE S6Hist); + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemClerk.ML --- a/src/HOL/TLA/Memory/MemClerk.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemClerk.ML Mon Feb 08 13:02:56 1999 +0100 @@ -3,69 +3,60 @@ Author: Stephan Merz Copyright: 1997 University of Munich - RPC-Memory example: Memory clerk specification (ML file) + RPC-Memory example: Memory clerk specification (theorems and proofs) *) val MC_action_defs = - [MClkInit_def RS inteq_reflection] - @ [MClkFwd_def, MClkRetry_def, MClkReply_def, MClkNext_def]; + [MClkInit_def, MClkFwd_def, MClkRetry_def, MClkReply_def, MClkNext_def]; val MC_temp_defs = [MClkIPSpec_def, MClkISpec_def]; +val mem_css = (claset(), simpset()); + (* The Clerk engages in an action for process p only if there is an outstanding, unanswered call for that process. *) qed_goal "MClkidle" MemClerk.thy - ".~ $(Calling send p) .& ($(cst@p) .= #clkA) .-> .~ MClkNext send rcv cst p" - (fn _ => [ auto_tac (claset(), - simpset() addsimps (MC_action_defs @ [Return_def])) - ]); + "|- ~$Calling send p & $(cst!p) = #clkA --> ~MClkNext send rcv cst p" + (fn _ => [ auto_tac (mem_css addsimps2 (Return_def::MC_action_defs)) ]); qed_goal "MClkbusy" MemClerk.thy - "$(Calling rcv p) .-> .~ MClkNext send rcv cst p" - (fn _ => [ auto_tac (claset(), - simpset() addsimps (MC_action_defs @ [Call_def])) - ]); - -(* unlifted versions as introduction rules *) - -bind_thm("MClkidleI", action_mp MClkidle); -bind_thm("MClkbusyI", action_mp MClkbusy); + "|- $Calling rcv p --> ~MClkNext send rcv cst p" + (fn _ => [ auto_tac (mem_css addsimps2 (MC_action_defs @ [Call_def])) ]); (* Enabledness of actions *) qed_goal "MClkFwd_enabled" MemClerk.thy - "!!p. base_var ==> \ -\ $(Calling send p) .& .~ $(Calling rcv p) .& ($(cst@p) .= #clkA) \ -\ .-> $(Enabled (MClkFwd send rcv cst p))" + "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> \ +\ |- Calling send p & ~Calling rcv p & cst!p = #clkA \ +\ --> Enabled (MClkFwd send rcv cst p)" (fn _ => [action_simp_tac (simpset() addsimps [MClkFwd_def,Call_def,caller_def,rtrner_def]) [] [base_enabled,Pair_inject] 1]); qed_goal "MClkFwd_ch_enabled" MemClerk.thy - "Enabled (MClkFwd send rcv cst p) s \ -\ ==> Enabled (_) s" - (fn [prem] => [auto_tac (claset() addSIs [prem RS enabled_mono], - simpset() addsimps [angle_def,MClkFwd_def]) - ]); + "|- Enabled (MClkFwd send rcv cst p) --> \ +\ Enabled (_(cst!p, rtrner send!p, caller rcv!p))" + (fn _ => [auto_tac (mem_css addSEs2 [enabled_mono] + addsimps2 [angle_def,MClkFwd_def]) + ]); qed_goal "MClkReply_change" MemClerk.thy - "MClkReply send rcv cst p .-> _" - (fn _ => [auto_tac (action_css addsimps2 [angle_def,MClkReply_def] - addEs2 [Return_changedE]) + "|- MClkReply send rcv cst p --> _(cst!p, rtrner send!p, caller rcv!p)" + (fn _ => [auto_tac (mem_css addsimps2 [angle_def,MClkReply_def] + addEs2 [Return_changed]) ]); qed_goal "MClkReply_enabled" MemClerk.thy - "!!p. base_var ==> \ -\ $(Calling send p) .& .~ $(Calling rcv p) .& ($(cst@p) .= #clkB) \ -\ .-> $(Enabled (_))" + "!!p. basevars (rtrner send!p, caller rcv!p, cst!p) ==> \ +\ |- Calling send p & ~Calling rcv p & cst!p = #clkB \ +\ --> Enabled (_(cst!p, rtrner send!p, caller rcv!p))" (fn _ => [action_simp_tac (simpset()) [MClkReply_change RSN (2,enabled_mono)] [] 1, action_simp_tac (simpset() addsimps [MClkReply_def,Return_def,caller_def,rtrner_def]) [] [base_enabled,Pair_inject] 1 ]); qed_goal "MClkReplyNotRetry" MemClerk.thy - "MClkReply send rcv cst p .-> .~(MClkRetry send rcv cst p)" - (fn _ => [ auto_tac (claset(), - simpset() addsimps [MClkReply_def,MClkRetry_def]) - ]); + "|- MClkReply send rcv cst p --> ~MClkRetry send rcv cst p" + (fn _ => [ auto_tac (mem_css addsimps2 [MClkReply_def,MClkRetry_def]) ]); + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemClerk.thy --- a/src/HOL/TLA/Memory/MemClerk.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemClerk.thy Mon Feb 08 13:02:56 1999 +0100 @@ -17,56 +17,53 @@ mClkRcvChType = "rpcSndChType" mClkStType = "(PrIds => mClkState) stfun" -consts +constdefs (* state predicates *) MClkInit :: "mClkRcvChType => mClkStType => PrIds => stpred" + "MClkInit rcv cst p == PRED (cst!p = #clkA & ~Calling rcv p)" (* actions *) MClkFwd :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" + "MClkFwd send rcv cst p == ACT + $Calling send p + & $(cst!p) = #clkA + & Call rcv p MClkRelayArg> + & (cst!p)$ = #clkB + & unchanged (rtrner send!p)" + MClkRetry :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" + "MClkRetry send rcv cst p == ACT + $(cst!p) = #clkB + & res<$(rcv!p)> = #RPCFailure + & Call rcv p MClkRelayArg> + & unchanged (cst!p, rtrner send!p)" + MClkReply :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" + "MClkReply send rcv cst p == ACT + ~$Calling rcv p + & $(cst!p) = #clkB + & Return send p MClkReplyVal> + & (cst!p)$ = #clkA + & unchanged (caller rcv!p)" + MClkNext :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => action" + "MClkNext send rcv cst p == ACT + ( MClkFwd send rcv cst p + | MClkRetry send rcv cst p + | MClkReply send rcv cst p)" + (* temporal *) MClkIPSpec :: "mClkSndChType => mClkRcvChType => mClkStType => PrIds => temporal" - MClkISpec :: "mClkSndChType => mClkRcvChType => mClkStType => temporal" - -rules - MClkInit_def "$(MClkInit rcv cst p) .= - ($(cst@p) .= #clkA .& .~ $(Calling rcv p))" - - MClkFwd_def "MClkFwd send rcv cst p == - $(Calling send p) - .& $(cst@p) .= #clkA - .& Call rcv p (MClkRelayArg[ arg[$(send@p)] ]) - .& (cst@p)$ .= #clkB - .& unchanged (rtrner send @ p)" - - MClkRetry_def "MClkRetry send rcv cst p == - $(cst@p) .= #clkB - .& res[$(rcv@p)] .= #RPCFailure - .& Call rcv p (MClkRelayArg[ arg[$(send@p)] ]) - .& unchanged " + "MClkIPSpec send rcv cst p == TEMP + Init MClkInit rcv cst p + & [][ MClkNext send rcv cst p ]_(cst!p, rtrner send!p, caller rcv!p) + & WF(MClkFwd send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p) + & SF(MClkReply send rcv cst p)_(cst!p, rtrner send!p, caller rcv!p)" - MClkReply_def "MClkReply send rcv cst p == - .~ $(Calling rcv p) - .& $(cst@p) .= #clkB - .& Return send p (MClkReplyVal[ res[$(rcv@p)] ]) - .& (cst@p)$ .= #clkA - .& unchanged (caller rcv @ p)" + MClkISpec :: "mClkSndChType => mClkRcvChType => mClkStType => temporal" + "MClkISpec send rcv cst == TEMP (!p. MClkIPSpec send rcv cst p)" - MClkNext_def "MClkNext send rcv cst p == - MClkFwd send rcv cst p - .| MClkRetry send rcv cst p - .| MClkReply send rcv cst p" - - MClkIPSpec_def "MClkIPSpec send rcv cst p == - Init($(MClkInit rcv cst p)) - .& [][ MClkNext send rcv cst p ]_ - .& WF(MClkFwd send rcv cst p)_ - .& SF(MClkReply send rcv cst p)_" - - MClkISpec_def "MClkISpec send rcv cst == RALL p. MClkIPSpec send rcv cst p" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemClerkParameters.ML --- a/src/HOL/TLA/Memory/MemClerkParameters.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemClerkParameters.ML Mon Feb 08 13:02:56 1999 +0100 @@ -6,6 +6,7 @@ RPC-Memory example: Memory clerk parameters (ML file) *) +(* val CP_simps = RP_simps @ mClkState.simps; - +*) diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemClerkParameters.thy --- a/src/HOL/TLA/Memory/MemClerkParameters.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemClerkParameters.thy Mon Feb 08 13:02:56 1999 +0100 @@ -16,19 +16,16 @@ types (* types sent on the clerk's send and receive channels are argument types of the memory and the RPC, respectively *) - mClkSndArgType = "memArgType" - mClkRcvArgType = "rpcArgType" + mClkSndArgType = "memOp" + mClkRcvArgType = "rpcOp" -consts +constdefs (* translate a memory call to an RPC call *) - MClkRelayArg :: "memArgType => rpcArgType" + MClkRelayArg :: "memOp => rpcOp" + "MClkRelayArg marg == memcall marg" (* translate RPC failures to memory failures *) MClkReplyVal :: "Vals => Vals" - -rules - MClkRelayArg_def "MClkRelayArg marg == Inl (remoteCall, marg)" - MClkReplyVal_def "MClkReplyVal v == - if v = RPCFailure then MemFailure else v" + "MClkReplyVal v == if v = RPCFailure then MemFailure else v" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/Memory.ML --- a/src/HOL/TLA/Memory/Memory.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/Memory.ML Mon Feb 08 13:02:56 1999 +0100 @@ -3,49 +3,45 @@ Author: Stephan Merz Copyright: 1997 University of Munich - RPC-Memory example: Memory specification (ML file) + RPC-Memory example: Memory specification (theorems and proofs) *) val RM_action_defs = - (map (fn t => t RS inteq_reflection) - [MInit_def, PInit_def, RdRequest_def, WrRequest_def, MemInv_def]) - @ [GoodRead_def, BadRead_def, ReadInner_def, Read_def, - GoodWrite_def, BadWrite_def, WriteInner_def, Write_def, - MemReturn_def, RNext_def]; + [MInit_def, PInit_def, RdRequest_def, WrRequest_def, MemInv_def, + GoodRead_def, BadRead_def, ReadInner_def, Read_def, + GoodWrite_def, BadWrite_def, WriteInner_def, Write_def, + MemReturn_def, RNext_def]; val UM_action_defs = RM_action_defs @ [MemFail_def, UNext_def]; val RM_temp_defs = [RPSpec_def, MSpec_def, IRSpec_def]; val UM_temp_defs = [UPSpec_def, MSpec_def, IUSpec_def]; -(* Make sure the simpset accepts non-boolean simplifications *) -simpset_ref() := simpset() setmksimps ((mksimps mksimps_pairs) o maybe_unlift); +val mem_css = (claset(), simpset()); (* -------------------- Proofs ---------------------------------------------- *) (* The reliable memory is an implementation of the unreliable one *) qed_goal "ReliableImplementsUnReliable" Memory.thy - "IRSpec ch mm rs .-> IUSpec ch mm rs" - (K [force_tac (temp_css addsimps2 ([square_def,UNext_def] @ - RM_temp_defs @ UM_temp_defs) addSEs2 [STL4E]) 1]); + "|- IRSpec ch mm rs --> IUSpec ch mm rs" + (K [force_tac (temp_css addsimps2 ([UNext_def,UPSpec_def,IUSpec_def] @ RM_temp_defs) + addSEs2 [STL4E,squareE]) 1]); (* The memory spec implies the memory invariant *) qed_goal "MemoryInvariant" Memory.thy - "(MSpec ch mm rs l) .-> []($(MemInv mm l))" - (fn _ => [ auto_inv_tac (simpset() addsimps RM_temp_defs @ - MP_simps @ RM_action_defs) 1 ]); + "|- MSpec ch mm rs l --> [](MemInv mm l)" + (fn _ => [ auto_inv_tac (simpset() addsimps RM_temp_defs @ RM_action_defs) 1 ]); (* The invariant is trivial for non-locations *) qed_goal "NonMemLocInvariant" Memory.thy - ".~ #(MemLoc l) .-> []($MemInv mm l)" - (K [ auto_tac (temp_css addsimps2 [MemInv_def] addSIs2 [necT RS tempD]) ]); + "|- #l ~: #MemLoc --> [](MemInv mm l)" + (K [ auto_tac (temp_css addsimps2 [MemInv_def] addSIs2 [necT]) ]); qed_goal "MemoryInvariantAll" Memory.thy - "((RALL l. #(MemLoc l) .-> MSpec ch mm rs l)) .-> (RALL l. []($MemInv mm l))" + "|- (!l. #l : #MemLoc --> MSpec ch mm rs l) --> (!l. [](MemInv mm l))" (K [step_tac temp_cs 1, - case_tac "MemLoc l" 1, - auto_tac (temp_css addSEs2 (map temp_mp [MemoryInvariant, - NonMemLocInvariant]))]); + case_tac "l : MemLoc" 1, + auto_tac (temp_css addSEs2 [MemoryInvariant,NonMemLocInvariant]) ]); (* The memory engages in an action for process p only if there is an unanswered call from p. @@ -53,35 +49,28 @@ *) qed_goal "Memoryidle" Memory.thy - ".~ $(Calling ch p) .-> .~ RNext ch mm rs p" - (K [ auto_tac (action_css addsimps2 (RM_action_defs @ [Return_def])) ]); - -bind_thm("MemoryidleI", action_mp Memoryidle); -bind_thm("MemoryidleE", action_impE Memoryidle); - + "|- ~$(Calling ch p) --> ~ RNext ch mm rs p" + (K [ auto_tac (mem_css addsimps2 (Return_def::RM_action_defs)) ]); (* Enabledness conditions *) qed_goal "MemReturn_change" Memory.thy - "MemReturn ch rs p .-> _" - (K [ force_tac (action_css addsimps2 [MemReturn_def,angle_def]) 1]); + "|- MemReturn ch rs p --> _(rtrner ch ! p, rs!p)" + (K [ force_tac (mem_css addsimps2 [MemReturn_def,angle_def]) 1]); qed_goal "MemReturn_enabled" Memory.thy - "!!p. base_var ==> \ -\ $(Calling ch p) .& ($(rs@p) .~= #NotAResult) \ -\ .-> $(Enabled (_))" + "!!p. basevars (rtrner ch ! p, rs!p) ==> \ +\ |- Calling ch p & (rs!p ~= #NotAResult) \ +\ --> Enabled (_(rtrner ch ! p, rs!p))" (K [action_simp_tac (simpset()) [MemReturn_change RSN (2,enabled_mono)] [] 1, action_simp_tac (simpset() addsimps [MemReturn_def,Return_def,rtrner_def]) [] [base_enabled,Pair_inject] 1 ]); qed_goal "ReadInner_enabled" Memory.thy - "!!p. base_var ==> \ -\ $(Calling ch p) .& (arg[$(ch@p)] .= #(Inl (read,l))) \ -\ .-> $(Enabled (ReadInner ch mm rs p l))" - (fn _ => [Action_simp_tac 1, -(* unlift before applying case_tac: case_split_thm expects boolean conclusion *) - case_tac "MemLoc l" 1, + "!!p. basevars (rtrner ch ! p, rs!p) ==> \ +\ |- Calling ch p & (arg = #(read l)) --> Enabled (ReadInner ch mm rs p l)" + (fn _ => [case_tac "l : MemLoc" 1, ALLGOALS (action_simp_tac (simpset() addsimps [ReadInner_def,GoodRead_def,BadRead_def, @@ -90,11 +79,10 @@ ]); qed_goal "WriteInner_enabled" Memory.thy - "!!p. base_var ==> \ -\ $(Calling ch p) .& (arg[$(ch@p)] .= #(Inr (write,l,v))) \ -\ .-> $(Enabled (WriteInner ch mm rs p l v))" - (fn _ => [Action_simp_tac 1, - case_tac "MemLoc l & MemVal v" 1, + "!!p. basevars (mm!l, rtrner ch ! p, rs!p) ==> \ +\ |- Calling ch p & (arg = #(write l v)) \ +\ --> Enabled (WriteInner ch mm rs p l v)" + (fn _ => [case_tac "l:MemLoc & v:MemVal" 1, ALLGOALS (action_simp_tac (simpset() addsimps [WriteInner_def,GoodWrite_def,BadWrite_def, WrRequest_def] delsimps [disj_not1]) @@ -102,57 +90,45 @@ ]); qed_goal "ReadResult" Memory.thy - "(Read ch mm rs p) .& (RALL l. $(MemInv mm l)) .-> (rs@p)$ .~= #NotAResult" - (fn _ => [action_simp_tac - (simpset() addsimps (MP_simps - @ [Read_def,ReadInner_def,GoodRead_def, - BadRead_def,MemInv_def])) - [] [] 1, - auto_tac (action_css addsimps2 MP_simps) ]); + "|- Read ch mm rs p & (!l. $(MemInv mm l)) --> (rs!p)` ~= #NotAResult" + (fn _ => [force_tac (mem_css addsimps2 + [Read_def,ReadInner_def,GoodRead_def,BadRead_def,MemInv_def]) 1]); qed_goal "WriteResult" Memory.thy - "(Write ch mm rs p l) .-> (rs@p)$ .~= #NotAResult" - (fn _ => [auto_tac (claset(), - simpset() addsimps (MP_simps @ - [Write_def,WriteInner_def,GoodWrite_def,BadWrite_def])) + "|- Write ch mm rs p l --> (rs!p)` ~= #NotAResult" + (fn _ => [auto_tac (mem_css addsimps2 ([Write_def,WriteInner_def,GoodWrite_def,BadWrite_def])) ]); qed_goal "ReturnNotReadWrite" Memory.thy - "(RALL l. $MemInv mm l) .& (MemReturn ch rs p) \ -\ .-> .~(Read ch mm rs p) .& (RALL l. .~(Write ch mm rs p l))" + "|- (!l. $MemInv mm l) & MemReturn ch rs p \ +\ --> ~ Read ch mm rs p & (!l. ~ Write ch mm rs p l)" (fn _ => [auto_tac - (action_css addsimps2 [MemReturn_def] - addSEs2 [action_impE WriteResult,action_conjimpE ReadResult]) + (mem_css addsimps2 [MemReturn_def] addSDs2 [WriteResult, ReadResult]) ]); qed_goal "RWRNext_enabled" Memory.thy - "($(rs@p) .= #NotAResult) .& (RALL l. $(MemInv mm l)) \ -\ .& $(Enabled (Read ch mm rs p .| (REX l. Write ch mm rs p l))) \ -\ .-> $(Enabled (_))" - (K [force_tac (action_css addsimps2 [RNext_def,angle_def] + "|- (rs!p = #NotAResult) & (!l. MemInv mm l) \ +\ & Enabled (Read ch mm rs p | (? l. Write ch mm rs p l)) \ +\ --> Enabled (_(rtrner ch ! p, rs!p))" + (K [force_tac (mem_css addsimps2 [RNext_def,angle_def] addSEs2 [enabled_mono2] - addEs2[action_conjimpE ReadResult,action_impE WriteResult]) 1]); + addDs2 [ReadResult, WriteResult]) 1]); (* Combine previous lemmas: the memory can make a visible step if there is an outstanding call for which no result has been produced. *) qed_goal "RNext_enabled" Memory.thy -"!!p. (ALL l. base_var ) ==> \ -\ ($(rs@p) .= #NotAResult) .& $(Calling ch p) .& (RALL l. $(MemInv mm l)) \ -\ .-> $(Enabled (_))" (K [ - auto_tac (action_css addsimps2 [enabled_disj] - addSIs2 [action_mp RWRNext_enabled]), - res_inst_tac [("s","arg(ch s p)")] sumE 1, - action_simp_tac (simpset()addsimps[Read_def,enabled_ex,base_pair]) - [action_mp ReadInner_enabled,exI] [] 1, - split_all_tac 1, rename_tac "a b" 1, induct_tac "a" 1, - (* introduce a trivial subgoal to solve flex-flex constraint?! *) - subgoal_tac "b = snd(a,b)" 1, - TRYALL Simp_tac, (* solves "read" case *) +"!!p. !l. basevars (mm!l, rtrner ch!p, rs!p) ==> \ +\ |- (rs!p = #NotAResult) & Calling ch p & (!l. MemInv mm l) \ +\ --> Enabled (_(rtrner ch ! p, rs!p))" (K [ + auto_tac (mem_css addsimps2 [enabled_disj] + addSIs2 [RWRNext_enabled]), + exhaust_tac "arg(ch w p)" 1, + action_simp_tac (simpset()addsimps[Read_def,enabled_ex]) + [ReadInner_enabled,exI] [] 1, + force_tac (mem_css addDs2 [base_pair]) 1, etac swap 1, - action_simp_tac (simpset()addsimps[Write_def,enabled_ex,base_pair]) - [action_mp WriteInner_enabled,exI] [] 1, - split_all_tac 1, rename_tac "a aa b" 1, induct_tac "a" 1, - subgoal_tac "(aa = fst(snd(a,aa,b))) & (b = snd(snd(a,aa,b)))" 1, - ALLGOALS Simp_tac ]); + action_simp_tac (simpset() addsimps [Write_def,enabled_ex]) + [WriteInner_enabled,exI] [] 1]); + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/Memory.thy --- a/src/HOL/TLA/Memory/Memory.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/Memory.thy Mon Feb 08 13:02:56 1999 +0100 @@ -12,7 +12,7 @@ Memory = MemoryParameters + ProcedureInterface + types - memChType = "(memArgType,Vals) channel" + memChType = "(memOp, Vals) channel" memType = "(Locs => Vals) stfun" (* intention: MemLocs => MemVals *) resType = "(PrIds => Vals) stfun" @@ -55,82 +55,83 @@ MemInv :: "memType => Locs => stpred" rules - MInit_def "$(MInit mm l) .= ($(mm@l) .= # InitVal)" - PInit_def "$(PInit rs p) .= ($(rs@p) .= # NotAResult)" + MInit_def "MInit mm l == PRED mm!l = #InitVal" + PInit_def "PInit rs p == PRED rs!p = #NotAResult" - RdRequest_def "$(RdRequest ch p l) .= - ($(Calling ch p) .& (arg[$(ch@p)] .= #(Inl (read,l))))" - WrRequest_def "$(WrRequest ch p l v) .= - ($(Calling ch p) .& (arg[$(ch@p)] .= #(Inr (write,l,v))))" + RdRequest_def "RdRequest ch p l == PRED + Calling ch p & (arg = #(read l))" + WrRequest_def "WrRequest ch p l v == PRED + Calling ch p & (arg = #(write l v))" (* a read that doesn't raise BadArg *) - GoodRead_def "GoodRead mm rs p l == - #(MemLoc l) .& (rs@p)$ .= $(mm@l)" + GoodRead_def "GoodRead mm rs p l == ACT + #l : #MemLoc & ((rs!p)$ = $(mm!l))" (* a read that raises BadArg *) - BadRead_def "BadRead mm rs p l == - .~ #(MemLoc l) .& (rs@p)$ .= #BadArg" + BadRead_def "BadRead mm rs p l == ACT + #l ~: #MemLoc & ((rs!p)$ = #BadArg)" (* the read action with l visible *) - ReadInner_def "ReadInner ch mm rs p l == + ReadInner_def "ReadInner ch mm rs p l == ACT $(RdRequest ch p l) - .& (GoodRead mm rs p l .| BadRead mm rs p l) - .& unchanged (rtrner ch @ p)" + & (GoodRead mm rs p l | BadRead mm rs p l) + & unchanged (rtrner ch ! p)" (* the read action with l quantified *) - Read_def "Read ch mm rs p == REX l. ReadInner ch mm rs p l" + Read_def "Read ch mm rs p == ACT (? l. ReadInner ch mm rs p l)" (* similar definitions for the write action *) - GoodWrite_def "GoodWrite mm rs p l v == - #(MemLoc l) .& #(MemVal v) - .& (mm@l)$ .= #v .& (rs@p)$ .= #OK" - BadWrite_def "BadWrite mm rs p l v == - .~ (#(MemLoc l) .& #(MemVal v)) - .& (rs@p)$ .= #BadArg .& unchanged (mm@l)" - WriteInner_def "WriteInner ch mm rs p l v == + GoodWrite_def "GoodWrite mm rs p l v == ACT + #l : #MemLoc & #v : #MemVal + & ((mm!l)$ = #v) & ((rs!p)$ = #OK)" + BadWrite_def "BadWrite mm rs p l v == ACT + ~ (#l : #MemLoc & #v : #MemVal) + & ((rs!p)$ = #BadArg) & unchanged (mm!l)" + WriteInner_def "WriteInner ch mm rs p l v == ACT $(WrRequest ch p l v) - .& (GoodWrite mm rs p l v .| BadWrite mm rs p l v) - .& unchanged (rtrner ch @ p)" - Write_def "Write ch mm rs p l == REX v. WriteInner ch mm rs p l v" + & (GoodWrite mm rs p l v | BadWrite mm rs p l v) + & unchanged (rtrner ch ! p)" + Write_def "Write ch mm rs p l == ACT (? v. WriteInner ch mm rs p l v)" (* the return action *) - MemReturn_def "MemReturn ch rs p == - $(rs@p) .~= #NotAResult - .& (rs@p)$ .= #NotAResult - .& Return ch p ($(rs@p))" + MemReturn_def "MemReturn ch rs p == ACT + ( ($(rs!p) ~= #NotAResult) + & ((rs!p)$ = #NotAResult) + & Return ch p (rs!p))" + (* the failure action of the unreliable memory *) - MemFail_def "MemFail ch rs p == + MemFail_def "MemFail ch rs p == ACT $(Calling ch p) - .& (rs@p)$ .= #MemFailure - .& unchanged (rtrner ch @ p)" - RNext_def "RNext ch mm rs p == - Read ch mm rs p - .| (REX l. Write ch mm rs p l) - .| MemReturn ch rs p" - UNext_def "UNext ch mm rs p == - RNext ch mm rs p .| MemFail ch rs p" + & ((rs!p)$ = #MemFailure) + & unchanged (rtrner ch ! p)" + (* next-state relations for reliable / unreliable memory *) + RNext_def "RNext ch mm rs p == ACT + ( Read ch mm rs p + | (? l. Write ch mm rs p l) + | MemReturn ch rs p)" + UNext_def "UNext ch mm rs p == ACT + (RNext ch mm rs p | MemFail ch rs p)" - RPSpec_def "RPSpec ch mm rs p == - Init($(PInit rs p)) - .& [][ RNext ch mm rs p ]_ - .& WF(RNext ch mm rs p)_ - .& WF(MemReturn ch rs p)_" - UPSpec_def "UPSpec ch mm rs p == - Init($(PInit rs p)) - .& [][ UNext ch mm rs p ]_ - .& WF(RNext ch mm rs p)_ - .& WF(MemReturn ch rs p)_" - MSpec_def "MSpec ch mm rs l == - Init($(MInit mm l)) - .& [][ REX p. Write ch mm rs p l ]_(mm@l)" - IRSpec_def "IRSpec ch mm rs == - (RALL p. RPSpec ch mm rs p) - .& (RALL l. #(MemLoc l) .-> MSpec ch mm rs l)" - IUSpec_def "IUSpec ch mm rs == - (RALL p. UPSpec ch mm rs p) - .& (RALL l. #(MemLoc l) .-> MSpec ch mm rs l)" + RPSpec_def "RPSpec ch mm rs p == TEMP + Init(PInit rs p) + & [][ RNext ch mm rs p ]_(rtrner ch ! p, rs!p) + & WF(RNext ch mm rs p)_(rtrner ch ! p, rs!p) + & WF(MemReturn ch rs p)_(rtrner ch ! p, rs!p)" + UPSpec_def "UPSpec ch mm rs p == TEMP + Init(PInit rs p) + & [][ UNext ch mm rs p ]_(rtrner ch ! p, rs!p) + & WF(RNext ch mm rs p)_(rtrner ch ! p, rs!p) + & WF(MemReturn ch rs p)_(rtrner ch ! p, rs!p)" + MSpec_def "MSpec ch mm rs l == TEMP + Init(MInit mm l) + & [][ ? p. Write ch mm rs p l ]_(mm!l)" + IRSpec_def "IRSpec ch mm rs == TEMP + (! p. RPSpec ch mm rs p) + & (! l. #l : #MemLoc --> MSpec ch mm rs l)" + IUSpec_def "IUSpec ch mm rs == TEMP + (! p. UPSpec ch mm rs p) + & (! l. #l : #MemLoc --> MSpec ch mm rs l)" - RSpec_def "RSpec ch rs == EEX mm. IRSpec ch mm rs" - USpec_def "USpec ch == EEX mm rs. IUSpec ch mm rs" + RSpec_def "RSpec ch rs == TEMP (EEX mm. IRSpec ch mm rs)" + USpec_def "USpec ch == TEMP (EEX mm rs. IUSpec ch mm rs)" - MemInv_def "$(MemInv mm l) .= - (#(MemLoc l) .-> MemVal[ $(mm@l)])" + MemInv_def "MemInv mm l == PRED #l : #MemLoc --> mm!l : #MemVal" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemoryImplementation.ML --- a/src/HOL/TLA/Memory/MemoryImplementation.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemoryImplementation.ML Mon Feb 08 13:02:56 1999 +0100 @@ -13,90 +13,50 @@ Steps are (roughly) numbered as in the hand proof. *) - -(* ------------------------------ HOL lemmas ------------------------------ *) -(* Add the following simple lemmas as default simplification rules. *) - -section "Auxiliary lemmas"; - -qed_goal "equal_false_not" HOL.thy "(P = False) = (~P)" - (fn _ => [fast_tac prop_cs 1]); - -Addsimps [equal_false_not]; - - -(* A variant of the implication elimination rule that keeps the antecedent. - Use "thm RS impdupE" to generate an unsafe (looping) elimination rule. -*) - -qed_goal "impdupE" HOL.thy - "[| P --> Q; P; [| P;Q |] ==> R |] ==> R" - (fn maj::prems => [REPEAT (resolve_tac ([maj RS mp] @ prems) 1)]); +(* --------------------------- automatic prover --------------------------- *) - -(* Introduction/elimination rules for if-then-else *) - -qed_goal "ifI" HOL.thy - "[| Q ==> P(x); ~Q ==> P(y) |] ==> P(if Q then x else y)" - (fn prems => [case_tac "Q" 1, ALLGOALS (Asm_simp_tac THEN' (eresolve_tac prems))]); - -qed_goal "ifE" HOL.thy - "[| P(if Q then x else y); [| Q; P(x) |] ==> R; [| ~Q; P(y) |] ==> R |] ==> R" - (fn (prem1::prems) => [case_tac "Q" 1, - ALLGOALS ((cut_facts_tac [prem1]) - THEN' Asm_full_simp_tac - THEN' (REPEAT o ((eresolve_tac prems) ORELSE' atac))) - ]); - -(* --------------------------- automatic prover --------------------------- *) -(* Set up a clasimpset that contains data-level simplifications. *) - -val MI_css = temp_css addsimps2 (CP_simps @ histState.simps - @ [slice_def,equal_false_not,if_cancel,sum_case_Inl, sum_case_Inr]); +val MI_css = (claset(), simpset()); (* A more aggressive variant that tries to solve subgoals by assumption or contradiction during the simplification. THIS IS UNSAFE, BECAUSE IT DOESN'T RECORD THE CHOICES!! - (but sometimes a lot faster than MI_css) + (but it can be a lot faster than MI_css) *) val MI_fast_css = let val (cs,ss) = MI_css in - (cs, ss addSSolver (fn thms => assume_tac ORELSE' (etac notE))) + (cs addSEs [squareE], ss addSSolver (fn thms => assume_tac ORELSE' (etac notE))) end; -(* Make sure the simpset accepts non-boolean simplifications *) -simpset_ref() := let val (_,ss) = MI_css in ss end; - +val temp_elim = make_elim o temp_use; (****************************** The history variable ******************************) section "History variable"; qed_goal "HistoryLemma" MemoryImplementation.thy - "Init(RALL p. $(ImpInit p)) .& [](RALL p. ImpNext p) \ -\ .-> (EEX rmhist. Init(RALL p. $(HInit rmhist p)) \ -\ .& [](RALL p. [HNext rmhist p]_))" - (fn _ => [Auto_tac, - rtac historyI 1, TRYALL atac, + "|- Init(!p. ImpInit p) & [](!p. ImpNext p) \ +\ --> (EEX rmhist. Init(! p. HInit rmhist p) \ +\ & [](!p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)))" + (fn _ => [Clarsimp_tac 1, + rtac historyI 1, TRYALL atac, rtac MI_base 1, action_simp_tac (simpset() addsimps [HInit_def]) [] [] 1, - res_inst_tac [("x","p")] fun_cong 1, atac 1, + etac fun_cong 1, action_simp_tac (simpset() addsimps [HNext_def]) [busy_squareI] [] 1, - res_inst_tac [("x","p")] fun_cong 1, atac 1 + etac fun_cong 1 ]); qed_goal "History" MemoryImplementation.thy - "Implementation .-> (EEX rmhist. Hist rmhist)" - (fn _ => [Auto_tac, - rtac ((temp_mp HistoryLemma) RS eex_mono) 1, - SELECT_GOAL - (auto_tac (MI_css - addsimps2 [Impl_def,MClkISpec_def,RPCISpec_def,IRSpec_def, - MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, - ImpInit_def,Init_def,ImpNext_def, - c_def,r_def,m_def,all_box,split_box_conj])) 1, - auto_tac (MI_css - addsimps2 [Hist_def,HistP_def,Init_def,all_box,split_box_conj]) + "|- Implementation --> (EEX rmhist. Hist rmhist)" + (fn _ => [Clarsimp_tac 1, + rtac ((temp_use HistoryLemma) RS eex_mono) 1, + force_tac (MI_css + addsimps2 [Hist_def,HistP_def,Init_def,all_box,split_box_conj]) 3, + auto_tac (MI_css + addsimps2 [Implementation_def,MClkISpec_def,RPCISpec_def,IRSpec_def, + MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, + ImpInit_def,Init_def,ImpNext_def, + c_def,r_def,m_def,all_box,split_box_conj]) ]); (******************************** The safety part *********************************) @@ -108,11 +68,12 @@ section "Correctness of predicate-action diagram"; + (* ========== Step 1.1 ================================================= *) (* The implementation's initial condition implies the state predicate S1 *) qed_goal "Step1_1" MemoryImplementation.thy - "$(ImpInit p) .& $(HInit rmhist p) .-> $(S1 rmhist p)" + "|- ImpInit p & HInit rmhist p --> S1 rmhist p" (fn _ => [auto_tac (MI_fast_css addsimps2 [MVNROKBA_def,MClkInit_def,RPCInit_def,PInit_def, HInit_def,ImpInit_def,S_def,S1_def]) @@ -122,84 +83,81 @@ (* Figure 16 is a predicate-action diagram for the implementation. *) qed_goal "Step1_2_1" MemoryImplementation.thy - "[HNext rmhist p]_ .& ImpNext p \ -\ .& .~ unchanged .& $(S1 rmhist p) \ -\ .-> (S2 rmhist p)$ .& (ENext p) .& unchanged " - (fn _ => [auto_tac (MI_css addsimps2 [ImpNext_def] - addSEs2 [S1ClerkUnchE,S1RPCUnchE,S1MemUnchE,S1HistE]), - ALLGOALS (action_simp_tac (simpset() addsimps [square_def]) [] [S1EnvE]) + "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p \ +\ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S1 rmhist p \ +\ --> (S2 rmhist p)$ & ENext p & unchanged (c p, r p, m p)" + (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) [] + (map temp_elim [S1ClerkUnch,S1RPCUnch,S1MemUnch,S1Hist]) 1, + auto_tac (MI_fast_css addSIs2 [S1Env]) ]); qed_goal "Step1_2_2" MemoryImplementation.thy - "[HNext rmhist p]_ .& ImpNext p \ -\ .& .~ unchanged .& $(S2 rmhist p) \ -\ .-> (S3 rmhist p)$ .& (MClkFwd memCh crCh cst p) .& unchanged " - (fn _ => [auto_tac (MI_css addsimps2 [ImpNext_def] - addSEs2 [S2EnvUnchE,S2RPCUnchE,S2MemUnchE,S2HistE]), - ALLGOALS (action_simp_tac (simpset() addsimps [square_def]) [] [S2ClerkE,S2ForwardE]) + "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p \ +\ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S2 rmhist p \ +\ --> (S3 rmhist p)$ & MClkFwd memCh crCh cst p & unchanged (e p, r p, m p, rmhist!p)" + (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) [] + (map temp_elim [S2EnvUnch,S2RPCUnch,S2MemUnch,S2Hist]) 1, + auto_tac (MI_fast_css addSIs2 [S2Clerk,S2Forward]) ]); qed_goal "Step1_2_3" MemoryImplementation.thy - "[HNext rmhist p]_ .& ImpNext p \ -\ .& .~ unchanged .& $(S3 rmhist p) \ -\ .-> ((S4 rmhist p)$ .& RPCFwd crCh rmCh rst p .& unchanged ) \ -\ .| ((S6 rmhist p)$ .& RPCFail crCh rmCh rst p .& unchanged )" - (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) - [] [S3EnvUnchE,S3ClerkUnchE,S3MemUnchE] 1, - ALLGOALS (action_simp_tac (simpset() addsimps [square_def]) - [] [S3RPCE,S3ForwardE,S3FailE]), - auto_tac (MI_css addEs2 [S3HistE]) + "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p \ +\ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S3 rmhist p \ +\ --> ((S4 rmhist p)$ & RPCFwd crCh rmCh rst p & unchanged (e p, c p, m p, rmhist!p)) \ +\ | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))" + (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) [] + (map temp_elim [S3EnvUnch,S3ClerkUnch,S3MemUnch]) 1, + action_simp_tac (simpset()) [] + (squareE::map temp_elim [S3RPC,S3Forward,S3Fail]) 1, + auto_tac (MI_css addDs2 [S3Hist]) ]); qed_goal "Step1_2_4" MemoryImplementation.thy - "[HNext rmhist p]_ .& ImpNext p \ -\ .& .~ unchanged \ -\ .& $(S4 rmhist p) .& (RALL l. $(MemInv mem l)) \ -\ .-> ((S4 rmhist p)$ .& Read rmCh mem ires p .& unchanged ) \ -\ .| ((S4 rmhist p)$ .& (REX l. Write rmCh mem ires p l) .& unchanged ) \ -\ .| ((S5 rmhist p)$ .& MemReturn rmCh ires p .& unchanged )" - (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) - [] [S4EnvUnchE,S4ClerkUnchE,S4RPCUnchE] 1, - ALLGOALS (action_simp_tac (simpset() addsimps [square_def,RNext_def]) - [] [S4ReadE,S4WriteE,S4ReturnE]), - auto_tac (MI_css addEs2 [S4HistE]) + "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p \ +\ & ~unchanged (e p, c p, r p, m p, rmhist!p) \ +\ & $S4 rmhist p & (!l. $(MemInv mm l)) \ +\ --> ((S4 rmhist p)$ & Read rmCh mm ires p & unchanged (e p, c p, r p, rmhist!p)) \ +\ | ((S4 rmhist p)$ & (? l. Write rmCh mm ires p l) & unchanged (e p, c p, r p, rmhist!p)) \ +\ | ((S5 rmhist p)$ & MemReturn rmCh ires p & unchanged (e p, c p, r p))" + (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) [] + (map temp_elim [S4EnvUnch,S4ClerkUnch,S4RPCUnch]) 1, + action_simp_tac (simpset() addsimps [RNext_def]) [] + (squareE::map temp_elim [S4Read,S4Write,S4Return]) 1, + auto_tac (MI_css addDs2 [S4Hist]) ]); qed_goal "Step1_2_5" MemoryImplementation.thy - "[HNext rmhist p]_ .& ImpNext p \ -\ .& .~ unchanged .& $(S5 rmhist p) \ -\ .-> ((S6 rmhist p)$ .& RPCReply crCh rmCh rst p .& unchanged ) \ -\ .| ((S6 rmhist p)$ .& RPCFail crCh rmCh rst p .& unchanged )" - (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) - [] [S5EnvUnchE,S5ClerkUnchE,S5MemUnchE,S5HistE] 1, - action_simp_tac (simpset() addsimps [square_def]) [] [S5RPCE] 1, - auto_tac (MI_fast_css addSEs2 [S5ReplyE,S5FailE]) + "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p \ +\ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S5 rmhist p \ +\ --> ((S6 rmhist p)$ & RPCReply crCh rmCh rst p & unchanged (e p, c p, m p)) \ +\ | ((S6 rmhist p)$ & RPCFail crCh rmCh rst p & unchanged (e p, c p, m p))" + (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) [] + (map temp_elim [S5EnvUnch,S5ClerkUnch,S5MemUnch,S5Hist]) 1, + action_simp_tac (simpset()) [] [squareE, temp_elim S5RPC] 1, + auto_tac (MI_fast_css addSDs2 [S5Reply,S5Fail]) ]); qed_goal "Step1_2_6" MemoryImplementation.thy - "[HNext rmhist p]_ .& ImpNext p \ -\ .& .~ unchanged .& $(S6 rmhist p) \ -\ .-> ((S1 rmhist p)$ .& (MClkReply memCh crCh cst p) .& unchanged )\ -\ .| ((S3 rmhist p)$ .& (MClkRetry memCh crCh cst p) .& unchanged )" - (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) - [] [S6EnvUnchE,S6RPCUnchE,S6MemUnchE] 1, - ALLGOALS (action_simp_tac (simpset() addsimps [square_def]) - [] [S6ClerkE,S6RetryE,S6ReplyE]), - auto_tac (MI_css addEs2 [S6HistE]) + "|- [HNext rmhist p]_(c p,r p,m p, rmhist!p) & ImpNext p \ +\ & ~unchanged (e p, c p, r p, m p, rmhist!p) & $S6 rmhist p \ +\ --> ((S1 rmhist p)$ & MClkReply memCh crCh cst p & unchanged (e p, r p, m p))\ +\ | ((S3 rmhist p)$ & MClkRetry memCh crCh cst p & unchanged (e p,r p,m p,rmhist!p))" + (fn _ => [action_simp_tac (simpset() addsimps [ImpNext_def]) [] + (map temp_elim [S6EnvUnch,S6RPCUnch,S6MemUnch]) 1, + action_simp_tac (simpset()) [] + (squareE::map temp_elim [S6Clerk,S6Retry,S6Reply]) 1, + auto_tac (MI_css addDs2 [S6Hist]) ]); - (* -------------------------------------------------------------------------- Step 1.3: S1 implies the barred initial condition. *) section "Initialization (Step 1.3)"; -val resbar_unl = rewrite_rule [slice_def] (action_unlift resbar_def); - qed_goal "Step1_3" MemoryImplementation.thy - "$(S1 rmhist p) .-> $(PInit (resbar rmhist) p)" - (fn _ => [action_simp_tac (simpset() addsimps [resbar_unl,PInit_def,S_def,S1_def]) + "|- S1 rmhist p --> PInit (resbar rmhist) p" + (fn _ => [action_simp_tac (simpset() addsimps [resbar_def,PInit_def,S_def,S1_def]) [] [] 1 ]); @@ -211,52 +169,51 @@ section "Step simulation (Step 1.4)"; qed_goal "Step1_4_1" MemoryImplementation.thy - "ENext p .& $(S1 rmhist p) .& (S2 rmhist p)$ .& unchanged \ -\ .-> unchanged " - (fn _ => [ auto_tac (MI_fast_css addsimps2 [c_def,r_def,m_def,resbar_unl]) ]); + "|- ENext p & $S1 rmhist p & (S2 rmhist p)$ & unchanged (c p, r p, m p) \ +\ --> unchanged (rtrner memCh!p, resbar rmhist!p)" + (fn _ => [ auto_tac (MI_fast_css addsimps2 [c_def,r_def,m_def,resbar_def]) ]); qed_goal "Step1_4_2" MemoryImplementation.thy - "MClkFwd memCh crCh cst p .& $(S2 rmhist p) .& (S3 rmhist p)$ \ -\ .& unchanged \ -\ .-> unchanged " - (fn _ => [auto_tac (MI_fast_css - addsimps2 [MClkFwd_def, e_def, r_def, m_def, resbar_unl, - S_def, S2_def, S3_def]) + "|- MClkFwd memCh crCh cst p & $S2 rmhist p & (S3 rmhist p)$ \ +\ & unchanged (e p, r p, m p, rmhist!p) \ +\ --> unchanged (rtrner memCh!p, resbar rmhist!p)" + (fn _ => [action_simp_tac + (simpset() addsimps [MClkFwd_def, e_def, r_def, m_def, resbar_def, + S_def, S2_def, S3_def]) [] [] 1 ]); qed_goal "Step1_4_3a" MemoryImplementation.thy - "RPCFwd crCh rmCh rst p .& $(S3 rmhist p) .& (S4 rmhist p)$ \ -\ .& unchanged \ -\ .-> unchanged " - (fn _ => [auto_tac (MI_fast_css addsimps2 [e_def,c_def,m_def,resbar_unl]), - (* NB: Adding S3_exclE,S4_exclE as safe elims above would loop, - adding them as unsafe elims doesn't help, - because auto_tac doesn't find the proof! *) - REPEAT (eresolve_tac [S3_exclE,S4_exclE] 1), - action_simp_tac (simpset() addsimps [S_def, S3_def]) [] [] 1 + "|- RPCFwd crCh rmCh rst p & $S3 rmhist p & (S4 rmhist p)$ \ +\ & unchanged (e p, c p, m p, rmhist!p) \ +\ --> unchanged (rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, + REPEAT (dresolve_tac (map temp_use [S3_excl,S4_excl]) 1), + action_simp_tac + (simpset() addsimps [e_def,c_def,m_def,resbar_def,S_def, S3_def]) [] [] 1 ]); qed_goal "Step1_4_3b" MemoryImplementation.thy - "RPCFail crCh rmCh rst p .& $(S3 rmhist p) .& (S6 rmhist p)$ .& unchanged \ -\ .-> MemFail memCh (resbar rmhist) p" - (fn _ => [auto_tac (MI_css addsimps2 [RPCFail_def,MemFail_def,e_def,c_def,m_def, - resbar_unl]), - (* It's faster not to expand S3 at once *) - action_simp_tac (simpset() addsimps [S3_def,S_def]) [] [] 1, - etac S6_exclE 1, - auto_tac (MI_fast_css addsimps2 [Return_def]) + "|- RPCFail crCh rmCh rst p & $S3 rmhist p & (S6 rmhist p)$ & unchanged (e p, c p, m p) \ +\ --> MemFail memCh (resbar rmhist) p" + (fn _ => [Clarsimp_tac 1, + dtac (temp_use S6_excl) 1, + auto_tac (MI_css addsimps2 [RPCFail_def,MemFail_def,e_def,c_def,m_def, + resbar_def]), + force_tac (MI_css addsimps2 [S3_def,S_def]) 1, + auto_tac (MI_css addsimps2 [Return_def]) ]); qed_goal "Step1_4_4a1" MemoryImplementation.thy - "$(S4 rmhist p) .& (S4 rmhist p)$ .& ReadInner rmCh mem ires p l \ -\ .& unchanged .& $(MemInv mem l) \ -\ .-> ReadInner memCh mem (resbar rmhist) p l" - (fn _ => [action_simp_tac + "|- $S4 rmhist p & (S4 rmhist p)$ & ReadInner rmCh mm ires p l \ +\ & unchanged (e p, c p, r p, rmhist!p) & $MemInv mm l \ +\ --> ReadInner memCh mm (resbar rmhist) p l" + (fn _ => [Clarsimp_tac 1, + REPEAT (dtac (temp_use S4_excl) 1), + action_simp_tac (simpset() addsimps [ReadInner_def,GoodRead_def,BadRead_def,e_def,c_def,m_def]) [] [] 1, - ALLGOALS (REPEAT o (etac S4_exclE)), - auto_tac (MI_css addsimps2 [resbar_unl]), + auto_tac (MI_css addsimps2 [resbar_def]), ALLGOALS (action_simp_tac (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def, S_def,S4_def,RdRequest_def,MemInv_def]) @@ -264,22 +221,22 @@ ]); qed_goal "Step1_4_4a" MemoryImplementation.thy - "Read rmCh mem ires p .& $(S4 rmhist p) .& (S4 rmhist p)$ \ -\ .& unchanged .& (RALL l. $(MemInv mem l)) \ -\ .-> Read memCh mem (resbar rmhist) p" - (fn _ => [ auto_tac (MI_css addsimps2 [Read_def] addSIs2 [action_mp Step1_4_4a1]) ]); + "|- Read rmCh mm ires p & $S4 rmhist p & (S4 rmhist p)$ \ +\ & unchanged (e p, c p, r p, rmhist!p) & (!l. $(MemInv mm l)) \ +\ --> Read memCh mm (resbar rmhist) p" + (fn _ => [ force_tac (MI_css addsimps2 [Read_def] addSEs2 [Step1_4_4a1]) 1 ]); qed_goal "Step1_4_4b1" MemoryImplementation.thy - "$(S4 rmhist p) .& (S4 rmhist p)$ .& WriteInner rmCh mem ires p l v \ -\ .& unchanged \ -\ .-> WriteInner memCh mem (resbar rmhist) p l v" - (fn _ => [action_simp_tac + "|- $S4 rmhist p & (S4 rmhist p)$ & WriteInner rmCh mm ires p l v \ +\ & unchanged (e p, c p, r p, rmhist!p) \ +\ --> WriteInner memCh mm (resbar rmhist) p l v" + (fn _ => [Clarsimp_tac 1, + REPEAT (dtac (temp_use S4_excl) 1), + action_simp_tac (simpset() addsimps [WriteInner_def, GoodWrite_def, BadWrite_def, e_def, c_def, m_def]) [] [] 1, - ALLGOALS (REPEAT o (etac S4_exclE)), - auto_tac (MI_css addsimps2 [resbar_unl]), - (* it's faster not to merge the two simplifications *) + auto_tac (MI_css addsimps2 [resbar_def]), ALLGOALS (action_simp_tac (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def, S_def,S4_def,WrRequest_def]) @@ -287,99 +244,93 @@ ]); qed_goal "Step1_4_4b" MemoryImplementation.thy - "Write rmCh mem ires p l .& $(S4 rmhist p) .& (S4 rmhist p)$ \ -\ .& unchanged \ -\ .-> Write memCh mem (resbar rmhist) p l" - (fn _ => [ auto_tac (MI_css addsimps2 [Write_def] addSIs2 [action_mp Step1_4_4b1]) ]); + "|- Write rmCh mm ires p l & $S4 rmhist p & (S4 rmhist p)$ \ +\ & unchanged (e p, c p, r p, rmhist!p) \ +\ --> Write memCh mm (resbar rmhist) p l" + (fn _ => [ force_tac (MI_css addsimps2 [Write_def] addSEs2 [Step1_4_4b1]) 1 ]); qed_goal "Step1_4_4c" MemoryImplementation.thy - "MemReturn rmCh ires p .& $(S4 rmhist p) .& (S5 rmhist p)$ .& unchanged \ -\ .-> unchanged " + "|- MemReturn rmCh ires p & $S4 rmhist p & (S5 rmhist p)$ & unchanged (e p, c p, r p) \ +\ --> unchanged (rtrner memCh!p, resbar rmhist!p)" (fn _ => [action_simp_tac - (simpset() addsimps [e_def,c_def,r_def,resbar_unl]) [] [] 1, - REPEAT (eresolve_tac [S4_exclE,S5_exclE] 1), + (simpset() addsimps [e_def,c_def,r_def,resbar_def]) [] [] 1, + REPEAT (dresolve_tac [temp_use S4_excl, temp_use S5_excl] 1), auto_tac (MI_fast_css addsimps2 [MemReturn_def,Return_def]) ]); qed_goal "Step1_4_5a" MemoryImplementation.thy - "RPCReply crCh rmCh rst p .& $(S5 rmhist p) .& (S6 rmhist p)$ .& unchanged \ -\ .-> unchanged " - (fn _ => [auto_tac (MI_css addsimps2 [e_def,c_def,m_def, resbar_unl]), - REPEAT (eresolve_tac [S5_exclE,S6_exclE] 1), - auto_tac (MI_css addsimps2 [RPCReply_def,Return_def,S5_def,S_def] - addSEs2 [MVOKBAnotRFE]) + "|- RPCReply crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$ & unchanged (e p, c p, m p) \ +\ --> unchanged (rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, + REPEAT (dresolve_tac [temp_use S5_excl, temp_use S6_excl] 1), + auto_tac (MI_css addsimps2 [e_def,c_def,m_def, resbar_def]), + auto_tac (MI_css addsimps2 [RPCReply_def,Return_def,S5_def,S_def] + addSDs2 [MVOKBAnotRF]) ]); qed_goal "Step1_4_5b" MemoryImplementation.thy - "RPCFail crCh rmCh rst p .& $(S5 rmhist p) .& (S6 rmhist p)$ .& unchanged \ -\ .-> MemFail memCh (resbar rmhist) p" - (fn _ => [action_simp_tac - (simpset() addsimps [e_def, c_def, m_def, RPCFail_def, Return_def, - MemFail_def, resbar_unl]) - [] [] 1, - action_simp_tac (simpset() addsimps [S5_def,S_def]) [] [] 1, - etac S6_exclE 1, - auto_tac MI_css + "|- RPCFail crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$ & unchanged (e p, c p, m p) \ +\ --> MemFail memCh (resbar rmhist) p" + (fn _ => [Clarsimp_tac 1, + dtac (temp_use S6_excl) 1, + auto_tac (MI_css addsimps2 [e_def, c_def, m_def, RPCFail_def, Return_def, + MemFail_def, resbar_def]), + auto_tac (MI_css addsimps2 [S5_def,S_def]) ]); qed_goal "Step1_4_6a" MemoryImplementation.thy - "MClkReply memCh crCh cst p .& $(S6 rmhist p) .& (S1 rmhist p)$ .& unchanged \ -\ .-> MemReturn memCh (resbar rmhist) p" - (fn _ => [action_simp_tac + "|- MClkReply memCh crCh cst p & $S6 rmhist p & (S1 rmhist p)$ & unchanged (e p, r p, m p) \ +\ --> MemReturn memCh (resbar rmhist) p" + (fn _ => [Clarsimp_tac 1, + dtac (temp_use S6_excl) 1, + action_simp_tac (simpset() addsimps [e_def, r_def, m_def, MClkReply_def, MemReturn_def, - Return_def, resbar_unl]) - [] [] 1, - ALLGOALS (etac S6_exclE), + Return_def, resbar_def]) [] [] 1, ALLGOALS Asm_full_simp_tac, (* simplify if-then-else *) ALLGOALS (action_simp_tac (simpset() addsimps [MClkReplyVal_def,S6_def,S_def]) - [] []), - rtac ifI 1, - ALLGOALS (action_simp_tac (simpset()) [] [MVOKBARFnotNRE]) + [] [MVOKBARFnotNR]) ]); qed_goal "Step1_4_6b" MemoryImplementation.thy - "MClkRetry memCh crCh cst p .& $(S6 rmhist p) .& (S3 rmhist p)$ \ -\ .& unchanged \ -\ .-> MemFail memCh (resbar rmhist) p" - (fn _ => [action_simp_tac - (simpset() addsimps [e_def, r_def, m_def, MClkRetry_def, MemFail_def, resbar_unl]) + "|- MClkRetry memCh crCh cst p & $S6 rmhist p & (S3 rmhist p)$ \ +\ & unchanged (e p, r p, m p, rmhist!p) \ +\ --> MemFail memCh (resbar rmhist) p" + (fn _ => [Clarsimp_tac 1, + dtac (temp_use S3_excl) 1, + action_simp_tac + (simpset() addsimps [e_def, r_def, m_def, MClkRetry_def, MemFail_def, resbar_def]) [] [] 1, - SELECT_GOAL (auto_tac (MI_css addsimps2 [S6_def,S_def])) 1, - etac S3_exclE 1, - Asm_full_simp_tac 1, - action_simp_tac (simpset() addsimps [S6_def,S3_def,S_def]) [] [] 1 + auto_tac (MI_css addsimps2 [S6_def,S_def]) ]); qed_goal "S_lemma" MemoryImplementation.thy - "unchanged \ -\ .-> unchanged (S rmhist ec cc rc cs rs hs1 hs2 p)" + "|- unchanged (e p, c p, r p, m p, rmhist!p) \ +\ --> unchanged (S rmhist ec cc rc cs rs hs1 hs2 p)" (fn _ => [auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,caller_def,rtrner_def, S_def,Calling_def]) ]); qed_goal "Step1_4_7H" MemoryImplementation.thy - "unchanged \ -\ .-> unchanged " - (fn _ => [Action_simp_tac 1, - SELECT_GOAL (auto_tac (MI_fast_css addsimps2 [c_def])) 1, - ALLGOALS (simp_tac (simpset() - addsimps [S1_def,S2_def,S3_def,S4_def,S5_def,S6_def])), - auto_tac (MI_css addSIs2 [action_mp S_lemma]) + "|- unchanged (e p, c p, r p, m p, rmhist!p) \ +\ --> unchanged (rtrner memCh!p, S1 rmhist p, S2 rmhist p, S3 rmhist p, \ +\ S4 rmhist p, S5 rmhist p, S6 rmhist p)" + (fn _ => [Clarsimp_tac 1, + rtac conjI 1, + force_tac (MI_css addsimps2 [c_def]) 1, + force_tac (MI_css addsimps2 [S1_def,S2_def,S3_def,S4_def,S5_def,S6_def] + addSIs2 [S_lemma]) 1 ]); -(* unlifted version as elimination rule *) -bind_thm("Step1_4_7h", - (rewrite_rule action_rews (Step1_4_7H RS actionD)) RS impdupE); - qed_goal "Step1_4_7" MemoryImplementation.thy - "unchanged .-> unchanged " + "|- unchanged (e p, c p, r p, m p, rmhist!p) \ +\ --> unchanged (rtrner memCh!p, resbar rmhist!p, S1 rmhist p, S2 rmhist p, S3 rmhist p, \ +\ S4 rmhist p, S5 rmhist p, S6 rmhist p)" (fn _ => [rtac actionI 1, rewrite_goals_tac action_rews, rtac impI 1, - etac Step1_4_7h 1, - auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,rtrner_def,resbar_unl]) + forward_tac [temp_use Step1_4_7H] 1, + auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,rtrner_def,resbar_def]) ]); @@ -387,10 +338,10 @@ steps of the implementation, and try to solve the idling case by simplification *) fun split_idle_tac simps i = - EVERY [rtac actionI i, - case_tac "(unchanged ) [[s,t]]" i, + EVERY [TRY (rtac actionI i), + case_tac "(s,t) |= unchanged (e p, c p, r p, m p, rmhist!p)" i, rewrite_goals_tac action_rews, - etac Step1_4_7h i, + forward_tac [temp_use Step1_4_7] i, asm_full_simp_tac (simpset() addsimps simps) i ]; @@ -402,75 +353,72 @@ (* Steps that leave all variables unchanged are safe, so I may assume that some variable changes in the proof that a step is safe. *) qed_goal "unchanged_safe" MemoryImplementation.thy - "(.~ (unchanged ) \ -\ .-> [UNext memCh mem (resbar rmhist) p]_) \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [rtac actionI 1, - case_tac "(unchanged ) [[s,t]]" 1, - rewrite_goals_tac action_rews, - auto_tac (MI_css addsimps2 [square_def] addSEs2 [action_impE Step1_4_7]) + "|- (~unchanged (e p, c p, r p, m p, rmhist!p) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [split_idle_tac [square_def] 1, + Force_tac 1 ]); (* turn into (unsafe, looping!) introduction rule *) -bind_thm("unchanged_safeI", impI RS (action_mp unchanged_safe)); +bind_thm("unchanged_safeI", impI RS (action_use unchanged_safe)); qed_goal "S1safe" MemoryImplementation.thy - "$(S1 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [Action_simp_tac 1, + "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, rtac unchanged_safeI 1, rtac idle_squareI 1, - auto_tac (MI_css addSEs2 (map action_conjimpE [Step1_2_1,Step1_4_1])) + auto_tac (MI_css addSDs2 [Step1_2_1,Step1_4_1]) ]); qed_goal "S2safe" MemoryImplementation.thy - "$(S2 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [Action_simp_tac 1, + "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, rtac unchanged_safeI 1, rtac idle_squareI 1, - auto_tac (MI_fast_css addSEs2 (map action_conjimpE [Step1_2_2,Step1_4_2])) + auto_tac (MI_css addSDs2 [Step1_2_2,Step1_4_2]) ]); qed_goal "S3safe" MemoryImplementation.thy - "$(S3 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [Action_simp_tac 1, + "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, rtac unchanged_safeI 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_3]), - auto_tac (MI_fast_css addsimps2 [square_def,UNext_def] - addSEs2 (map action_conjimpE [Step1_4_3a,Step1_4_3b])) + auto_tac (MI_css addSDs2 [Step1_2_3]), + auto_tac (MI_css addsimps2 [square_def,UNext_def] + addSDs2 [Step1_4_3a,Step1_4_3b]) ]); qed_goal "S4safe" MemoryImplementation.thy - "$(S4 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .& (RALL l. $(MemInv mem l)) \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [Action_simp_tac 1, + "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ & (!l. $(MemInv mm l)) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, rtac unchanged_safeI 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_4]), - ALLGOALS (action_simp_tac (simpset() addsimps [square_def,UNext_def,RNext_def]) [] []), - auto_tac (MI_fast_css addSEs2 (map action_conjimpE - [Step1_4_4a,Step1_4_4b,Step1_4_4c])) + auto_tac (MI_css addSDs2 [Step1_2_4]), + auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def] + addSDs2 [Step1_4_4a,Step1_4_4b,Step1_4_4c]) ]); qed_goal "S5safe" MemoryImplementation.thy - "$(S5 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [Action_simp_tac 1, + "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, rtac unchanged_safeI 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_5]), - auto_tac (MI_fast_css addsimps2 [square_def,UNext_def] - addSEs2 (map action_conjimpE [Step1_4_5a,Step1_4_5b])) + auto_tac (MI_css addSDs2 [Step1_2_5]), + auto_tac (MI_css addsimps2 [square_def,UNext_def] + addSDs2 [Step1_4_5a,Step1_4_5b]) ]); qed_goal "S6safe" MemoryImplementation.thy - "$(S6 rmhist p) .& ImpNext p .& [HNext rmhist p]_ \ -\ .-> [UNext memCh mem (resbar rmhist) p]_" - (fn _ => [Action_simp_tac 1, + "|- $S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ +\ --> [UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [Clarsimp_tac 1, rtac unchanged_safeI 1, - auto_tac (MI_css addSEs2 [action_conjimpE Step1_2_6]), - auto_tac (MI_fast_css addsimps2 [square_def,UNext_def,RNext_def] - addSEs2 (map action_conjimpE [Step1_4_6a,Step1_4_6b])) + auto_tac (MI_css addSDs2 [Step1_2_6]), + auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def] + addSDs2 [Step1_4_6a,Step1_4_6b]) ]); (* ---------------------------------------------------------------------- @@ -488,113 +436,99 @@ b. "implementation invariant": always in states S1,...,S6 *) qed_goal "Step1_5_1a" MemoryImplementation.thy - "IPImp p .-> (RALL l. []$(MemInv mem l))" - (fn _ => [auto_tac (MI_css addsimps2 [IPImp_def] - addSIs2 [temp_mp MemoryInvariantAll]) + "|- IPImp p --> (!l. []$MemInv mm l)" + (fn _ => [auto_tac (MI_css addsimps2 [IPImp_def,box_stp_act] + addSIs2 [MemoryInvariantAll]) ]); -bind_thm("MemInvI", (rewrite_rule intensional_rews (Step1_5_1a RS tempD)) RS impdupE); qed_goal "Step1_5_1b" MemoryImplementation.thy - " Init($(ImpInit p) .& $(HInit rmhist p)) .& [](ImpNext p) \ -\ .& [][HNext rmhist p]_ .& [](RALL l. $(MemInv mem l)) \ -\ .-> []($(ImpInv rmhist p))" + "|- Init(ImpInit p & HInit rmhist p) & [](ImpNext p) \ +\ & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) & [](!l. $MemInv mm l) \ +\ --> []ImpInv rmhist p" (fn _ => [inv_tac MI_css 1, - auto_tac (MI_css - addsimps2 [Init_def, ImpInv_def] - addSEs2 [action_impE Step1_1] - addEs2 (map action_conjimpE - [S1_successors,S2_successors,S3_successors, - S4_successors,S5_successors,S6_successors])) + auto_tac (MI_css addsimps2 [Init_def, ImpInv_def, box_stp_act] + addSDs2 [Step1_1] + addDs2 [S1_successors,S2_successors,S3_successors, + S4_successors,S5_successors,S6_successors]) ]); -bind_thm("ImpInvI", (rewrite_rule intensional_rews (Step1_5_1b RS tempD)) RS impdupE); (*** Initialization ***) qed_goal "Step1_5_2a" MemoryImplementation.thy - "Init($(ImpInit p) .& $(HInit rmhist p)) .-> Init($PInit (resbar rmhist) p)" + "|- Init(ImpInit p & HInit rmhist p) --> Init(PInit (resbar rmhist) p)" (fn _ => [auto_tac (MI_css addsimps2 [Init_def] - addSIs2 (map action_mp [Step1_1,Step1_3])) + addSIs2 [Step1_1,Step1_3]) ]); (*** step simulation ***) qed_goal "Step1_5_2b" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_ \ -\ .& $(ImpInv rmhist p) .& (RALL l. $(MemInv mem l))) \ -\ .-> [][UNext memCh mem (resbar rmhist) p]_" - (fn _ => [auto_tac (MI_fast_css - addsimps2 [ImpInv_def] - addSEs2 (STL4E::(map action_conjimpE - [S1safe,S2safe,S3safe,S4safe,S5safe,S6safe]))) + "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p) \ +\ & $ImpInv rmhist p & (!l. $MemInv mm l)) \ +\ --> [][UNext memCh mm (resbar rmhist) p]_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [auto_tac (MI_css + addsimps2 [ImpInv_def] addSEs2 [STL4E] + addSDs2 [S1safe,S2safe,S3safe,S4safe,S5safe,S6safe]) ]); (*** Liveness ***) qed_goal "GoodImpl" MemoryImplementation.thy - "IPImp p .& HistP rmhist p \ -\ .-> Init($(ImpInit p) .& $(HInit rmhist p)) \ -\ .& [](ImpNext p .& [HNext rmhist p]_) \ -\ .& [](RALL l. $(MemInv mem l)) .& []($(ImpInv rmhist p)) \ -\ .& ImpLive p" - (fn _ => [(* need some subgoals to prove [](ImpInv p), avoid duplication *) - rtac tempI 1, rewrite_goals_tac intensional_rews, rtac impI 1, - subgoal_tac - "sigma |= Init($(ImpInit p) .& $(HInit rmhist p)) \ -\ .& [](ImpNext p) \ -\ .& [][HNext rmhist p]_ \ -\ .& [](RALL l. $(MemInv mem l))" 1, - auto_tac (MI_css addsimps2 [split_box_conj] - addSEs2 [temp_conjimpE Step1_5_1b]), - SELECT_GOAL - (auto_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, - ImpLive_def,c_def,r_def,m_def])) 1, - SELECT_GOAL - (auto_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, - HistP_def,Init_def,action_unlift ImpInit_def])) 1, - SELECT_GOAL - (auto_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, - ImpNext_def,c_def,r_def,m_def, - split_box_conj])) 1, - SELECT_GOAL (auto_tac (MI_css addsimps2 [HistP_def])) 1, - etac ((temp_mp Step1_5_1a) RS ((temp_unlift allT) RS iffD1)) 1 + "|- IPImp p & HistP rmhist p \ +\ --> Init(ImpInit p & HInit rmhist p) \ +\ & [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ +\ & [](!l. $MemInv mm l) & []($ImpInv rmhist p) \ +\ & ImpLive p" + (fn _ => [Clarsimp_tac 1, + subgoal_tac + "sigma |= Init(ImpInit p & HInit rmhist p) \ +\ & [](ImpNext p) \ +\ & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) \ +\ & [](!l. $MemInv mm l)" 1, + auto_tac (MI_css addsimps2 [split_box_conj,box_stp_act] addSDs2 [Step1_5_1b]), + force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, + ImpLive_def,c_def,r_def,m_def]) 1, + force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, + HistP_def,Init_def,ImpInit_def]) 1, + force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, + ImpNext_def,c_def,r_def,m_def,split_box_conj]) 1, + force_tac (MI_css addsimps2 [HistP_def]) 1, + force_tac (MI_css addsimps2 [temp_use allT] addSDs2 [Step1_5_1a]) 1 ]); -(* The implementation is infinitely often in state S1 *) +(* The implementation is infinitely often in state S1... *) qed_goal "Step1_5_3a" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_) \ -\ .& [](RALL l. $(MemInv mem l)) \ -\ .& []($(ImpInv rmhist p)) .& ImpLive p \ -\ .-> []<>($(S1 rmhist p))" - (fn _ => [auto_tac (MI_css addsimps2 [ImpLive_def]), + "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ +\ & [](!l. $MemInv mm l) \ +\ & []($ImpInv rmhist p) & ImpLive p \ +\ --> []<>S1 rmhist p" + (fn _ => [clarsimp_tac (MI_css addsimps2 [ImpLive_def]) 1, rtac S1Infinite 1, - SELECT_GOAL - (auto_tac (MI_css - addsimps2 [split_box_conj] - addSIs2 (NotS1LeadstoS6:: - map temp_mp [S2_live,S3_live,S4a_live,S4b_live,S5_live]))) 1, - auto_tac (MI_css addsimps2 [split_box_conj] addSIs2 [temp_mp S6_live]) + force_tac (MI_css + addsimps2 [split_box_conj,box_stp_act] + addSIs2 [NotS1LeadstoS6,S2_live,S3_live,S4a_live,S4b_live,S5_live]) 1, + auto_tac (MI_css addsimps2 [split_box_conj] addSIs2 [S6_live]) ]); -(* Hence, it satisfies the fairness requirements of the specification *) +(* ... which implies that it satisfies the fairness requirements of the specification *) qed_goal "Step1_5_3b" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_) \ -\ .& [](RALL l. $(MemInv mem l)) .& []($(ImpInv rmhist p)) .& ImpLive p \ -\ .-> WF(RNext memCh mem (resbar rmhist) p)_" - (fn _ => [ auto_tac (MI_fast_css addSIs2 [RNext_fair,temp_mp Step1_5_3a]) ]); + "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ +\ & [](!l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p \ +\ --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [ auto_tac (MI_css addSIs2 [RNext_fair,Step1_5_3a]) ]); qed_goal "Step1_5_3c" MemoryImplementation.thy - "[](ImpNext p .& [HNext rmhist p]_) \ -\ .& [](RALL l. $(MemInv mem l)) .& []($(ImpInv rmhist p)) .& ImpLive p \ -\ .-> WF(MemReturn memCh (resbar rmhist) p)_" - (fn _ => [ auto_tac (MI_fast_css addSIs2 [Return_fair,temp_mp Step1_5_3a]) ]); + "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ +\ & [](!l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p \ +\ --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + (fn _ => [ auto_tac (MI_css addSIs2 [Return_fair,Step1_5_3a]) ]); (* QED step of step 1 *) qed_goal "Step1" MemoryImplementation.thy - "IPImp p .& HistP rmhist p .-> UPSpec memCh mem (resbar rmhist) p" + "|- IPImp p & HistP rmhist p --> UPSpec memCh mm (resbar rmhist) p" (fn _ => [auto_tac (MI_css addsimps2 [UPSpec_def,split_box_conj] - addSEs2 [temp_impE GoodImpl] - addSIs2 (map temp_mp [Step1_5_2a,Step1_5_2b, - Step1_5_3b,Step1_5_3c])) + addSDs2 [GoodImpl] + addSIs2 [Step1_5_2a,Step1_5_2b,Step1_5_3b,Step1_5_3c]) ]); @@ -602,46 +536,46 @@ section "Step 2"; qed_goal "Step2_2a" MemoryImplementation.thy - "ImpNext p .& [HNext rmhist p]_ \ -\ .& $(S4 rmhist p) .& Write rmCh mem ires p l \ -\ .-> (S4 rmhist p)$ .& unchanged " - (fn _ => [split_idle_tac [] 1, - action_simp_tac (simpset() addsimps [ImpNext_def]) - [] [S4EnvUnchE,S4ClerkUnchE,S4RPCUnchE] 1, - TRYALL (action_simp_tac (simpset() addsimps [square_def]) [] [S4WriteE]), - Auto_tac + "|- Write rmCh mm ires p l & ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p) \ +\ & $ImpInv rmhist p \ +\ --> (S4 rmhist p)` & unchanged (e p, c p, r p, rmhist!p)" + (fn _ => [Clarsimp_tac 1, + dtac (action_use WriteS4) 1, atac 1, + split_idle_tac [] 1, + auto_tac (MI_css addsimps2 [ImpNext_def] + addSDs2 [S4EnvUnch,S4ClerkUnch,S4RPCUnch]), + auto_tac (MI_css addsimps2 [square_def] addDs2 [S4Write]) ]); qed_goal "Step2_2" MemoryImplementation.thy - " (RALL p. ImpNext p) \ -\ .& (RALL p. [HNext rmhist p]_) \ -\ .& (RALL p. $(ImpInv rmhist p)) \ -\ .& [REX q. Write rmCh mem ires q l]_(mem@l) \ -\ .-> [REX q. Write memCh mem (resbar rmhist) q l]_(mem@l)" - (fn _ => [auto_tac (MI_css addsimps2 [square_def] - addSIs2 [action_mp Step1_4_4b] - addSEs2 [WriteS4E, action_conjimpE Step2_2a]) + "|- (!p. ImpNext p) \ +\ & (!p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ +\ & (!p. $ImpInv rmhist p) \ +\ & [? q. Write rmCh mm ires q l]_(mm!l) \ +\ --> [? q. Write memCh mm (resbar rmhist) q l]_(mm!l)" + (fn _ => [auto_tac (MI_css addSIs2 [squareCI] addSEs2 [squareE]), + REPEAT (ares_tac [exI, action_use Step1_4_4b] 1), + force_tac (MI_css addSIs2 [WriteS4]) 1, + auto_tac (MI_css addSDs2 [Step2_2a]) ]); qed_goal "Step2_lemma" MemoryImplementation.thy - " []( (RALL p. ImpNext p) \ -\ .& (RALL p. [HNext rmhist p]_) \ -\ .& (RALL p. $(ImpInv rmhist p)) \ -\ .& [REX q. Write rmCh mem ires q l]_(mem@l)) \ -\ .-> [][REX q. Write memCh mem (resbar rmhist) q l]_(mem@l)" - (fn _ => [ auto_tac (MI_css addSEs2 [STL4E, action_conjimpE Step2_2]) ]); + "|- []( (!p. ImpNext p) \ +\ & (!p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ +\ & (!p. $ImpInv rmhist p) \ +\ & [? q. Write rmCh mm ires q l]_(mm!l)) \ +\ --> [][? q. Write memCh mm (resbar rmhist) q l]_(mm!l)" + (fn _ => [ force_tac (MI_css addSEs2 [STL4E] addSDs2 [Step2_2]) 1 ]); qed_goal "Step2" MemoryImplementation.thy - "#(MemLoc l) .& (RALL p. IPImp p .& HistP rmhist p) \ -\ .-> MSpec memCh mem (resbar rmhist) l" + "|- #l : #MemLoc & (!p. IPImp p & HistP rmhist p) \ +\ --> MSpec memCh mm (resbar rmhist) l" (fn _ => [auto_tac (MI_css addsimps2 [MSpec_def]), - (* prove initial condition, don't expand IPImp in other subgoal *) - SELECT_GOAL (auto_tac (MI_css addsimps2 [IPImp_def,MSpec_def])) 1, - auto_tac (MI_css addSIs2 [temp_mp Step2_lemma] + force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 1, + auto_tac (MI_css addSIs2 [Step2_lemma] addsimps2 [split_box_conj,all_box]), - SELECT_GOAL (auto_tac (MI_css addsimps2 [IPImp_def,MSpec_def])) 4, - auto_tac (MI_css addsimps2 [split_box_conj] - addSEs2 [temp_impE GoodImpl]) + force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 4, + auto_tac (MI_css addsimps2 [split_box_conj] addSEs2 [allE] addSDs2 [GoodImpl]) ]); (* ----------------------------- Main theorem --------------------------------- *) @@ -655,18 +589,20 @@ and history variable with explicit refinement mapping *) qed_goal "Impl_IUSpec" MemoryImplementation.thy - "Implementation .& Hist rmhist .-> IUSpec memCh mem (resbar rmhist)" - (fn _ => [auto_tac (MI_css addsimps2 [IUSpec_def,Impl_def,IPImp_def,MClkISpec_def, - RPCISpec_def,IRSpec_def,Hist_def] - addSIs2 (map temp_mp [Step1,Step2])) + "|- Implementation & Hist rmhist --> IUSpec memCh mm (resbar rmhist)" + (fn _ => [auto_tac (MI_css addsimps2 [IUSpec_def,Implementation_def,IPImp_def, + MClkISpec_def,RPCISpec_def,IRSpec_def,Hist_def] + addSIs2 [Step1,Step2]) ]); (* The main theorem: introduce hiding and eliminate history variable. *) qed_goal "Implementation" MemoryImplementation.thy - "Implementation .-> USpec memCh" - (fn _ => [Auto_tac, - forward_tac [temp_mp History] 1, + "|- Implementation --> USpec memCh" + (fn _ => [Clarsimp_tac 1, + forward_tac [temp_use History] 1, auto_tac (MI_css addsimps2 [USpec_def] - addIs2 (map temp_mp [eexI, Impl_IUSpec]) + addIs2 [eexI, Impl_IUSpec, MI_base] addSEs2 [eexE]) ]); + + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemoryImplementation.thy --- a/src/HOL/TLA/Memory/MemoryImplementation.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Mon Feb 08 13:02:56 1999 +0100 @@ -9,7 +9,9 @@ RPC-Memory example: Memory implementation *) -MemoryImplementation = Memory + RPC + MemClerk + MIParameters + +MemoryImplementation = Memory + RPC + MemClerk + Datatype + + +datatype histState = histA | histB types histType = "(PrIds => histState) stfun" (* the type of the history variable *) @@ -19,8 +21,7 @@ (* channel (external) *) memCh :: "memChType" (* internal variables *) - mem :: "memType" - resbar :: "histType => resType" (* defined by refinement mapping *) + mm :: "memType" (* the state variables of the implementation *) (* channels *) @@ -28,7 +29,7 @@ crCh :: "rpcSndChType" rmCh :: "rpcRcvChType" (* internal variables *) - (* identity refinement mapping for mem -- simply reused *) + (* identity refinement mapping for mm -- simply reused *) rst :: "rpcStType" cst :: "mClkStType" ires :: "resType" @@ -36,153 +37,145 @@ rmhist :: "histType" *) +constdefs + (* auxiliary predicates *) + MVOKBARF :: "Vals => bool" + "MVOKBARF v == (v : MemVal) | (v = OK) | (v = BadArg) | (v = RPCFailure)" + MVOKBA :: "Vals => bool" + "MVOKBA v == (v : MemVal) | (v = OK) | (v = BadArg)" + MVNROKBA :: "Vals => bool" + "MVNROKBA v == (v : MemVal) | (v = NotAResult) | (v = OK) | (v = BadArg)" + + (* tuples of state functions changed by the various components *) + e :: "PrIds => (bit * memOp) stfun" + "e p == PRED (caller memCh!p)" + c :: "PrIds => (mClkState * (bit * Vals) * (bit * rpcOp)) stfun" + "c p == PRED (cst!p, rtrner memCh!p, caller crCh!p)" + r :: "PrIds => (rpcState * (bit * Vals) * (bit * memOp)) stfun" + "r p == PRED (rst!p, rtrner crCh!p, caller rmCh!p)" + m :: "PrIds => ((bit * Vals) * Vals) stfun" + "m p == PRED (rtrner rmCh!p, ires!p)" + (* the environment action *) ENext :: "PrIds => action" + "ENext p == ACT (? l. #l : #MemLoc & Call memCh p #(read l))" + (* specification of the history variable *) HInit :: "histType => PrIds => stpred" + "HInit rmhist p == PRED rmhist!p = #histA" + HNext :: "histType => PrIds => action" + "HNext rmhist p == ACT (rmhist!p)$ = + (if (MemReturn rmCh ires p | RPCFail crCh rmCh rst p) + then #histB + else if (MClkReply memCh crCh cst p) + then #histA + else $(rmhist!p))" + HistP :: "histType => PrIds => temporal" + "HistP rmhist p == TEMP Init HInit rmhist p + & [][HNext rmhist p]_(c p,r p,m p, rmhist!p)" + Hist :: "histType => temporal" + "Hist rmhist == TEMP (!p. HistP rmhist p)" (* the implementation *) + IPImp :: "PrIds => temporal" + "IPImp p == TEMP ( Init ~Calling memCh p & [][ENext p]_(e p) + & MClkIPSpec memCh crCh cst p + & RPCIPSpec crCh rmCh rst p + & RPSpec rmCh mm ires p + & (! l. #l : #MemLoc --> MSpec rmCh mm ires l))" + ImpInit :: "PrIds => stpred" + "ImpInit p == PRED ( ~Calling memCh p + & MClkInit crCh cst p + & RPCInit rmCh rst p + & PInit ires p)" + ImpNext :: "PrIds => action" + "ImpNext p == ACT [ENext p]_(e p) + & [MClkNext memCh crCh cst p]_(c p) + & [RPCNext crCh rmCh rst p]_(r p) + & [RNext rmCh mm ires p]_(m p)" + ImpLive :: "PrIds => temporal" - IPImp :: "PrIds => temporal" + "ImpLive p == TEMP WF(MClkFwd memCh crCh cst p)_(c p) + & SF(MClkReply memCh crCh cst p)_(c p) + & WF(RPCNext crCh rmCh rst p)_(r p) + & WF(RNext rmCh mm ires p)_(m p) + & WF(MemReturn rmCh ires p)_(m p)" + Implementation :: "temporal" - ImpInv :: "histType => PrIds => stpred" - - (* tuples of state functions changed by the various components *) - e :: "PrIds => (bit * memArgType) stfun" - c :: "PrIds => (mClkState * (bit * Vals) * (bit * rpcArgType)) stfun" - r :: "PrIds => (rpcState * (bit * Vals) * (bit * memArgType)) stfun" - m :: "PrIds => ((bit * Vals) * Vals) stfun" + "Implementation == TEMP ( (!p. Init (~Calling memCh p) & [][ENext p]_(e p)) + & MClkISpec memCh crCh cst + & RPCISpec crCh rmCh rst + & IRSpec rmCh mm ires)" (* the predicate S describes the states of the implementation. - slight simplification: two "histState" parameters instead of a (one- or - two-element) set. *) - S :: "histType => bool => bool => bool => mClkState => rpcState => histState => histState => PrIds => stpred" + slight simplification: two "histState" parameters instead of a + (one- or two-element) set. + NB: The second conjunct of the definition in the paper is taken care of by + the type definitions. The last conjunct is asserted separately as the memory + invariant MemInv, proved in Memory.ML. *) + S :: "histType => bool => bool => bool => mClkState => rpcState => histState => histState => PrIds => stpred" + "S rmhist ecalling ccalling rcalling cs rs hs1 hs2 p == PRED + Calling memCh p = #ecalling + & Calling crCh p = #ccalling + & (#ccalling --> arg = MClkRelayArg>) + & (~ #ccalling & cst!p = #clkB --> MVOKBARF>) + & Calling rmCh p = #rcalling + & (#rcalling --> arg = RPCRelayArg>) + & (~ #rcalling --> ires!p = #NotAResult) + & (~ #rcalling & rst!p = #rpcB --> MVOKBA>) + & cst!p = #cs + & rst!p = #rs + & (rmhist!p = #hs1 | rmhist!p = #hs2) + & MVNROKBA" (* predicates S1 -- S6 define special instances of S *) S1 :: "histType => PrIds => stpred" + "S1 rmhist p == S rmhist False False False clkA rpcA histA histA p" S2 :: "histType => PrIds => stpred" + "S2 rmhist p == S rmhist True False False clkA rpcA histA histA p" S3 :: "histType => PrIds => stpred" + "S3 rmhist p == S rmhist True True False clkB rpcA histA histB p" S4 :: "histType => PrIds => stpred" + "S4 rmhist p == S rmhist True True True clkB rpcB histA histB p" S5 :: "histType => PrIds => stpred" + "S5 rmhist p == S rmhist True True False clkB rpcB histB histB p" S6 :: "histType => PrIds => stpred" + "S6 rmhist p == S rmhist True False False clkB rpcA histB histB p" - (* auxiliary predicates *) - MVOKBARF :: "Vals => bool" - MVOKBA :: "Vals => bool" - MVNROKBA :: "Vals => bool" + (* The invariant asserts that the system is always in one of S1 - S6, for every p *) + ImpInv :: "histType => PrIds => stpred" + "ImpInv rmhist p == PRED ( S1 rmhist p | S2 rmhist p | S3 rmhist p + | S4 rmhist p | S5 rmhist p | S6 rmhist p)" + + resbar :: "histType => resType" (* refinement mapping *) + "resbar rmhist s p == + (if (S1 rmhist p s | S2 rmhist p s) + then ires s p + else if S3 rmhist p s + then if rmhist s p = histA + then ires s p else MemFailure + else if S4 rmhist p s + then if (rmhist s p = histB & ires s p = NotAResult) + then MemFailure else ires s p + else if S5 rmhist p s + then res (rmCh s p) + else if S6 rmhist p s + then if res (crCh s p) = RPCFailure + then MemFailure else res (crCh s p) + else NotAResult)" (* dummy value *) rules - MVOKBARF_def "MVOKBARF v == (MemVal v) | (v = OK) | (v = BadArg) | (v = RPCFailure)" - MVOKBA_def "MVOKBA v == (MemVal v) | (v = OK) | (v = BadArg)" - MVNROKBA_def "MVNROKBA v == (MemVal v) | (v = NotAResult) | (v = OK) | (v = BadArg)" - (* the "base" variables: everything except resbar and hist (for any index) *) - MI_base "base_var " - - (* Environment's next-state relation *) - ENext_def "ENext p == REX l. #(MemLoc l) .& Call memCh p (#(Inl (read,l)))" - - (* Specification of the history variable used in the proof *) - HInit_def "$(HInit rmhist p) .= ($(rmhist@p) .= #histA)" - HNext_def "HNext rmhist p == - (rmhist@p)$ .= - (.if (MemReturn rmCh ires p .| RPCFail crCh rmCh rst p) - .then #histB - .else .if (MClkReply memCh crCh cst p) - .then #histA - .else $(rmhist@p))" - HistP_def "HistP rmhist p == - Init($(HInit rmhist p)) - .& [][HNext rmhist p]_" - Hist_def "Hist rmhist == RALL p. HistP rmhist p" - - (* definitions of e,c,r,m *) - e_def "e p == caller memCh @ p" - c_def "c p == " - r_def "r p == " - m_def "m p == " - - (* definition of the implementation (without the history variable) *) - IPImp_def "IPImp p == Init(.~ $(Calling memCh p)) .& [][ENext p]_(e p) - .& MClkIPSpec memCh crCh cst p - .& RPCIPSpec crCh rmCh rst p - .& RPSpec rmCh mem ires p - .& (RALL l. #(MemLoc l) .-> MSpec rmCh mem ires l)" - - ImpInit_def "$(ImpInit p) .= ( .~ $(Calling memCh p) \ -\ .& $(MClkInit crCh cst p) \ -\ .& $(RPCInit rmCh rst p) \ -\ .& $(PInit ires p))" - - ImpNext_def "ImpNext p == [ENext p]_(e p) - .& [MClkNext memCh crCh cst p]_(c p) - .& [RPCNext crCh rmCh rst p]_(r p) - .& [RNext rmCh mem ires p]_(m p)" - - ImpLive_def "ImpLive p == WF(MClkFwd memCh crCh cst p)_(c p) - .& SF(MClkReply memCh crCh cst p)_(c p) - .& WF(RPCNext crCh rmCh rst p)_(r p) - .& WF(RNext rmCh mem ires p)_(m p) - .& WF(MemReturn rmCh ires p)_(m p)" - - Impl_def "Implementation == (RALL p. Init(.~ $(Calling memCh p)) .& [][ENext p]_(e p)) - .& MClkISpec memCh crCh cst - .& RPCISpec crCh rmCh rst - .& IRSpec rmCh mem ires" - - ImpInv_def "$(ImpInv rmhist p) .= ($(S1 rmhist p) .| $(S2 rmhist p) .| $(S3 rmhist p) .| - $(S4 rmhist p) .| $(S5 rmhist p) .| $(S6 rmhist p))" - - (* Definition of predicate S. - NB: The second conjunct of the definition in the paper is taken care of by - the type definitions. The last conjunct is asserted separately as the memory - invariant MemInv, proved in Memory.ML. *) - S_def "$(S rmhist ecalling ccalling rcalling cs rs hs1 hs2 p) .= - ( ($(Calling memCh p) .= # ecalling) - .& ($(Calling crCh p) .= # ccalling) - .& (# ccalling .-> (arg[ $(crCh@p)] .= MClkRelayArg[ arg[$(memCh@p)] ])) - .& ((.~ # ccalling .& ($(cst@p) .= # clkB)) .-> MVOKBARF[ res[$(crCh@p)] ]) - .& ($(Calling rmCh p) .= # rcalling) - .& (# rcalling .-> (arg[ $(rmCh@p)] .= RPCRelayArg[ arg[$(crCh@p)] ])) - .& (.~ # rcalling .-> ($(ires@p) .= # NotAResult)) - .& ((.~ # rcalling .& ($(rst@p) .= # rpcB)) .-> MVOKBA[ res[$(rmCh@p)] ]) - .& ($(cst@p) .= # cs) - .& ($(rst@p) .= # rs) - .& (($(rmhist@p) .= #hs1) .| ($(rmhist@p) .= #hs2)) - .& (MVNROKBA[ $(ires@p)]))" - - S1_def "$(S1 rmhist p) .= $(S rmhist False False False clkA rpcA histA histA p)" - S2_def "$(S2 rmhist p) .= $(S rmhist True False False clkA rpcA histA histA p)" - S3_def "$(S3 rmhist p) .= $(S rmhist True True False clkB rpcA histA histB p)" - S4_def "$(S4 rmhist p) .= $(S rmhist True True True clkB rpcB histA histB p)" - S5_def "$(S5 rmhist p) .= $(S rmhist True True False clkB rpcB histB histB p)" - S6_def "$(S6 rmhist p) .= $(S rmhist True False False clkB rpcA histB histB p)" - - (* Definition of the refinement mapping resbar for result *) - resbar_def "$((resbar rmhist) @ p) .= - (.if ($(S1 rmhist p) .| $(S2 rmhist p)) - .then $(ires@p) - .else .if $(S3 rmhist p) - .then .if $(rmhist@p) .= #histA - .then $(ires@p) .else # MemFailure - .else .if $(S4 rmhist p) - .then .if ($(rmhist@p) .= #histB) .& ($(ires@p) .= # NotAResult) - .then #MemFailure .else $(ires@p) - .else .if $(S5 rmhist p) - .then res[$(rmCh@p)] - .else .if $(S6 rmhist p) - .then .if res[$(crCh@p)] .= #RPCFailure - .then #MemFailure .else res[$(crCh@p)] - .else #NotAResult)" (* dummy value *) + MI_base "basevars (caller memCh!p, + (rtrner memCh!p, caller crCh!p, cst!p), + (rtrner crCh!p, caller rmCh!p, rst!p), + (mm!l, rtrner rmCh!p, ires!p))" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemoryParameters.ML --- a/src/HOL/TLA/Memory/MemoryParameters.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemoryParameters.ML Mon Feb 08 13:02:56 1999 +0100 @@ -6,19 +6,24 @@ RPC-Memory example: memory parameters (ML file) *) +(* val MP_simps = [BadArgNoMemVal,MemFailNoMemVal,InitValMemVal,NotAResultNotVal, NotAResultNotOK, NotAResultNotBA, NotAResultNotMF] @ (map (fn x => x RS not_sym) [NotAResultNotOK, NotAResultNotBA, NotAResultNotMF]); +*) +Addsimps ([BadArgNoMemVal,MemFailNoMemVal,InitValMemVal,NotAResultNotVal, + NotAResultNotOK, NotAResultNotBA, NotAResultNotMF] + @ (map (fn x => x RS not_sym) + [NotAResultNotOK, NotAResultNotBA, NotAResultNotMF])); (* Auxiliary rules *) qed_goal "MemValNotAResultE" MemoryParameters.thy - "[| MemVal x; (x ~= NotAResult ==> P) |] ==> P" - (fn [min,maj] => [rtac maj 1, - case_tac "x = NotAResult" 1, - cut_facts_tac [min,NotAResultNotVal] 1, - ALLGOALS Asm_full_simp_tac - ]); + "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P" + (fn prems => [resolve_tac prems 1, + cut_facts_tac (NotAResultNotVal::prems) 1, + Force_tac 1 + ]); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/MemoryParameters.thy --- a/src/HOL/TLA/Memory/MemoryParameters.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/MemoryParameters.thy Mon Feb 08 13:02:56 1999 +0100 @@ -11,21 +11,24 @@ MemoryParameters = Datatype + RPCMemoryParams + -(* the memory operations. nb: data types must be defined in theories - that do not include Intensional -- otherwise the induction rule - can't be type-checked unambiguously. -*) +(* the memory operations *) +(*** datatype Rd = read datatype Wr = write +***) +datatype memOp = read Locs | write Locs Vals + +(*** types (* legal arguments for the memory *) memArgType = "(Rd * Locs) + (Wr * Locs * Vals)" +***) consts (* memory locations and contents *) - MemLoc :: "Locs => bool" - MemVal :: "Vals => bool" + MemLoc :: Locs set + MemVal :: Vals set (* some particular values *) OK :: "Vals" @@ -38,13 +41,11 @@ rules (* basic assumptions about the above constants and predicates *) - BadArgNoMemVal "~MemVal(BadArg)" - MemFailNoMemVal "~MemVal(MemFailure)" - InitValMemVal "MemVal(InitVal)" - NotAResultNotVal "~MemVal(NotAResult)" + BadArgNoMemVal "BadArg ~: MemVal" + MemFailNoMemVal "MemFailure ~: MemVal" + InitValMemVal "InitVal : MemVal" + NotAResultNotVal "NotAResult ~: MemVal" NotAResultNotOK "NotAResult ~= OK" NotAResultNotBA "NotAResult ~= BadArg" NotAResultNotMF "NotAResult ~= MemFailure" end - - diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/ProcedureInterface.ML --- a/src/HOL/TLA/Memory/ProcedureInterface.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/ProcedureInterface.ML Mon Feb 08 13:02:56 1999 +0100 @@ -3,14 +3,15 @@ Author: Stephan Merz Copyright: 1997 University of Munich - Procedure interface (ML file) + Procedure interface (theorems and proofs) *) Addsimps [slice_def]; +val mem_css = (claset(), simpset()); (* ---------------------------------------------------------------------------- *) -val Procedure_defs = [caller_def, rtrner_def, action_rewrite Calling_def, +val Procedure_defs = [caller_def, rtrner_def, Calling_def, Call_def, Return_def, PLegalCaller_def, LegalCaller_def, PLegalReturner_def, LegalReturner_def]; @@ -18,22 +19,29 @@ (* sample theorems (not used in the proof): 1. calls and returns are mutually exclusive -qed_goal "CallReturnMutex" ProcedureInterface.thy - "Call ch p v .-> .~ Return ch p w" - (fn prems => [ auto_tac (action_css addsimps2 [Call_def,Return_def]) ]); +qed_goal "CallNotReturn" ProcedureInterface.thy + "|- Call ch p v --> ~ Return ch p w" + (fn prems => [ auto_tac (temp_css addsimps2 [Call_def,Return_def]) ]); 2. enabledness of calls and returns - NB: action_simp_tac is significantly faster than auto_tac qed_goal "Call_enabled" ProcedureInterface.thy - "!!p. base_var ((caller ch)@p) ==> (.~ $(Calling ch p) .-> $(Enabled (Call ch p (#v))))" + "!!p. basevars ((caller ch)!p) ==> |- ~ Calling ch p --> Enabled (Call ch p v)" (fn _ => [action_simp_tac (simpset() addsimps [caller_def, Call_def]) [] [base_enabled,Pair_inject] 1 ]); +qed_goal "Call_enabled_rew" ProcedureInterface.thy + "basevars ((caller ch)!p) ==> |- Enabled (Call ch p v) = (~Calling ch p)" + (fn [prem] => [auto_tac (mem_css addsimps2 [Call_def]), + force_tac (mem_css addsimps2 [enabled_def]) 1, + enabled_tac prem 1, + action_simp_tac (simpset() addsimps [caller_def]) [] [Pair_inject] 1 + ]); + qed_goal "Return_enabled" ProcedureInterface.thy - "!!p. base_var ((rtrner ch)@p) ==> $(Calling ch p) .-> $(Enabled (Return ch p (#v)))" + "!!p. basevars ((rtrner ch)!p) ==> |- Calling ch p --> Enabled (Return ch p v)" (fn _ => [action_simp_tac (simpset() addsimps [rtrner_def, Return_def]) [] [base_enabled,Pair_inject] 1 ]); @@ -42,21 +50,11 @@ (* Calls and returns change their subchannel *) qed_goal "Call_changed" ProcedureInterface.thy - "Call ch p v .-> _((caller ch)@p)" - (fn _ => [auto_tac (claset(), - simpset() addsimps [angle_def,Call_def,caller_def, - action_rewrite Calling_def]) - ]); + "|- Call ch p v --> _((caller ch)!p)" + (fn _ => [ auto_tac (mem_css addsimps2 [angle_def,Call_def,caller_def,Calling_def]) ]); qed_goal "Return_changed" ProcedureInterface.thy - "Return ch p v .-> _((rtrner ch)@p)" - (fn _ => [auto_tac (claset(), - simpset() addsimps [angle_def,Return_def,rtrner_def, - action_rewrite Calling_def]) - ]); + "|- Return ch p v --> _((rtrner ch)!p)" + (fn _ => [ auto_tac (mem_css addsimps2 [angle_def,Return_def,rtrner_def,Calling_def]) ]); -(* For convenience, generate elimination rules. - These rules loop if angle_def is active! *) -bind_thm("Call_changedE", action_impE Call_changed); -bind_thm("Return_changedE", action_impE Return_changed); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/ProcedureInterface.thy --- a/src/HOL/TLA/Memory/ProcedureInterface.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/ProcedureInterface.thy Mon Feb 08 13:02:56 1999 +0100 @@ -29,9 +29,6 @@ arg :: "('a,'r) chan => 'a" res :: "('a,'r) chan => 'r" - (* slice through array-valued state function *) - "@" :: "('a => 'b) stfun => 'a => 'b stfun" (infixl 20) - (* state functions *) caller :: "('a,'r) channel => (PrIds => (bit * 'a)) stfun" rtrner :: "('a,'r) channel => (PrIds => (bit * 'r)) stfun" @@ -40,8 +37,8 @@ Calling :: "('a,'r) channel => PrIds => stpred" (* actions *) - Call :: "('a,'r) channel => PrIds => 'a trfct => action" - Return :: "('a,'r) channel => PrIds => 'r trfct => action" + ACall :: "('a,'r) channel => PrIds => 'a stfun => action" + AReturn :: "('a,'r) channel => PrIds => 'r stfun => action" (* temporal formulas *) PLegalCaller :: "('a,'r) channel => PrIds => temporal" @@ -49,27 +46,42 @@ PLegalReturner :: "('a,'r) channel => PrIds => temporal" LegalReturner :: "('a,'r) channel => temporal" -rules - slice_def "(x@i) s == x s i" + (* slice through array-valued state function *) + slice :: "('a => 'b) stfun => 'a => 'b stfun" + +syntax + "_slice" :: [lift, 'a] => lift ("(_!_)" [70,70] 70) + + "_Call" :: ['a, 'b, lift] => lift ("(Call _ _ _)" [90,90,90] 90) + "_Return" :: ['a, 'b, lift] => lift ("(Return _ _ _)" [90,90,90] 90) - caller_def "caller ch s p == (cbit (ch s p), arg (ch s p))" - rtrner_def "rtrner ch s p == (rbit (ch s p), res (ch s p))" +translations + "_slice" == "slice" + + "_Call" == "ACall" + "_Return" == "AReturn" + +rules + slice_def "(PRED (x!i)) s == x s i" - Calling_def "$(Calling ch p) .= (cbit[$(ch@p)] .~= rbit[$(ch@p)])" - Call_def "Call ch p v == .~ $(Calling ch p) - .& (cbit[$(ch@p)])` .~= rbit[$(ch@p)] - .& (arg[$(ch@p)])` .= v" - Return_def "Return ch p v == $(Calling ch p) - .& (rbit[$(ch@p)])` .= cbit[$(ch@p)] - .& (res[$(ch@p)])` .= v" + caller_def "caller ch == %s p. (cbit (ch s p), arg (ch s p))" + rtrner_def "rtrner ch == %s p. (rbit (ch s p), res (ch s p))" - PLegalCaller_def "PLegalCaller ch p == - Init(.~ $(Calling ch p)) - .& [][ REX a. Call ch p (#a) ]_((caller ch)@p)" - LegalCaller_def "LegalCaller ch == RALL p. PLegalCaller ch p" - PLegalReturner_def "PLegalReturner ch p == - [][ REX v. Return ch p (#v) ]_((rtrner ch)@p)" - LegalReturner_def "LegalReturner ch == RALL p. PLegalReturner ch p" + Calling_def "Calling ch p == PRED cbit< ch!p > ~= rbit< ch!p >" + Call_def "(ACT Call ch p v) == ACT ~ $Calling ch p + & (cbit` ~= $rbit) + & (arg` = $v)" + Return_def "(ACT Return ch p v) == ACT $Calling ch p + & (rbit` = $cbit) + & (res` = $v)" + PLegalCaller_def "PLegalCaller ch p == TEMP + Init(~ Calling ch p) + & [][ ? a. Call ch p a ]_((caller ch)!p)" + LegalCaller_def "LegalCaller ch == TEMP (! p. PLegalCaller ch p)" + PLegalReturner_def "PLegalReturner ch p == TEMP + [][ ? v. Return ch p v ]_((rtrner ch)!p)" + LegalReturner_def "LegalReturner ch == TEMP (! p. PLegalReturner ch p)" end + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/ROOT.ML --- a/src/HOL/TLA/Memory/ROOT.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/ROOT.ML Mon Feb 08 13:02:56 1999 +0100 @@ -1,2 +1,2 @@ -use_thy "Memory"; +use_thy "MemoryImplementation"; diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/RPC.ML --- a/src/HOL/TLA/Memory/RPC.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/RPC.ML Mon Feb 08 13:02:56 1999 +0100 @@ -3,57 +3,54 @@ Author: Stephan Merz Copyright: 1997 University of Munich - RPC-Memory example: RPC specification (ML file) + RPC-Memory example: RPC specification (theorems and proofs) *) -val RPC_action_defs = - [RPCInit_def RS inteq_reflection] - @ [RPCFwd_def, RPCReject_def, RPCFail_def, RPCReply_def, RPCNext_def]; +val RPC_action_defs = [RPCInit_def, RPCFwd_def, RPCReject_def, RPCFail_def, + RPCReply_def, RPCNext_def]; val RPC_temp_defs = [RPCIPSpec_def, RPCISpec_def]; +val mem_css = (claset(), simpset()); + (* The RPC component engages in an action for process p only if there is an outstanding, unanswered call for that process. *) qed_goal "RPCidle" RPC.thy - ".~ $(Calling send p) .-> .~ RPCNext send rcv rst p" - (fn _ => [ auto_tac (action_css addsimps2 (Return_def::RPC_action_defs)) ]); + "|- ~$(Calling send p) --> ~RPCNext send rcv rst p" + (fn _ => [ auto_tac (mem_css addsimps2 (Return_def::RPC_action_defs)) ]); qed_goal "RPCbusy" RPC.thy - "$(Calling rcv p) .& ($(rst@p) .= #rpcB) .-> .~ RPCNext send rcv rst p" - (fn _ => [ auto_tac (action_css addsimps2 (RP_simps @ RPC_action_defs)) ]); - -(* unlifted versions as introduction rules *) - -bind_thm("RPCidleI", action_mp RPCidle); -bind_thm("RPCbusyI", action_mp RPCbusy); + "|- $(Calling rcv p) & $(rst!p) = #rpcB --> ~RPCNext send rcv rst p" + (fn _ => [ auto_tac (mem_css addsimps2 RPC_action_defs) ]); (* RPC failure actions are visible. *) qed_goal "RPCFail_vis" RPC.thy - "RPCFail send rcv rst p .-> _" - (fn _ => [auto_tac (claset() addSEs [Return_changedE], + "|- RPCFail send rcv rst p --> \ +\ _(rst!p, rtrner send!p, caller rcv!p)" + (fn _ => [auto_tac (claset() addSDs [Return_changed], simpset() addsimps [angle_def,RPCNext_def,RPCFail_def]) ]); qed_goal "RPCFail_Next_enabled" RPC.thy - "Enabled (RPCFail send rcv rst p) s \ -\ ==> Enabled (_) s" - (fn [prem] => [REPEAT (resolve_tac [prem RS enabled_mono,RPCFail_vis] 1)]); + "|- Enabled (RPCFail send rcv rst p) --> \ +\ Enabled (_(rst!p, rtrner send!p, caller rcv!p))" + (fn _ => [force_tac (mem_css addSEs2 [enabled_mono,RPCFail_vis]) 1]); (* Enabledness of some actions *) qed_goal "RPCFail_enabled" RPC.thy - "!!p. base_var ==> \ -\ .~ $(Calling rcv p) .& $(Calling send p) .-> $(Enabled (RPCFail send rcv rst p))" + "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \ +\ |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)" (fn _ => [action_simp_tac (simpset() addsimps [RPCFail_def,Return_def,caller_def,rtrner_def]) [] [base_enabled,Pair_inject] 1 ]); qed_goal "RPCReply_enabled" RPC.thy - "!!p. base_var ==> \ -\ .~ $(Calling rcv p) .& $(Calling send p) .& $(rst@p) .= #rpcB \ -\ .-> $(Enabled (RPCReply send rcv rst p))" + "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \ +\ |- ~Calling rcv p & Calling send p & rst!p = #rpcB \ +\ --> Enabled (RPCReply send rcv rst p)" (fn _ => [action_simp_tac (simpset() addsimps [RPCReply_def,Return_def,caller_def,rtrner_def]) [] [base_enabled,Pair_inject] 1]); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/RPC.thy --- a/src/HOL/TLA/Memory/RPC.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/RPC.thy Mon Feb 08 13:02:56 1999 +0100 @@ -7,15 +7,13 @@ Logic Image: TLA RPC-Memory example: RPC specification - For simplicity, specify the instance of RPC that is used in the - memory implementation (ignoring the BadCall exception). *) -RPC = RPCParameters + ProcedureInterface + +RPC = RPCParameters + ProcedureInterface + Memory + types - rpcSndChType = "(rpcArgType,Vals) channel" - rpcRcvChType = "(memArgType,Vals) channel" + rpcSndChType = "(rpcOp,Vals) channel" + rpcRcvChType = "memChType" rpcStType = "(PrIds => rpcState) stfun" consts @@ -34,49 +32,47 @@ RPCISpec :: "rpcSndChType => rpcRcvChType => rpcStType => temporal" rules - RPCInit_def "$(RPCInit rcv rst p) .= - ($(rst@p) .= # rpcA - .& .~ $(Calling rcv p))" + RPCInit_def "RPCInit rcv rst p == PRED ((rst!p = #rpcA) & ~Calling rcv p)" - RPCFwd_def "RPCFwd send rcv rst p == + RPCFwd_def "RPCFwd send rcv rst p == ACT $(Calling send p) - .& $(rst@p) .= # rpcA - .& IsLegalRcvArg[ arg[ $(send@p) ] ] - .& Call rcv p (RPCRelayArg[ arg[ $(send@p)] ]) - .& (rst@p)$ .= # rpcB - .& unchanged (rtrner send @ p)" + & $(rst!p) = # rpcA + & IsLegalRcvArg> + & Call rcv p RPCRelayArg> + & (rst!p)$ = # rpcB + & unchanged (rtrner send!p)" - RPCReject_def "RPCReject send rcv rst p == - $(rst@p) .= # rpcA - .& .~ IsLegalRcvArg[ arg[ $(send@p) ] ] - .& Return send p (#BadCall) - .& unchanged <(rst@p), (caller rcv @ p)>" + RPCReject_def "RPCReject send rcv rst p == ACT + $(rst!p) = # rpcA + & ~IsLegalRcvArg> + & Return send p #BadCall + & unchanged ((rst!p), (caller rcv!p))" - RPCFail_def "RPCFail send rcv rst p == - .~ $(Calling rcv p) - .& Return send p (#RPCFailure) - .& (rst@p)$ .= #rpcA - .& unchanged (caller rcv @ p)" + RPCFail_def "RPCFail send rcv rst p == ACT + ~$(Calling rcv p) + & Return send p #RPCFailure + & (rst!p)$ = #rpcA + & unchanged (caller rcv!p)" - RPCReply_def "RPCReply send rcv rst p == - .~ $(Calling rcv p) - .& $(rst@p) .= #rpcB - .& Return send p (res[$(rcv@p)]) - .& (rst@p)$ .= #rpcA - .& unchanged (caller rcv @ p)" + RPCReply_def "RPCReply send rcv rst p == ACT + ~$(Calling rcv p) + & $(rst!p) = #rpcB + & Return send p res + & (rst!p)$ = #rpcA + & unchanged (caller rcv!p)" - RPCNext_def "RPCNext send rcv rst p == - RPCFwd send rcv rst p - .| RPCReject send rcv rst p - .| RPCFail send rcv rst p - .| RPCReply send rcv rst p" + RPCNext_def "RPCNext send rcv rst p == ACT + ( RPCFwd send rcv rst p + | RPCReject send rcv rst p + | RPCFail send rcv rst p + | RPCReply send rcv rst p)" - RPCIPSpec_def "RPCIPSpec send rcv rst p == - Init($(RPCInit rcv rst p)) - .& [][ RPCNext send rcv rst p ]_ - .& WF(RPCNext send rcv rst p)_" + RPCIPSpec_def "RPCIPSpec send rcv rst p == TEMP + Init RPCInit rcv rst p + & [][ RPCNext send rcv rst p ]_(rst!p, rtrner send!p, caller rcv!p) + & WF(RPCNext send rcv rst p)_(rst!p, rtrner send!p, caller rcv!p)" - RPCISpec_def "RPCISpec send rcv rst == RALL p. RPCIPSpec send rcv rst p" + RPCISpec_def "RPCISpec send rcv rst == TEMP (! p. RPCIPSpec send rcv rst p)" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/RPCMemoryParams.thy --- a/src/HOL/TLA/Memory/RPCMemoryParams.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/RPCMemoryParams.thy Mon Feb 08 13:02:56 1999 +0100 @@ -12,9 +12,9 @@ RPCMemoryParams = HOL + types - bit = "bool" (* signal wires for the procedure interface *) - (* Defined as bool for simplicity. All I should really need is *) - (* the existence of two distinct values. *) + bit = "bool" (* Signal wires for the procedure interface. + Defined as bool for simplicity. All I should really need is + the existence of two distinct values. *) Locs (* "syntactic" value type *) Vals (* "syntactic" value type *) PrIds (* process id's *) diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/RPCParameters.ML --- a/src/HOL/TLA/Memory/RPCParameters.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/RPCParameters.ML Mon Feb 08 13:02:56 1999 +0100 @@ -3,10 +3,15 @@ Author: Stephan Merz Copyright: 1997 University of Munich - RPC-Memory example: RPC parameters (ML file) + RPC-Memory example: RPC parameters (theorems and proofs) *) +(* val RP_simps = MP_simps @ [RFNoMemVal, NotAResultNotRF, OKNotRF, BANotRF] @ (map (fn x => x RS not_sym) [NotAResultNotRF, OKNotRF, BANotRF]) - @ rpcOps.simps @ rpcState.simps; + @ rpcState.simps @ rpcOp.simps; +*) + +Addsimps ([RFNoMemVal, NotAResultNotRF, OKNotRF, BANotRF] + @ (map (fn x => x RS not_sym) [NotAResultNotRF, OKNotRF, BANotRF])); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Memory/RPCParameters.thy --- a/src/HOL/TLA/Memory/RPCParameters.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Memory/RPCParameters.thy Mon Feb 08 13:02:56 1999 +0100 @@ -13,9 +13,10 @@ RPCParameters = MemoryParameters + -datatype rpcOps = remoteCall +datatype rpcOp = memcall memOp | othercall Vals datatype rpcState = rpcA | rpcB +(*** types (* type of RPC arguments other than memory calls *) noMemArgType @@ -24,31 +25,44 @@ arities noMemArgType :: term +***) consts (* some particular return values *) - RPCFailure :: "Vals" - BadCall :: "Vals" + RPCFailure :: Vals + BadCall :: Vals (* Translate an rpc call to a memory call and test if the current argument is legal for the receiver (i.e., the memory). This can now be a little simpler than for the generic RPC component. RelayArg returns an arbitrary memory call for illegal arguments. *) - IsLegalRcvArg :: "rpcArgType => bool" - RPCRelayArg :: "rpcArgType => memArgType" +(*** + IsLegalRcvArg :: rpcArgType => bool + RPCRelayArg :: rpcArgType => memArgType +***) + IsLegalRcvArg :: rpcOp => bool + RPCRelayArg :: rpcOp => memOp rules (* RPCFailure is different from MemVals and exceptions *) - RFNoMemVal "~(MemVal RPCFailure)" + RFNoMemVal "RPCFailure ~: MemVal" NotAResultNotRF "NotAResult ~= RPCFailure" OKNotRF "OK ~= RPCFailure" BANotRF "BadArg ~= RPCFailure" +(*** IsLegalRcvArg_def "IsLegalRcvArg ra == EX marg. ra = Inl (remoteCall,marg)" RPCRelayArg_def "RPCRelayArg ra == case ra of Inl (rm) => (snd rm) - | Inr (rn) => Inl (read, @ l. True)" - + | Inr (rn) => (read, @ l. True)" +***) +defs + IsLegalRcvArg_def "IsLegalRcvArg ra == + case ra of (memcall m) => True + | (othercall v) => False" + RPCRelayArg_def "RPCRelayArg ra == + case ra of (memcall m) => m + | (othercall v) => arbitrary" end diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/README.html --- a/src/HOL/TLA/README.html Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/README.html Mon Feb 08 13:02:56 1999 +0100 @@ -1,46 +1,59 @@ -HOL/TLA/README +HOL/TLA -

TLA: A formalization of TLA in HOL

- -Author: Stephan Merz
-Copyright 1997 Universität München

+

TLA: Lamport's Temporal Logic of Actions

-The distribution contains a representation of Lamport's - -Temporal Logic of Actions -in Isabelle/HOL. - -

+TLA +is a linear-time temporal logic introduced by Leslie Lamport in +The Temporal Logic of Actions (ACM TOPLAS 16(3), 1994, +872-923). Unlike other temporal logics, both systems and properties +are represented as logical formulas, and logical connectives such as +implication, conjunction, and existential quantification represent +structural relations such as refinement, parallel composition, and +hiding. TLA has been applied to numerous case studies. -The encoding is mainly oriented towards practical verification -examples. It does not contain a formalization of TLA's semantics, -although it could be an interesting exercise to add such a formalization -to the existing representation. Instead, it is based on a -complete axiomatization -of the "raw" (stuttering-sensitive) variant of propositional TLA. -There is also a -design note -that explains the basic setup and use of the prover. - -

- -The distribution includes the following examples: +

This directory formalizes TLA in Isabelle/HOL, as follows:

    -
  • a verification of Lamport's increment example - (subdirectory inc),

    - -

  • a proof that two buffers in a row implement a single buffer - (subdirectory buffer), and

    - -

  • the verification of Broy and Lamport's RPC-Memory example. For details see:
    - - Martín Abadi, Leslie Lamport, and Stephan Merz: - - A TLA Solution to the RPC-Memory Specification Problem. - In: Formal System Specification, LNCS 1169, 1996, 21-69. +
  • Theory Intensional prepares the + ground by introducing basic syntax for "lifted", possibl-world based + logics. +
  • Theories Stfun and + Action represent the state and transition + level formulas of TLA, evaluated over single states and pairs of + states. +
  • Theory Init introduces temporal logic + and defines conversion functions from nontemporal to temporal + formulas. +
  • Theory TLA axiomatizes proper temporal + logic.
-If you use Isabelle/TLA and have any comments, suggestions or contributions, -please contact Stephan Merz. +Please consult the +design notes +for further information regarding the setup and use of this encoding +of TLA. - +

+The theories are accompanied by a small number of examples: +

    +
  • Inc: Lamport's increment + example, a standard TLA benchmark, illustrates an elementary TLA + proof. +
  • Buffer: a proof that two buffers + in a row implement a single buffer, uses a simple refinement + mapping. +
  • Memory: a verification of (the + untimed part of) Broy and Lamport's RPC-Memory case study, + more fully explained in LNCS 1169 (the + TLA + solution is available separately). +
+ +
+ +
+Stephan Merz +
+ +Last modified: Mon Jan 25 14:06:43 MET 1999 + + \ No newline at end of file diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/ROOT.ML --- a/src/HOL/TLA/ROOT.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/ROOT.ML Mon Feb 08 13:02:56 1999 +0100 @@ -5,23 +5,6 @@ val banner = "Temporal Logic of Actions"; -(* - raise the ambiguity level to avoid ambiguity warnings; - since Trueprop and TrueInt have both empty syntax, there is - an unavoidable ambiguity in the TLA (actually, Intensional) grammar. -*) -Syntax.ambiguity_level := 10000; - -(*FIXME: the old auto_tac is sometimes needed!*) -fun old_auto_tac (cs,ss) = - let val cs' = cs addss ss - in EVERY [TRY (safe_tac cs'), - REPEAT (FIRSTGOAL (fast_tac cs')), - TRY (safe_tac (cs addSss ss)), - prune_params_tac] - end; - - use_thy "TLA"; val TLA_build_completed = (); diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Stfun.ML --- a/src/HOL/TLA/Stfun.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Stfun.ML Mon Feb 08 13:02:56 1999 +0100 @@ -1,25 +1,21 @@ (* File: Stfun.ML Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Lemmas and tactics for states and state functions. *) -(* A stronger version of existential elimination (goal needn't be boolean) *) -qed_goalw "exE_prop" HOL.thy [Ex_def] - "[| ? x::'a. P(x); !!x. P(x) ==> PROP R |] ==> PROP R" - (fn prems => [REPEAT(resolve_tac prems 1)]); +(* [| basevars v; !!x. v x = c ==> Q |] ==> Q *) +bind_thm("baseE", (standard (basevars RS exE))); -(* Might as well use that version in automated proofs *) -AddSEs [exE_prop]; +(* ------------------------------------------------------------------------------- + The following shows that there should not be duplicates in a "stvars" tuple: -(* [| base_var v; !!x. v x = c ==> PROP R |] ==> PROP R *) -bind_thm("baseE", (standard (base_var RS exE_prop))); +Goal "!!v. basevars (v::bool stfun, v) ==> False"; +by (etac baseE 1); +by (subgoal_tac "(LIFT (v,v)) x = (True, False)" 1); +by (atac 2); +by (Asm_full_simp_tac 1); -qed_goal "PairVarE" Stfun.thy - "[| u = (x,y); [| v u = x; w u = y |] ==> PROP R |] ==> PROP R" - (fn prems => [cut_facts_tac prems 1, resolve_tac prems 1, - ALLGOALS (asm_full_simp_tac (simpset() addsimps [pairSF_def])) - ]); - +------------------------------------------------------------------------------- *) diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/Stfun.thy --- a/src/HOL/TLA/Stfun.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/Stfun.thy Mon Feb 08 13:02:56 1999 +0100 @@ -1,15 +1,15 @@ (* File: TLA/Stfun.thy Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Theory Name: Stfun Logic Image: HOL -States and state functions for TLA +States and state functions for TLA as an "intensional" logic. *) -Stfun = Prod + +Stfun = Intensional + types state @@ -17,40 +17,49 @@ stpred = "bool stfun" arities - state :: term + state :: term + +instance + state :: world consts - (* For simplicity, we do not syntactically distinguish between state variables - and state functions, and treat "state" as an anonymous type. But we need a - "meta-predicate" to identify "base" state variables that represent the state - components of a system, in particular to define the enabledness of actions. + (* Formalizing type "state" would require formulas to be tagged with + their underlying state space and would result in a system that is + much harder to use. (Unlike Hoare logic or Unity, TLA has quantification + over state variables, and therefore one usually works with different + state spaces within a single specification.) Instead, "state" is just + an anonymous type whose only purpose is to provide "Skolem" constants. + Moreover, we do not define a type of state variables separate from that + of arbitrary state functions, again in order to simplify the definition + of flexible quantification later on. Nevertheless, we need to distinguish + state variables, mainly to define the enabledness of actions. The user + identifies (tuples of) "base" state variables in a specification via the + "meta predicate" stvars. + NOTE: There should not be duplicates in the tuple! *) - base_var :: "'a stfun => bool" - - (* lift tupling to state functions *) - pairSF :: "['a stfun, 'b stfun] => ('a * 'b) stfun" + stvars :: "'a stfun => bool" syntax - "@tupleSF" :: "args => ('a * 'b) stfun" ("(1<_>)") + "PRED" :: lift => 'a ("PRED _") + "_stvars" :: lift => bool ("basevars _") translations - "" == " >" - "" == "pairSF x y" - "" => "x" + "PRED P" => "(P::state => _)" + "_stvars" == "stvars" rules - (* tupling *) - pairSF_def "(s) = (v(s),w(s))" + (* Base variables may be assigned arbitrary (type-correct) values. + Note that vs may be a tuple of variables. The rule would be unsound + if vs contained duplicates. + *) + basevars "basevars vs ==> EX u. vs u = c" + base_pair "basevars (x,y) ==> basevars x & basevars y" + (* Since the unit type has just one value, any state function can be + regarded as "base". The following axiom can sometimes be useful + because it gives a trivial solution for "basevars" premises. + *) + unit_base "basevars (v::unit stfun)" - (* "base" variables may be assigned arbitrary values by states. - NB: It's really stronger than that because "u" doesn't depend - on either c or v. In particular, if "==>" were replaced - with "==", base_pair would (still) not be derivable. - *) - base_var "base_var v ==> EX u. v u = c" - - (* a tuple of variables is "base" if each variable is "base" *) - base_pair "base_var = (base_var v & base_var w)" end ML diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/TLA.ML --- a/src/HOL/TLA/TLA.ML Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/TLA.ML Mon Feb 08 13:02:56 1999 +0100 @@ -1,94 +1,72 @@ (* File: TLA/TLA.ML Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Lemmas and tactics for temporal reasoning. *) -(* Specialize intensional introduction/elimination rules to temporal formulas *) +(* Specialize intensional introduction/elimination rules for temporal formulas *) -qed_goal "tempI" TLA.thy "(!!sigma. (sigma |= (F::temporal))) ==> F" +qed_goal "tempI" TLA.thy "(!!sigma. sigma |= (F::temporal)) ==> |- F" (fn [prem] => [ REPEAT (resolve_tac [prem,intI] 1) ]); -qed_goal "tempD" TLA.thy "F::temporal ==> (sigma |= F)" - (fn [prem] => [ REPEAT (resolve_tac [prem,intD] 1) ]); +qed_goal "tempD" TLA.thy "|- (F::temporal) ==> sigma |= F" + (fn [prem] => [ rtac (prem RS intD) 1 ]); -(* ======== Functions to "unlift" temporal implications into HOL rules ====== *) - -(* Basic unlifting introduces a parameter "sigma" and applies basic rewrites, e.g. - F .= G gets (sigma |= F) = (sigma |= G) - F .-> G gets (sigma |= F) --> (sigma |= G) -*) -fun temp_unlift th = rewrite_rule intensional_rews (th RS tempD); - -(* F .-> G becomes sigma |= F ==> sigma |= G *) -fun temp_mp th = zero_var_indexes ((temp_unlift th) RS mp); +(* ======== Functions to "unlift" temporal theorems ====== *) -(* F .-> G becomes [| sigma |= F; sigma |= G ==> R |] ==> R - so that it can be used as an elimination rule +(* The following functions are specialized versions of the corresponding + functions defined in Intensional.ML in that they introduce a + "world" parameter of type "behavior". *) -fun temp_impE th = zero_var_indexes ((temp_unlift th) RS impE); +fun temp_unlift th = + (rewrite_rule action_rews (th RS tempD)) + handle _ => action_unlift th; -(* F .& G .-> H becomes [| sigma |= F; sigma |= G |] ==> sigma |= H *) -fun temp_conjmp th = zero_var_indexes (conjI RS (temp_mp th)); +(* Turn |- F = G into meta-level rewrite rule F == G *) +val temp_rewrite = int_rewrite; -(* F .& G .-> H becomes [| sigma |= F; sigma |= G; (sigma |= H ==> R) |] ==> R *) -fun temp_conjimpE th = zero_var_indexes (conjI RS (temp_impE th)); - -(* Turn F .= G into meta-level rewrite rule F == G *) -fun temp_rewrite th = (rewrite_rule intensional_rews (th RS inteq_reflection)); - +fun temp_use th = + case (concl_of th) of + Const _ $ (Const ("Intensional.Valid", _) $ _) => + ((flatten (temp_unlift th)) handle _ => th) + | _ => th; (* Update classical reasoner---will be updated once more below! *) AddSIs [tempI]; AddDs [tempD]; -val temp_css = action_css addSIs2 [tempI] addDs2 [tempD]; +val temp_css = (claset(), simpset()); val temp_cs = op addss temp_css; -(* ========================================================================= *) -section "Init"; - -(* Push logical connectives through Init. *) -qed_goal "Init_true" TLA.thy "Init(#True) .= #True" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); +(* Modify the functions that add rules to simpsets, classical sets, + and clasimpsets in order to accept "lifted" theorems +*) -qed_goal "Init_false" TLA.thy "Init(#False) .= #False" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); - -qed_goal "Init_not" TLA.thy "Init(.~P) .= (.~Init(P))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); - -qed_goal "Init_and" TLA.thy "Init(P .& Q) .= (Init(P) .& Init(Q))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); - -qed_goal "Init_or" TLA.thy "Init(P .| Q) .= (Init(P) .| Init(Q))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); +local + fun try_rewrite th = + (temp_rewrite th) handle _ => temp_use th +in + val op addsimps = fn (ss, ts) => ss addsimps (map try_rewrite ts) + val op addsimps2 = fn (css, ts) => css addsimps2 (map try_rewrite ts) +end; -qed_goal "Init_imp" TLA.thy "Init(P .-> Q) .= (Init(P) .-> Init(Q))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); - -qed_goal "Init_iff" TLA.thy "Init(P .= Q) .= (Init(P) .= Init(Q))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); - -qed_goal "Init_all" TLA.thy "Init(RALL x. P(x)) .= (RALL x. Init(P(x)))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); +val op addSIs = fn (cs, ts) => cs addSIs (map temp_use ts); +val op addSEs = fn (cs, ts) => cs addSEs (map temp_use ts); +val op addSDs = fn (cs, ts) => cs addSDs (map temp_use ts); +val op addIs = fn (cs, ts) => cs addIs (map temp_use ts); +val op addEs = fn (cs, ts) => cs addEs (map temp_use ts); +val op addDs = fn (cs, ts) => cs addDs (map temp_use ts); -qed_goal "Init_ex" TLA.thy "Init(REX x. P(x)) .= (REX x. Init(P(x)))" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def]) ]); - -val Init_simps = map temp_rewrite - [Init_true,Init_false,Init_not,Init_and,Init_or, - Init_imp,Init_iff,Init_all,Init_ex]; - - -(* Temporal lemmas *) - -qed_goalw "DmdAct" TLA.thy [dmd_def,boxact_def] "(<>(F::action)) .= (<> Init F)" - (fn _ => [auto_tac (temp_css addsimps2 Init_simps)]); +val op addSIs2 = fn (css, ts) => css addSIs2 (map temp_use ts); +val op addSEs2 = fn (css, ts) => css addSEs2 (map temp_use ts); +val op addSDs2 = fn (css, ts) => css addSDs2 (map temp_use ts); +val op addIs2 = fn (css, ts) => css addIs2 (map temp_use ts); +val op addEs2 = fn (css, ts) => css addEs2 (map temp_use ts); +val op addDs2 = fn (css, ts) => css addDs2 (map temp_use ts); (* ------------------------------------------------------------------------- *) @@ -96,29 +74,55 @@ (* ------------------------------------------------------------------------- *) section "Simple temporal logic"; +(* []~F == []~Init F *) +bind_thm("boxNotInit", rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] boxInit)); + +qed_goalw "dmdInit" TLA.thy [dmd_def] "TEMP <>F == TEMP <> Init F" + (fn _ => [rewtac (read_instantiate [("F", "LIFT ~F")] boxInit), + simp_tac (simpset() addsimps Init_simps) 1]); + +bind_thm("dmdNotInit", rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] dmdInit)); + +(* boxInit and dmdInit cannot be used as rewrites, because they loop. + Non-looping instances for state predicates and actions are occasionally useful. +*) +bind_thm("boxInit_stp", read_instantiate [("'a","state")] boxInit); +bind_thm("boxInit_act", read_instantiate [("'a","state * state")] boxInit); +bind_thm("dmdInit_stp", read_instantiate [("'a","state")] dmdInit); +bind_thm("dmdInit_act", read_instantiate [("'a","state * state")] dmdInit); + +(* The symmetric equations can be used to get rid of Init *) +bind_thm("boxInitD", symmetric boxInit); +bind_thm("dmdInitD", symmetric dmdInit); +bind_thm("boxNotInitD", symmetric boxNotInit); +bind_thm("dmdNotInitD", symmetric dmdNotInit); + +val Init_simps = Init_simps @ [boxInitD, dmdInitD, boxNotInitD, dmdNotInitD]; + (* ------------------------ STL2 ------------------------------------------- *) bind_thm("STL2", reflT); -bind_thm("STL2D", temp_mp STL2); -(* The action variants. *) -qed_goalw "STL2b" TLA.thy [boxact_def] "[]P .-> Init P" - (fn _ => [rtac STL2 1]); -bind_thm("STL2bD", temp_mp STL2b); -(* see also STL2b_pr below: "[]P .-> Init(P .& P`)" *) +(* The "polymorphic" (generic) variant *) +qed_goal "STL2_gen" TLA.thy "|- []F --> Init F" + (fn _ => [rewtac (read_instantiate [("F", "F")] boxInit), + rtac STL2 1]); + +(* see also STL2_pr below: "|- []P --> Init P & Init (P`)" *) + (* Dual versions for <> *) -qed_goalw "ImplDmd" TLA.thy [dmd_def] "F .-> <>F" - (fn _ => [ auto_tac (temp_css addSDs2 [STL2D]) ]); -bind_thm ("ImplDmdD", temp_mp ImplDmd); +qed_goalw "InitDmd" TLA.thy [dmd_def] "|- F --> <> F" + (fn _ => [ auto_tac (temp_css addSDs2 [STL2]) ]); -qed_goalw "InitDmd" TLA.thy [dmd_def] "Init(P) .-> <>P" - (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2bD]) ]); -bind_thm("InitDmdD", temp_mp InitDmd); +qed_goal "InitDmd_gen" TLA.thy "|- Init F --> <>F" + (fn _ => [Clarsimp_tac 1, + dtac (temp_use InitDmd) 1, + asm_full_simp_tac (simpset() addsimps [dmdInitD]) 1]); (* ------------------------ STL3 ------------------------------------------- *) -qed_goal "STL3" TLA.thy "([][]F) .= ([]F)" - (K [force_tac (temp_css addIs2 [temp_mp transT,temp_mp STL2]) 1]); +qed_goal "STL3" TLA.thy "|- ([][]F) = ([]F)" + (K [force_tac (temp_css addEs2 [transT,STL2]) 1]); (* corresponding elimination rule introduces double boxes: [| (sigma |= []F); (sigma |= [][]F) ==> PROP W |] ==> PROP W @@ -127,23 +131,31 @@ bind_thm("dup_boxD", (temp_unlift STL3) RS iffD1); (* dual versions for <> *) -qed_goalw "DmdDmd" TLA.thy [dmd_def] "(<><>F) .= (<>F)" +qed_goalw "DmdDmd" TLA.thy [dmd_def] "|- (<><>F) = (<>F)" (fn _ => [ auto_tac (temp_css addsimps2 [STL3]) ]); bind_thm("dup_dmdE", make_elim((temp_unlift DmdDmd) RS iffD2)); bind_thm("dup_dmdD", (temp_unlift DmdDmd) RS iffD1); (* ------------------------ STL4 ------------------------------------------- *) -qed_goal "STL4" TLA.thy "(F .-> G) ==> ([]F .-> []G)" - (fn [prem] => [Auto_tac, - rtac ((temp_mp normalT) RS mp) 1, - REPEAT (ares_tac [prem, necT RS tempD] 1) +qed_goal "STL4" TLA.thy "|- F --> G ==> |- []F --> []G" + (fn [prem] => [Clarsimp_tac 1, + rtac (temp_use normalT) 1, + rtac (temp_use (prem RS necT)) 1, + atac 1 ]); -(* A more practical variant as an (unlifted) elimination rule *) +(* Unlifted version as an elimination rule *) qed_goal "STL4E" TLA.thy - "[| (sigma |= []F); F .-> G |] ==> (sigma |= []G)" - (fn prems => [ REPEAT (resolve_tac (prems @ [temp_mp STL4]) 1) ]); + "[| sigma |= []F; |- F --> G |] ==> sigma |= []G" + (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use STL4]) 1) ]); + +qed_goal "STL4_gen" TLA.thy "|- Init F --> Init G ==> |- []F --> []G" + (fn [prem] => [rtac (rewrite_rule [boxInitD] (prem RS STL4)) 1]); + +qed_goal "STL4E_gen" TLA.thy + "[| sigma |= []F; |- Init F --> Init G |] ==> sigma |= []G" + (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use STL4_gen]) 1) ]); (* see also STL4Edup below, which allows an auxiliary boxed formula: []A /\ F => G @@ -153,19 +165,19 @@ (* The dual versions for <> *) qed_goalw "DmdImpl" TLA.thy [dmd_def] - "(F .-> G) ==> (<>F .-> <>G)" - (fn [prem] => [fast_tac (temp_cs addSIs [int_mp prem] addSEs [STL4E]) 1]); + "|- F --> G ==> |- <>F --> <>G" + (fn [prem] => [fast_tac (temp_cs addSIs [prem] addSEs [STL4E]) 1]); qed_goal "DmdImplE" TLA.thy - "[| (sigma |= <>F); F .-> G |] ==> (sigma |= <>G)" - (fn prems => [ REPEAT (resolve_tac (prems @ [temp_mp DmdImpl]) 1) ]); + "[| sigma |= <>F; |- F --> G |] ==> sigma |= <>G" + (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use DmdImpl]) 1) ]); (* ------------------------ STL5 ------------------------------------------- *) -qed_goal "STL5" TLA.thy "([]F .& []G) .= ([](F .& G))" +qed_goal "STL5" TLA.thy "|- ([]F & []G) = ([](F & G))" (fn _ => [Auto_tac, - subgoal_tac "sigma |= [](G .-> (F .& G))" 1, - etac ((temp_mp normalT) RS mp) 1, atac 1, + subgoal_tac "sigma |= [](G --> (F & G))" 1, + etac (temp_use normalT) 1, atac 1, ALLGOALS (fast_tac (temp_cs addSEs [STL4E])) ]); (* rewrite rule to split conjunctions under boxes *) @@ -173,63 +185,65 @@ (* the corresponding elimination rule allows to combine boxes in the hypotheses (NB: F and G must have the same type, i.e., both actions or temporals.) + Use "addSE2" etc. if you want to add this to a claset, otherwise it will loop! *) qed_goal "box_conjE" TLA.thy - "[| (sigma |= []F); (sigma |= []G); (sigma |= [](F.&G)) ==> PROP R |] ==> PROP R" + "[| sigma |= []F; sigma |= []G; sigma |= [](F&G) ==> PROP R |] ==> PROP R" (fn prems => [ REPEAT (resolve_tac (prems @ [(temp_unlift STL5) RS iffD1, conjI]) 1) ]); +(* Instances of box_conjE for state predicates, actions, and temporals + in case the general rule is "too polymorphic". +*) +bind_thm("box_conjE_temp", read_instantiate [("'a","behavior")] box_conjE); +bind_thm("box_conjE_stp", read_instantiate [("'a","state")] box_conjE); +bind_thm("box_conjE_act", read_instantiate [("'a","state * state")] box_conjE); + (* Define a tactic that tries to merge all boxes in an antecedent. The definition is - a bit kludgy: how do you simulate "double elim-resolution"? - Note: If there are boxed hypotheses of different types, the tactic may delete the - wrong formulas. We therefore also define less polymorphic tactics for - temporals and actions. + a bit kludgy in order to simulate "double elim-resolution". *) -qed_goal "box_thin" TLA.thy "[| (sigma |= []F); PROP W |] ==> PROP W" - (fn prems => [resolve_tac prems 1]); + +Goal "[| sigma |= []F; PROP W |] ==> PROP W"; +by (atac 1); +val box_thin = result(); fun merge_box_tac i = REPEAT_DETERM (EVERY [etac box_conjE i, atac i, etac box_thin i]); -qed_goal "temp_box_conjE" TLA.thy - "[| (sigma |= [](F::temporal)); (sigma |= []G); (sigma |= [](F.&G)) ==> PROP R |] ==> PROP R" - (fn prems => [ REPEAT (resolve_tac - (prems @ [(temp_unlift STL5) RS iffD1, conjI]) 1) ]); -qed_goal "temp_box_thin" TLA.thy "[| (sigma |= [](F::temporal)); PROP W |] ==> PROP W" - (fn prems => [resolve_tac prems 1]); fun merge_temp_box_tac i = - REPEAT_DETERM (EVERY [etac temp_box_conjE i, atac i, etac temp_box_thin i]); + REPEAT_DETERM (EVERY [etac box_conjE_temp i, atac i, + eres_inst_tac [("'a","behavior")] box_thin i]); -qed_goal "act_box_conjE" TLA.thy - "[| (sigma |= [](A::action)); (sigma |= []B); (sigma |= [](A.&B)) ==> PROP R |] ==> PROP R" - (fn prems => [ REPEAT (resolve_tac - (prems @ [(temp_unlift STL5) RS iffD1, conjI]) 1) ]); -qed_goal "act_box_thin" TLA.thy "[| (sigma |= [](A::action)); PROP W |] ==> PROP W" - (fn prems => [resolve_tac prems 1]); +fun merge_stp_box_tac i = + REPEAT_DETERM (EVERY [etac box_conjE_stp i, atac i, + eres_inst_tac [("'a","state")] box_thin i]); + fun merge_act_box_tac i = - REPEAT_DETERM (EVERY [etac act_box_conjE i, atac i, etac act_box_thin i]); + REPEAT_DETERM (EVERY [etac box_conjE_act i, atac i, + eres_inst_tac [("'a","state * state")] box_thin i]); + (* rewrite rule to push universal quantification through box: - (sigma |= [](RALL x. F x)) = (! x. (sigma |= []F x)) + (sigma |= [](! x. F x)) = (! x. (sigma |= []F x)) *) bind_thm("all_box", standard((temp_unlift allT) RS sym)); -qed_goal "DmdOr" TLA.thy "(<>(F .| G)) .= (<>F .| <>G)" +qed_goal "DmdOr" TLA.thy "|- (<>(F | G)) = (<>F | <>G)" (fn _ => [auto_tac (temp_css addsimps2 [dmd_def,split_box_conj]), TRYALL (EVERY' [etac swap, merge_box_tac, fast_tac (temp_cs addSEs [STL4E])]) ]); -qed_goal "exT" TLA.thy "(REX x. <>(F x)) .= (<>(REX x. F x))" - (fn _ => [ auto_tac (temp_css addsimps2 [dmd_def,temp_rewrite Not_rex,all_box]) ]); +qed_goal "exT" TLA.thy "|- (? x. <>(F x)) = (<>(? x. F x))" + (fn _ => [ auto_tac (temp_css addsimps2 [dmd_def,Not_Rex,all_box]) ]); bind_thm("ex_dmd", standard((temp_unlift exT) RS sym)); qed_goal "STL4Edup" TLA.thy - "!!sigma. [| (sigma |= []A); (sigma |= []F); F .& []A .-> G |] ==> (sigma |= []G)" + "!!sigma. [| sigma |= []A; sigma |= []F; |- F & []A --> G |] ==> sigma |= []G" (fn _ => [etac dup_boxE 1, merge_box_tac 1, etac STL4E 1, @@ -237,7 +251,7 @@ ]); qed_goalw "DmdImpl2" TLA.thy [dmd_def] - "!!sigma. [| (sigma |= <>F); (sigma |= [](F .-> G)) |] ==> (sigma |= <>G)" + "!!sigma. [| sigma |= <>F; sigma |= [](F --> G) |] ==> sigma |= <>G" (fn _ => [Auto_tac, etac notE 1, merge_box_tac 1, @@ -245,41 +259,51 @@ ]); qed_goal "InfImpl" TLA.thy - "[| (sigma |= []<>F); (sigma |= []G); F .& G .-> H |] ==> (sigma |= []<>H)" + "[| sigma |= []<>F; sigma |= []G; |- F & G --> H |] ==> sigma |= []<>H" (fn [prem1,prem2,prem3] => [cut_facts_tac [prem1,prem2] 1, eres_inst_tac [("F","G")] dup_boxE 1, merge_box_tac 1, - fast_tac (temp_cs addSEs [STL4E,DmdImpl2] addSIs [int_mp prem3]) 1 + fast_tac (temp_cs addSEs [STL4E,DmdImpl2] addSIs [prem3]) 1 ]); (* ------------------------ STL6 ------------------------------------------- *) (* Used in the proof of STL6, but useful in itself. *) -qed_goalw "BoxDmdT" TLA.thy [dmd_def] "[]F .& <>G .-> <>([]F .& G)" - (fn _ => [ Auto_tac, +qed_goalw "BoxDmd" TLA.thy [dmd_def] "|- []F & <>G --> <>([]F & G)" + (fn _ => [ Clarsimp_tac 1, etac dup_boxE 1, merge_box_tac 1, etac swap 1, fast_tac (temp_cs addSEs [STL4E]) 1 ]); -bind_thm("BoxDmd", temp_conjmp BoxDmdT); (* weaker than BoxDmd, but more polymorphic (and often just right) *) -qed_goalw "BoxDmdT2" TLA.thy [dmd_def] "<>F .& []G .-> <>(F .& G)" - (fn _ => [ Auto_tac, +qed_goalw "BoxDmd_simple" TLA.thy [dmd_def] "|- []F & <>G --> <>(F & G)" + (fn _ => [ Clarsimp_tac 1, + merge_box_tac 1, + fast_tac (temp_cs addSEs [notE,STL4E]) 1 + ]); + +qed_goalw "BoxDmd2_simple" TLA.thy [dmd_def] "|- []F & <>G --> <>(G & F)" + (fn _ => [ Clarsimp_tac 1, merge_box_tac 1, fast_tac (temp_cs addSEs [notE,STL4E]) 1 ]); -qed_goal "STL6" TLA.thy "<>[]F .& <>[]G .-> <>[](F .& G)" +qed_goal "DmdImpldup" TLA.thy + "[| sigma |= []A; sigma |= <>F; |- []A & F --> G |] ==> sigma |= <>G" + (fn [p1,p2,p3] => [rtac ((p2 RS (p1 RS (temp_use BoxDmd))) RS DmdImplE) 1, + rtac p3 1]); + +qed_goal "STL6" TLA.thy "|- <>[]F & <>[]G --> <>[](F & G)" (fn _ => [auto_tac (temp_css addsimps2 [symmetric (temp_rewrite STL5)]), - etac (temp_conjimpE linT) 1, atac 1, etac thin_rl 1, + dtac (temp_use linT) 1, atac 1, etac thin_rl 1, rtac ((temp_unlift DmdDmd) RS iffD1) 1, etac disjE 1, - etac DmdImplE 1, rtac BoxDmdT 1, - (* the second subgoal needs commutativity of .&, which complicates the proof *) + etac DmdImplE 1, rtac BoxDmd 1, + (* the second subgoal needs commutativity of &, which complicates the proof *) etac DmdImplE 1, Auto_tac, - etac (temp_conjimpE BoxDmdT) 1, atac 1, etac thin_rl 1, + dtac (temp_use BoxDmd) 1, atac 1, etac thin_rl 1, fast_tac (temp_cs addSEs [DmdImplE]) 1 ]); @@ -287,55 +311,19 @@ (* ------------------------ True / False ----------------------------------------- *) section "Simplification of constants"; -qed_goal "BoxTrue" TLA.thy "[](#True)" - (fn _ => [ fast_tac (temp_cs addSIs [necT]) 1 ]); - -qed_goal "BoxTrue_simp" TLA.thy "([](#True)) .= #True" - (fn _ => [ fast_tac (temp_cs addSIs [BoxTrue RS tempD]) 1 ]); - -qed_goal "DmdFalse_simp" TLA.thy "(<>(#False)) .= #False" - (fn _ => [ auto_tac (temp_css addsimps2 [dmd_def, BoxTrue_simp]) ]); - -qed_goal "DmdTrue_simp" TLA.thy "(<>((#True)::temporal)) .= #True" - (fn _ => [ fast_tac (temp_cs addSIs [ImplDmdD]) 1 ]); - -qed_goal "DmdActTrue_simp" TLA.thy "(<>((#True)::action)) .= #True" - (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSIs2 [InitDmdD]) ]); - -qed_goal "BoxFalse_simp" TLA.thy "([]((#False)::temporal)) .= #False" - (fn _ => [ fast_tac (temp_cs addSDs [STL2D]) 1 ]); - -qed_goal "BoxActFalse_simp" TLA.thy "([]((#False)::action)) .= #False" - (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2bD]) ]); +qed_goal "BoxConst" TLA.thy "|- ([]#P) = #P" + (fn _ => [rtac tempI 1, + case_tac "P" 1, + auto_tac (temp_css addSIs2 [necT] addDs2 [STL2_gen] + addsimps2 Init_simps) + ]); -qed_goal "BoxConst_simp" TLA.thy "([]((#P)::temporal)) .= #P" - (fn _ => [rtac tempI 1, - case_tac "P" 1, - auto_tac (temp_css addsimps2 [BoxTrue_simp,BoxFalse_simp]) - ]); - -qed_goal "BoxActConst_simp" TLA.thy "([]((#P)::action)) .= #P" - (fn _ => [rtac tempI 1, - case_tac "P" 1, - auto_tac (temp_css addsimps2 [BoxTrue_simp,BoxActFalse_simp]) - ]); +qed_goalw "DmdConst" TLA.thy [dmd_def] "|- (<>#P) = #P" + (fn _ => [case_tac "P" 1, + ALLGOALS (asm_full_simp_tac (simpset() addsimps [BoxConst])) + ]); -qed_goal "DmdConst_simp" TLA.thy "(<>((#P)::temporal)) .= #P" - (fn _ => [rtac tempI 1, - case_tac "P" 1, - auto_tac (temp_css addsimps2 [DmdTrue_simp,DmdFalse_simp]) - ]); - -qed_goal "DmdActConst_simp" TLA.thy "(<>((#P)::action)) .= #P" - (fn _ => [rtac tempI 1, - case_tac "P" 1, - auto_tac (temp_css addsimps2 [DmdActTrue_simp,DmdFalse_simp]) - ]); - -val temp_simps = map temp_rewrite - [BoxTrue_simp,DmdFalse_simp,DmdTrue_simp, - DmdActTrue_simp, BoxFalse_simp, BoxActFalse_simp, - BoxConst_simp,BoxActConst_simp,DmdConst_simp,DmdActConst_simp]; +val temp_simps = map temp_rewrite [BoxConst, DmdConst]; (* Make these rewrites active by default *) Addsimps temp_simps; @@ -346,31 +334,31 @@ (* ------------------------ Further rewrites ----------------------------------------- *) section "Further rewrites"; -qed_goalw "NotBox" TLA.thy [dmd_def] "(.~[]F) .= (<>.~F)" - (fn _ => [ Auto_tac ]); +qed_goalw "NotBox" TLA.thy [dmd_def] "|- (~[]F) = (<>~F)" + (fn _ => [ Simp_tac 1 ]); -qed_goalw "NotDmd" TLA.thy [dmd_def] "(.~<>F) .= ([].~F)" - (fn _ => [ Auto_tac ]); +qed_goalw "NotDmd" TLA.thy [dmd_def] "|- (~<>F) = ([]~F)" + (fn _ => [ Simp_tac 1 ]); (* These are not by default included in temp_css, because they could be harmful, - e.g. []F .& .~[]F becomes []F .& <>.~F !! *) + e.g. []F & ~[]F becomes []F & <>~F !! *) val more_temp_simps = (map temp_rewrite [STL3, DmdDmd, NotBox, NotDmd]) @ (map (fn th => (temp_unlift th) RS eq_reflection) [NotBox, NotDmd]); -qed_goal "BoxDmdBox" TLA.thy "([]<>[]F) .= (<>[]F)" - (fn _ => [ auto_tac (temp_css addSDs2 [STL2D]), +qed_goal "BoxDmdBox" TLA.thy "|- ([]<>[]F) = (<>[]F)" + (fn _ => [ auto_tac (temp_css addSDs2 [STL2]), rtac ccontr 1, - subgoal_tac "sigma |= <>[][]F .& <>[].~[]F" 1, + subgoal_tac "sigma |= <>[][]F & <>[]~[]F" 1, etac thin_rl 1, Auto_tac, - etac (temp_conjimpE STL6) 1, atac 1, + dtac (temp_use STL6) 1, atac 1, Asm_full_simp_tac 1, ALLGOALS (asm_full_simp_tac (simpset() addsimps more_temp_simps)) ]); -qed_goalw "DmdBoxDmd" TLA.thy [dmd_def] "(<>[]<>F) .= ([]<>F)" - (fn _ => [auto_tac (temp_css addsimps2 [temp_rewrite (rewrite_rule [dmd_def] BoxDmdBox)])]); +qed_goalw "DmdBoxDmd" TLA.thy [dmd_def] "|- (<>[]<>F) = ([]<>F)" + (fn _ => [ auto_tac (temp_css addsimps2 [rewrite_rule [dmd_def] BoxDmdBox]) ]); val more_temp_simps = more_temp_simps @ (map temp_rewrite [BoxDmdBox, DmdBoxDmd]); @@ -378,31 +366,25 @@ (* ------------------------ Miscellaneous ----------------------------------- *) qed_goal "BoxOr" TLA.thy - "!!sigma. [| (sigma |= []F .| []G) |] ==> (sigma |= [](F .| G))" + "!!sigma. [| sigma |= []F | []G |] ==> sigma |= [](F | G)" (fn _ => [ fast_tac (temp_cs addSEs [STL4E]) 1 ]); -qed_goal "DBImplBD" TLA.thy "<>[](F::temporal) .-> []<>F" - (fn _ => [Auto_tac, +(* "persistently implies infinitely often" *) +qed_goal "DBImplBD" TLA.thy "|- <>[]F --> []<>F" + (fn _ => [Clarsimp_tac 1, rtac ccontr 1, - old_auto_tac (temp_css addsimps2 more_temp_simps - addEs2 [temp_conjimpE STL6]) - ]); - -(* Although the script is the same, the derivation isn't polymorphic and doesn't - work for other types of formulas (uses STL2). -*) -qed_goal "DBImplBDAct" TLA.thy "<>[](A::action) .-> []<>A" - (fn _ => [Auto_tac, - rtac ccontr 1, - old_auto_tac (temp_css addsimps2 more_temp_simps addEs2 [temp_conjimpE STL6]) + asm_full_simp_tac (simpset() addsimps more_temp_simps) 1, + dtac (temp_use STL6) 1, atac 1, + Asm_full_simp_tac 1 ]); qed_goal "BoxDmdDmdBox" TLA.thy - "!!sigma. [| (sigma |= []<>F); (sigma |= <>[]G) |] ==> (sigma |= []<>(F .& G))" - (fn _ => [rtac ccontr 1, + "|- []<>F & <>[]G --> []<>(F & G)" + (fn _ => [Clarsimp_tac 1, + rtac ccontr 1, rewrite_goals_tac more_temp_simps, - etac (temp_conjimpE STL6) 1, atac 1, - subgoal_tac "sigma |= <>[].~F" 1, + dtac (temp_use STL6) 1, atac 1, + subgoal_tac "sigma |= <>[]~F" 1, force_tac (temp_css addsimps2 [dmd_def]) 1, fast_tac (temp_cs addEs [DmdImplE,STL4E]) 1 ]); @@ -414,65 +396,78 @@ section "priming"; (* ------------------------ TLA2 ------------------------------------------- *) -qed_goal "STL2bD_pr" TLA.thy - "!!sigma. (sigma |= []P) ==> (sigma |= Init(P .& P`))" - (fn _ => [rewrite_goals_tac Init_simps, - fast_tac (temp_cs addSIs [temp_mp primeI, STL2bD]) 1]); +qed_goal "STL2_pr" TLA.thy + "|- []P --> Init P & Init P`" + (fn _ => [fast_tac (temp_cs addSIs [primeI, STL2_gen]) 1]); (* Auxiliary lemma allows priming of boxed actions *) -qed_goal "BoxPrime" TLA.thy "[]P .-> [](P .& P`)" - (fn _ => [Auto_tac, +qed_goal "BoxPrime" TLA.thy "|- []P --> []($P & P$)" + (fn _ => [Clarsimp_tac 1, etac dup_boxE 1, - auto_tac (temp_css addsimps2 [boxact_def] - addSIs2 [STL2bD_pr] addSEs2 [STL4E]) + rewtac boxInit_act, + etac STL4E 1, + auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_pr]) ]); -qed_goal "TLA2" TLA.thy "P .& P` .-> Q ==> []P .-> []Q" - (fn prems => [fast_tac (temp_cs addSIs prems addDs [temp_mp BoxPrime] addEs [STL4E]) 1]); +qed_goal "TLA2" TLA.thy "|- $P & P$ --> A ==> |- []P --> []A" + (fn prems => [Clarsimp_tac 1, + dtac (temp_use BoxPrime) 1, + auto_tac (temp_css addsimps2 [Init_stp_act_rev] addSIs2 prems addSEs2 [STL4E]) + ]); qed_goal "TLA2E" TLA.thy - "[| (sigma |= []P); P .& P` .-> Q |] ==> (sigma |= []Q)" - (fn prems => [REPEAT (resolve_tac (prems @ (prems RL [temp_mp TLA2])) 1)]); + "[| sigma |= []P; |- $P & P$ --> A |] ==> sigma |= []A" + (fn prems => [REPEAT (resolve_tac (prems @ (prems RL [temp_use TLA2])) 1)]); -qed_goalw "DmdPrime" TLA.thy [dmd_def] "(<>P`) .-> (<>P)" +qed_goalw "DmdPrime" TLA.thy [dmd_def] "|- (<>P`) --> (<>P)" (fn _ => [ fast_tac (temp_cs addSEs [TLA2E]) 1 ]); +bind_thm("PrimeDmd", (temp_use InitDmd_gen) RS (temp_use DmdPrime)); (* ------------------------ INV1, stable --------------------------------------- *) section "stable, invariant"; qed_goal "ind_rule" TLA.thy - "[| (sigma |= []H); (sigma |= Init(P)); H .-> (Init(P) .& .~[]F .-> Init(P`) .& F) |] \ -\ ==> (sigma |= []F)" - (fn prems => [rtac ((temp_mp indT) RS mp) 1, + "[| sigma |= []H; sigma |= Init P; |- H --> (Init P & ~[]F --> Init(P`) & F) |] \ +\ ==> sigma |= []F" + (fn prems => [rtac (temp_use indT) 1, REPEAT (resolve_tac (prems @ (prems RL [STL4E])) 1)]); - + +qed_goalw "box_stp_act" TLA.thy [boxInit_act] "|- ([]$P) = ([]P)" + (K [simp_tac (simpset() addsimps Init_simps) 1]); +bind_thm("box_stp_actI", zero_var_indexes ((temp_use box_stp_act) RS iffD2)); +bind_thm("box_stp_actD", zero_var_indexes ((temp_use box_stp_act) RS iffD1)); -qed_goalw "INV1" TLA.thy [stable_def,boxact_def] - "Init(P) .& stable(P) .-> []P" - (K [force_tac (temp_css addsimps2 Init_simps addEs2 [ind_rule]) 1]); -bind_thm("INV1I", temp_conjmp INV1); +val more_temp_simps = (temp_rewrite box_stp_act)::more_temp_simps; -qed_goalw "StableL" TLA.thy [stable_def] - "(P .& A .-> P`) ==> ([]A .-> stable(P))" - (fn [prem] => [fast_tac (temp_cs addSIs [action_mp prem] addSEs [STL4E]) 1]); +qed_goalw "INV1" TLA.thy [stable_def,boxInit_stp,boxInit_act] + "|- (Init P) --> (stable P) --> []P" + (K [Clarsimp_tac 1, + etac ind_rule 1, + auto_tac (temp_css addsimps2 Init_simps addEs2 [ind_rule]) + ]); + +qed_goalw "StableT" TLA.thy [stable_def] + "|- $P & A --> P` ==> |- []A --> stable P" + (fn [prem] => [fast_tac (temp_cs addSEs [STL4E] addIs [prem]) 1]); qed_goal "Stable" TLA.thy - "[| (sigma |= []A); P .& A .-> P` |] ==> (sigma |= stable P)" - (fn prems => [ REPEAT (resolve_tac (prems @ [temp_mp StableL]) 1) ]); + "[| sigma |= []A; |- $P & A --> P` |] ==> sigma |= stable P" + (fn prems => [ REPEAT (resolve_tac (prems @ [temp_use StableT]) 1) ]); (* Generalization of INV1 *) qed_goalw "StableBox" TLA.thy [stable_def] - "!!sigma. (sigma |= stable P) ==> (sigma |= [](Init P .-> []P))" - (K [etac dup_boxE 1, - force_tac (temp_css addsimps2 [stable_def] addEs2 [STL4E, INV1I]) 1]); - -(* useful for WF2 / SF2 *) + "|- (stable P) --> [](Init P --> []P)" + (K [Clarsimp_tac 1, + etac dup_boxE 1, + force_tac (temp_css addsimps2 [stable_def] addEs2 [STL4E, INV1]) 1]); + qed_goal "DmdStable" TLA.thy - "!!sigma. [| (sigma |= stable P); (sigma |= <>P) |] ==> (sigma |= <>[]P)" - (fn _ => [rtac DmdImpl2 1, - etac StableBox 2, - auto_tac (temp_css addsimps2 [DmdAct]) + "|- (stable P) & <>P --> <>[]P" + (fn _ => [Clarsimp_tac 1, + rtac DmdImpl2 1, + etac (temp_use StableBox) 2, + asm_simp_tac (simpset() addsimps [dmdInitD]) 1 ]); (* ---------------- (Semi-)automatic invariant tactics ---------------------- *) @@ -481,11 +476,11 @@ fun inv_tac css = SELECT_GOAL (EVERY [auto_tac css, TRY (merge_box_tac 1), - rtac INV1I 1, (* fail if the goal is not a box *) + rtac (temp_use INV1) 1, (* fail if the goal is not a box *) TRYALL (etac Stable)]); (* auto_inv_tac applies inv_tac and then tries to attack the subgoals; - in simple cases it may be able to handle goals like MyProg .-> []Inv. + in simple cases it may be able to handle goals like |- MyProg --> []Inv. In these simple cases the simplifier seems to be more useful than the auto-tactic, which applies too much propositional logic and simplifies too late. @@ -493,13 +488,14 @@ fun auto_inv_tac ss = SELECT_GOAL ((inv_tac (claset(),ss) 1) THEN - (TRYALL (action_simp_tac (ss addsimps [Init_def,square_def]) [] []))); + (TRYALL (action_simp_tac (ss addsimps [Init_stp,Init_act]) [] [squareE]))); qed_goalw "unless" TLA.thy [dmd_def] - "!!sigma. (sigma |= [](P .-> P` .| Q`)) ==> (sigma |= stable P .| <>Q`)" - (fn _ => [action_simp_tac (simpset()) [disjCI] [] 1, + "|- []($P --> P` | Q`) --> (stable P) | <>Q" + (fn _ => [clarsimp_tac (temp_css addSDs2 [BoxPrime]) 1, merge_box_tac 1, + etac swap 1, fast_tac (temp_cs addSEs [Stable]) 1 ]); @@ -507,73 +503,72 @@ (* --------------------- Recursive expansions --------------------------------------- *) section "recursive expansions"; -(* Recursive expansions of [] and <>, restricted to state predicates to avoid looping *) -qed_goal "BoxRec" TLA.thy "([]$P) .= (Init($P) .& ([]P$))" - (fn _ => [auto_tac (temp_css addSIs2 [STL2bD]), +(* Recursive expansions of [] and <> for state predicates *) +qed_goal "BoxRec" TLA.thy "|- ([]P) = (Init P & []P`)" + (fn _ => [auto_tac (temp_css addSIs2 [STL2_gen]), fast_tac (temp_cs addSEs [TLA2E]) 1, - auto_tac (temp_css addsimps2 [stable_def] addSEs2 [INV1I,STL4E]) + auto_tac (temp_css addsimps2 [stable_def] addSEs2 [INV1,STL4E]) ]); -qed_goalw "DmdRec" TLA.thy [dmd_def] "(<>$P) .= (Init($P) .| (<>P$))" (K [ - Auto_tac, - etac notE 1, - force_tac (temp_css addsimps2 (stable_def::Init_simps) - addIs2 [INV1I] addEs2 [STL4E]) 1, - force_tac (temp_css addsimps2 Init_simps addSDs2 [STL2bD]) 1, - fast_tac (temp_cs addSEs [notE,TLA2E]) 1 - ]); +qed_goalw "DmdRec" TLA.thy [dmd_def, temp_rewrite BoxRec] "|- (<>P) = (Init P | <>P`)" + (K [ auto_tac (temp_css addsimps2 Init_simps) ]); qed_goal "DmdRec2" TLA.thy - "!!sigma. [| (sigma |= <>($P)); (sigma |= [](.~P$)) |] ==> (sigma |= Init($P))" - (K [ dtac ((temp_unlift DmdRec) RS iffD1) 1, - force_tac (temp_css addsimps2 [dmd_def]) 1]); + "!!sigma. [| sigma |= <>P; sigma |= []~P` |] ==> sigma |= Init P" + (K [ force_tac (temp_css addsimps2 [DmdRec,dmd_def]) 1]); -(* The "=>" part of the following is a little intricate. *) -qed_goal "InfinitePrime" TLA.thy "([]<>$P) .= ([]<>P$)" +(* The "-->" part of the following is a little intricate. *) +qed_goal "InfinitePrime" TLA.thy "|- ([]<>P) = ([]<>P`)" (fn _ => [Auto_tac, rtac classical 1, - rtac (temp_mp DBImplBDAct) 1, - subgoal_tac "sigma |= <>[]$P" 1, + rtac (temp_use DBImplBD) 1, + subgoal_tac "sigma |= <>[]P" 1, fast_tac (temp_cs addSEs [DmdImplE,TLA2E]) 1, - subgoal_tac "sigma |= <>[](<>$P .& [].~P$)" 1, - force_tac (temp_css addsimps2 [boxact_def] + subgoal_tac "sigma |= <>[](<>P & []~P`)" 1, + force_tac (temp_css addsimps2 [boxInit_stp] addSEs2 [DmdImplE,STL4E,DmdRec2]) 1, - force_tac (temp_css addSIs2 [temp_mp STL6] addsimps2 more_temp_simps) 1, - fast_tac (temp_cs addIs [temp_mp DmdPrime] addSEs [STL4E]) 1 + force_tac (temp_css addSIs2 [STL6] addsimps2 more_temp_simps) 1, + fast_tac (temp_cs addIs [DmdPrime] addSEs [STL4E]) 1 ]); +qed_goal "InfiniteEnsures" TLA.thy + "[| sigma |= []N; sigma |= []<>A; |- A & N --> P` |] ==> sigma |= []<>P" + (fn prems => [rewtac (temp_rewrite InfinitePrime), + rtac InfImpl 1, + REPEAT (resolve_tac prems 1) + ]); + (* ------------------------ fairness ------------------------------------------- *) section "fairness"; (* alternative definitions of fairness *) qed_goalw "WF_alt" TLA.thy [WF_def,dmd_def] - "WF(A)_v .= (([]<>.~$(Enabled(_v))) .| []<>_v)" + "|- WF(A)_v = ([]<>~Enabled(_v) | []<>_v)" (fn _ => [ fast_tac temp_cs 1 ]); qed_goalw "SF_alt" TLA.thy [SF_def,dmd_def] - "SF(A)_v .= ((<>[].~$(Enabled(_v))) .| []<>_v)" + "|- SF(A)_v = (<>[]~Enabled(_v) | []<>_v)" (fn _ => [ fast_tac temp_cs 1 ]); (* theorems to "box" fairness conditions *) -qed_goal "BoxWFI" TLA.thy - "!!sigma. (sigma |= WF(A)_v) ==> (sigma |= []WF(A)_v)" - (fn _ => [ auto_tac (temp_css addsimps2 (temp_rewrite WF_alt::more_temp_simps) addSIs2 [BoxOr]) ]); +qed_goal "BoxWFI" TLA.thy "|- WF(A)_v --> []WF(A)_v" + (fn _ => [ auto_tac (temp_css addsimps2 (WF_alt::more_temp_simps) + addSIs2 [BoxOr]) ]); -qed_goal "WF_Box" TLA.thy "([]WF(A)_v) .= WF(A)_v" - (fn prems => [ fast_tac (temp_cs addSIs [BoxWFI] addSDs [STL2D]) 1 ]); +qed_goal "WF_Box" TLA.thy "|- ([]WF(A)_v) = WF(A)_v" + (fn prems => [ fast_tac (temp_cs addSIs [BoxWFI] addSDs [STL2]) 1 ]); -qed_goal "BoxSFI" TLA.thy - "!!sigma. (sigma |= SF(A)_v) ==> (sigma |= []SF(A)_v)" - (fn _ => [ auto_tac (temp_css addsimps2 (temp_rewrite SF_alt::more_temp_simps) addSIs2 [BoxOr]) ]); +qed_goal "BoxSFI" TLA.thy "|- SF(A)_v --> []SF(A)_v" + (fn _ => [ auto_tac (temp_css addsimps2 (SF_alt::more_temp_simps) + addSIs2 [BoxOr]) ]); -qed_goal "SF_Box" TLA.thy "([]SF(A)_v) .= SF(A)_v" - (fn prems => [ fast_tac (temp_cs addSIs [BoxSFI] addSDs [STL2D]) 1 ]); +qed_goal "SF_Box" TLA.thy "|- ([]SF(A)_v) = SF(A)_v" + (fn prems => [ fast_tac (temp_cs addSIs [BoxSFI] addSDs [STL2]) 1 ]); val more_temp_simps = more_temp_simps @ (map temp_rewrite [WF_Box, SF_Box]); -qed_goalw "SFImplWF" TLA.thy [SF_def,WF_def] - "!!sigma. (sigma |= SF(A)_v) ==> (sigma |= WF(A)_v)" - (fn _ => [ fast_tac (temp_cs addSDs [temp_mp DBImplBDAct]) 1 ]); +qed_goalw "SFImplWF" TLA.thy [SF_def,WF_def] "|- SF(A)_v --> WF(A)_v" + (fn _ => [ fast_tac (temp_cs addSDs [DBImplBD]) 1 ]); (* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *) val box_fair_tac = SELECT_GOAL (REPEAT (dresolve_tac [BoxWFI,BoxSFI] 1)); @@ -583,313 +578,321 @@ section "~>"; -qed_goalw "leadsto_init" TLA.thy [leadsto] - "!!sigma. [| (sigma |= Init P); (sigma |= P ~> Q) |] ==> (sigma |= <>Q)" - (fn _ => [ fast_tac (temp_cs addSDs [temp_mp STL2]) 1 ]); +qed_goalw "leadsto_init" TLA.thy [leadsto_def] + "|- (Init F) & (F ~> G) --> <>G" + (fn _ => [ auto_tac (temp_css addSDs2 [STL2]) ]); -qed_goalw "streett_leadsto" TLA.thy [leadsto] - "([]<>P .-> []<>Q) .= (<>(P ~> Q))" (K [ +(* |- F & (F ~> G) --> <>G *) +bind_thm("leadsto_init_temp", + rewrite_rule Init_simps (read_instantiate [("'a","behavior")] leadsto_init)); + +qed_goalw "streett_leadsto" TLA.thy [leadsto_def] + "|- ([]<>Init F --> []<>G) = (<>(F ~> G))" (K [ Auto_tac, - asm_full_simp_tac (simpset() addsimps boxact_def::more_temp_simps) 1, - force_tac (temp_css addSEs2 [DmdImplE,STL4E] - addsimps2 Init_simps) 1, - force_tac (temp_css addSIs2 [ImplDmdD] addSEs2 [STL4E]) 1, - subgoal_tac "sigma |= []<><>Q" 1, asm_full_simp_tac (simpset() addsimps more_temp_simps) 1, - rewtac (temp_rewrite DmdAct), - dtac BoxDmdDmdBox 1, atac 1, - auto_tac (temp_css addSEs2 [DmdImplE,STL4E]) + fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1, + fast_tac (temp_cs addSIs [InitDmd] addSEs [STL4E]) 1, + subgoal_tac "sigma |= []<><>G" 1, + asm_full_simp_tac (simpset() addsimps more_temp_simps) 1, + dtac (temp_use BoxDmdDmdBox) 1, atac 1, + fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1 ]); qed_goal "leadsto_infinite" TLA.thy - "!!sigma. [| (sigma |= []<>P); (sigma |= P ~> Q) |] ==> (sigma |= []<>Q)" - (fn _ => [rtac ((temp_unlift streett_leadsto) RS iffD2 RS mp) 1, - auto_tac (temp_css addSIs2 [ImplDmdD]) + "|- []<>F & (F ~> G) --> []<>G" + (fn _ => [Clarsimp_tac 1, + etac ((temp_use InitDmd) RS + ((temp_unlift streett_leadsto) RS iffD2 RS mp)) 1, + asm_simp_tac (simpset() addsimps [dmdInitD]) 1 ]); (* In particular, strong fairness is a Streett condition. The following rules are sometimes easier to use than WF2 or SF2 below. *) qed_goalw "leadsto_SF" TLA.thy [SF_def] - "!!sigma. (sigma |= $(Enabled(_v)) ~> _v) ==> sigma |= SF(A)_v" - (K [REPEAT(step_tac temp_cs 1), - rtac leadsto_infinite 1, - ALLGOALS atac]); + "|- (Enabled(_v) ~> _v) --> SF(A)_v" + (K [clarsimp_tac (temp_css addSEs2 [leadsto_infinite]) 1]); -bind_thm("leadsto_WF", leadsto_SF RS SFImplWF); +qed_goal "leadsto_WF" TLA.thy + "|- (Enabled(_v) ~> _v) --> WF(A)_v" + (K [ clarsimp_tac (temp_css addSIs2 [SFImplWF, leadsto_SF]) 1 ]); (* introduce an invariant into the proof of a leadsto assertion. - []I => ((P ~> Q) = (P /\ I ~> Q)) + []I --> ((P ~> Q) = (P /\ I ~> Q)) *) -qed_goalw "INV_leadsto" TLA.thy [leadsto] - "!!sigma. [| (sigma |= []I); (sigma |= (P .& I) ~> Q) |] ==> (sigma |= P ~> Q)" - (fn _ => [etac STL4Edup 1, atac 1, - auto_tac (temp_css addsimps2 [Init_def] addSDs2 [STL2bD]) +qed_goalw "INV_leadsto" TLA.thy [leadsto_def] + "|- []I & (P & I ~> Q) --> (P ~> Q)" + (fn _ => [Clarsimp_tac 1, + etac STL4Edup 1, atac 1, + auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_gen]) ]); -qed_goalw "leadsto_classical" TLA.thy [leadsto,dmd_def] - "!!sigma. (sigma |= [](Init P .& [].~Q .-> <>Q)) ==> (sigma |= P ~> Q)" - (fn _ => [fast_tac (temp_cs addSEs [STL4E]) 1]); +qed_goalw "leadsto_classical" TLA.thy [leadsto_def,dmd_def] + "|- (Init F & []~G ~> G) --> (F ~> G)" + (fn _ => [force_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) 1]); -qed_goalw "leadsto_false" TLA.thy [leadsto] - "(P ~> #False) .= ([] .~P)" - (fn _ => [ auto_tac (temp_css addsimps2 boxact_def::Init_simps) ]); +qed_goalw "leadsto_false" TLA.thy [leadsto_def] + "|- (F ~> #False) = ([]~F)" + (fn _ => [ simp_tac (simpset() addsimps [boxNotInitD]) 1 ]); + +qed_goalw "leadsto_exists" TLA.thy [leadsto_def] + "|- ((? x. F x) ~> G) = (!x. (F x ~> G))" + (K [auto_tac (temp_css addsimps2 allT::Init_simps addSEs2 [STL4E])]); + (* basic leadsto properties, cf. Unity *) -qed_goal "ImplLeadsto" TLA.thy - "!!sigma. (sigma |= [](P .-> Q)) ==> (sigma |= (P ~> Q))" - (fn _ => [etac INV_leadsto 1, rewtac leadsto, - rtac (temp_unlift necT) 1, - auto_tac (temp_css addSIs2 [InitDmdD] addsimps2 [Init_def]) +qed_goalw "ImplLeadsto_gen" TLA.thy [leadsto_def] + "|- [](Init F --> Init G) --> (F ~> G)" + (fn _ => [auto_tac (temp_css addSIs2 [InitDmd_gen] + addSEs2 [STL4E_gen] addsimps2 Init_simps) ]); -qed_goal "EnsuresLeadsto" TLA.thy - "A .& P .-> Q` ==> []A .-> (P ~> Q)" (fn [prem] => [ - auto_tac (temp_css addSEs2 [INV_leadsto]), - rewtac leadsto, - auto_tac (temp_css addSIs2 [temp_unlift necT]), - rtac (temp_mp DmdPrime) 1, - rtac InitDmdD 1, - force_tac (action_css addsimps2 [Init_def] - addSIs2 [action_mp prem]) 1]); +bind_thm("ImplLeadsto", + rewrite_rule Init_simps + (read_instantiate [("'a","behavior"), ("'b","behavior")] ImplLeadsto_gen)); + +qed_goal "ImplLeadsto_simple" TLA.thy + "|- F --> G ==> |- F ~> G" + (fn [prem] => [auto_tac (temp_css addsimps2 [Init_def] + addSIs2 [ImplLeadsto_gen,necT,prem])]); + +qed_goalw "EnsuresLeadsto" TLA.thy [leadsto_def] + "|- A & $P --> Q` ==> |- []A --> (P ~> Q)" (fn [prem] => [ + clarsimp_tac (temp_css addSEs2 [INV_leadsto]) 1, + etac STL4E_gen 1, + auto_tac (temp_css addsimps2 Init_defs + addSIs2 [PrimeDmd, prem]) + ]); + +qed_goalw "EnsuresLeadsto2" TLA.thy [leadsto_def] + "|- []($P --> Q`) --> (P ~> Q)" + (fn _ => [Clarsimp_tac 1, + etac STL4E_gen 1, + auto_tac (temp_css addsimps2 Init_simps addSIs2 [PrimeDmd]) + ]); -qed_goalw "EnsuresLeadsto2" TLA.thy [leadsto] - "!!sigma. sigma |= [](P .-> Q`) ==> sigma |= P ~> Q" - (fn _ => [subgoal_tac "sigma |= []Init(P .-> Q`)" 1, - etac STL4E 1, - ALLGOALS (force_tac (temp_css addsimps2 boxact_def::Init_simps - addIs2 [(temp_mp InitDmd) RS (temp_mp DmdPrime)]))]); - +qed_goalw "ensures" TLA.thy [leadsto_def] + "[| |- $P & N --> P` | Q`; \ +\ |- ($P & N) & A --> Q` \ +\ |] ==> |- []N & []([]P --> <>A) --> (P ~> Q)" + (fn [p1,p2] => [Clarsimp_tac 1, + etac STL4Edup 1, atac 1, + Clarsimp_tac 1, + subgoal_tac "sigmaa |= []($P --> P` | Q`)" 1, + dtac (temp_use unless) 1, + clarsimp_tac (temp_css addSDs2 [INV1]) 1, + rtac ((temp_use (p2 RS DmdImpl)) RS (temp_use DmdPrime)) 1, + force_tac (temp_css addSIs2 [BoxDmd_simple] + addsimps2 [split_box_conj,box_stp_act]) 1, + force_tac (temp_css addEs2 [STL4E] addDs2 [p1]) 1 + ]); + +qed_goal "ensures_simple" TLA.thy + "[| |- $P & N --> P` | Q`; \ +\ |- ($P & N) & A --> Q` \ +\ |] ==> |- []N & []<>A --> (P ~> Q)" + (fn prems => [Clarsimp_tac 1, + rtac (temp_use ensures) 1, + TRYALL (ares_tac prems), + force_tac (temp_css addSEs2 [STL4E]) 1 + ]); + qed_goal "EnsuresInfinite" TLA.thy - "[| (sigma |= []<>P); (sigma |= []A); A .& P .-> Q` |] ==> (sigma |= []<>Q)" - (fn prems => [REPEAT (resolve_tac (prems @ [leadsto_infinite, - temp_mp EnsuresLeadsto]) 1)]); + "[| sigma |= []<>P; sigma |= []A; |- A & $P --> Q` |] ==> sigma |= []<>Q" + (fn prems => [REPEAT (resolve_tac (prems @ [temp_use leadsto_infinite, + temp_use EnsuresLeadsto]) 1)]); + (*** Gronning's lattice rules (taken from TLP) ***) section "Lattice rules"; -qed_goalw "LatticeReflexivity" TLA.thy [leadsto] "F ~> F" - (fn _ => [REPEAT (resolve_tac [necT,InitDmd] 1)]); +qed_goalw "LatticeReflexivity" TLA.thy [leadsto_def] "|- F ~> F" + (fn _ => [REPEAT (resolve_tac [necT,InitDmd_gen] 1)]); -qed_goalw "LatticeTransitivity" TLA.thy [leadsto] - "!!sigma. [| (sigma |= G ~> H); (sigma |= F ~> G) |] ==> (sigma |= F ~> H)" - (fn _ => [etac dup_boxE 1, (* [][](Init G .-> H) *) +qed_goalw "LatticeTransitivity" TLA.thy [leadsto_def] + "|- (G ~> H) & (F ~> G) --> (F ~> H)" + (fn _ => [Clarsimp_tac 1, + etac dup_boxE 1, (* [][](Init G --> H) *) merge_box_tac 1, - auto_tac (temp_css addSEs2 [STL4E]), - rewtac (temp_rewrite DmdAct), - subgoal_tac "sigmaa |= <><> Init H" 1, - asm_full_simp_tac (simpset() addsimps more_temp_simps) 1, - fast_tac (temp_cs addSEs [DmdImpl2]) 1 + clarsimp_tac (temp_css addSEs2 [STL4E]) 1, + rtac dup_dmdD 1, + subgoal_tac "sigmaa |= <>Init G" 1, + etac DmdImpl2 1, atac 1, + asm_simp_tac (simpset() addsimps [dmdInitD]) 1 ]); -qed_goalw "LatticeDisjunctionElim1" TLA.thy [leadsto] - "!!sigma. (sigma |= (F .| G) ~> H) ==> (sigma |= F ~> H)" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def] addSEs2 [STL4E]) ]); +qed_goalw "LatticeDisjunctionElim1" TLA.thy [leadsto_def] + "|- (F | G ~> H) --> (F ~> H)" + (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) ]); -qed_goalw "LatticeDisjunctionElim2" TLA.thy [leadsto] - "!!sigma. (sigma |= (F .| G) ~> H) ==> (sigma |= G ~> H)" - (fn _ => [ auto_tac (temp_css addsimps2 [Init_def] addSEs2 [STL4E]) ]); +qed_goalw "LatticeDisjunctionElim2" TLA.thy [leadsto_def] + "|- (F | G ~> H) --> (G ~> H)" + (fn _ => [ auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) ]); -qed_goalw "LatticeDisjunctionIntro" TLA.thy [leadsto] - "!!sigma. [| (sigma |= F ~> H); (sigma |= G ~> H) |] ==> (sigma |= (F .| G) ~> H)" - (fn _ => [merge_box_tac 1, - auto_tac (temp_css addsimps2 [Init_def] addSEs2 [STL4E]) +qed_goalw "LatticeDisjunctionIntro" TLA.thy [leadsto_def] + "|- (F ~> H) & (G ~> H) --> (F | G ~> H)" + (fn _ => [Clarsimp_tac 1, + merge_box_tac 1, + auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) ]); +qed_goal "LatticeDisjunction" TLA.thy + "|- (F | G ~> H) = ((F ~> H) & (G ~> H))" + (fn _ => [auto_tac (temp_css addIs2 [LatticeDisjunctionIntro, + LatticeDisjunctionElim1, LatticeDisjunctionElim2])]); + qed_goal "LatticeDiamond" TLA.thy - "!!sigma. [| (sigma |= B ~> D); (sigma |= A ~> (B .| C)); (sigma |= C ~> D) |] \ -\ ==> (sigma |= A ~> D)" - (fn _ => [subgoal_tac "sigma |= (B .| C) ~> D" 1, - eres_inst_tac [("G", "B .| C")] LatticeTransitivity 1, + "|- (A ~> B | C) & (B ~> D) & (C ~> D) --> (A ~> D)" + (fn _ => [Clarsimp_tac 1, + subgoal_tac "sigma |= (B | C) ~> D" 1, + eres_inst_tac [("G", "LIFT (B | C)")] (temp_use LatticeTransitivity) 1, ALLGOALS (fast_tac (temp_cs addSIs [LatticeDisjunctionIntro])) ]); qed_goal "LatticeTriangle" TLA.thy - "!!sigma. [| (sigma |= B ~> D); (sigma |= A ~> (B .| D)) |] ==> (sigma |= A ~> D)" - (fn _ => [subgoal_tac "sigma |= (B .| D) ~> D" 1, - eres_inst_tac [("G", "B .| D")] LatticeTransitivity 1, atac 1, - auto_tac (temp_css addSIs2 [LatticeDisjunctionIntro] addIs2 [ImplLeadsto]) + "|- (A ~> D | B) & (B ~> D) --> (A ~> D)" + (fn _ => [Clarsimp_tac 1, + subgoal_tac "sigma |= (D | B) ~> D" 1, + eres_inst_tac [("G", "LIFT (D | B)")] (temp_use LatticeTransitivity) 1, atac 1, + auto_tac (temp_css addSIs2 [LatticeDisjunctionIntro] + addIs2 [LatticeReflexivity]) + ]); + +qed_goal "LatticeTriangle2" TLA.thy + "|- (A ~> B | D) & (B ~> D) --> (A ~> D)" + (fn _ => [Clarsimp_tac 1, + subgoal_tac "sigma |= B | D ~> D" 1, + eres_inst_tac [("G", "LIFT (B | D)")] (temp_use LatticeTransitivity) 1, atac 1, + auto_tac (temp_css addSIs2 [LatticeDisjunctionIntro] + addIs2 [LatticeReflexivity]) ]); (*** Lamport's fairness rules ***) section "Fairness rules"; -qed_goalw "WF1" TLA.thy [leadsto] - "[| P .& N .-> P` .| Q`; \ -\ P .& N .& _v .-> Q`; \ -\ P .& N .-> $(Enabled(_v)) |] \ -\ ==> []N .& WF(A)_v .-> (P ~> Q)" (fn [prem1,prem2,prem3] => [ - auto_tac (temp_css addSDs2 [BoxWFI]), - etac STL4Edup 1, atac 1, - Auto_tac, - subgoal_tac "sigmaa |= [](P .-> P` .| Q`)" 1, - auto_tac (temp_css addSDs2 [unless]), - etac (temp_conjimpE INV1) 1, atac 1, - merge_box_tac 1, - rtac STL2D 1, - rtac EnsuresInfinite 1, atac 2, - SELECT_GOAL (old_auto_tac (temp_css addsimps2 [WF_alt])) 1, - atac 2, - subgoal_tac "sigmaa |= [][]$(Enabled(_v))" 1, - merge_box_tac 1, - force_tac (temp_css addsimps2 [dmd_def]) 1, - SELECT_GOAL (rewtac (temp_rewrite STL3)) 1, - force_tac (temp_css addSEs2 [STL4E] addSIs2 [action_mp prem3]) 1, - fast_tac (action_cs addSIs [action_mp prem2]) 1, - fast_tac (temp_cs addIs [temp_mp DmdPrime]) 1, - fast_tac (temp_cs addSEs [STL4E] addSIs [action_mp prem1]) 1 - ]); +qed_goal "WF1" TLA.thy + "[| |- $P & N --> P` | Q`; \ +\ |- ($P & N) & _v --> Q`; \ +\ |- $P & N --> $(Enabled(_v)) |] \ +\ ==> |- []N & WF(A)_v --> (P ~> Q)" (fn prems => [ + clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1, + rtac (temp_use ensures) 1, + TRYALL (ares_tac prems), + etac STL4Edup 1, atac 1, + clarsimp_tac (temp_css addsimps2 [WF_def]) 1, + rtac (temp_use STL2) 1, + clarsimp_tac (temp_css addSEs2 [mp] addSIs2 [InitDmd]) 1, + resolve_tac ((map temp_use (prems RL [STL4])) RL [box_stp_actD]) 1, + asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_actI]) 1 + ]); (* Sometimes easier to use; designed for action B rather than state predicate Q *) -qed_goalw "WF_leadsto" TLA.thy [leadsto] - "[| N .& P .-> $Enabled (_v); \ -\ N .& _v .-> B; \ -\ [](N .& [.~A]_v) .-> stable P |] \ -\ ==> []N .& WF(A)_v .-> (P ~> B)" +qed_goalw "WF_leadsto" TLA.thy [leadsto_def] + "[| |- N & $P --> $Enabled (_v); \ +\ |- N & _v --> B; \ +\ |- [](N & [~A]_v) --> stable P |] \ +\ ==> |- []N & WF(A)_v --> (P ~> B)" (fn [prem1,prem2,prem3] - => [auto_tac (temp_css addSDs2 [BoxWFI]), + => [clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1, etac STL4Edup 1, atac 1, - Auto_tac, - subgoal_tac "sigmaa |= <>_v" 1, - force_tac (temp_css addSEs2 [DmdImpl2,STL4E] addSIs2 [action_mp prem2]) 1, + Clarsimp_tac 1, + rtac (temp_use (prem2 RS DmdImpl)) 1, + rtac (temp_use BoxDmd_simple) 1, atac 1, rtac classical 1, - rtac STL2D 1, - auto_tac (temp_css addsimps2 WF_def::more_temp_simps addSEs2 [mp]), - rtac ImplDmdD 1, - rtac (temp_mp (prem1 RS STL4)) 1, - auto_tac (temp_css addsimps2 [split_box_conj]), - etac INV1I 1, - merge_act_box_tac 1, - auto_tac (temp_css addsimps2 [temp_rewrite not_angle] addSEs2 [temp_mp prem3]) + rtac (temp_use STL2) 1, + clarsimp_tac (temp_css addsimps2 [WF_def] addSEs2 [mp] addSIs2 [InitDmd]) 1, + rtac ((temp_use (prem1 RS STL4)) RS box_stp_actD) 1, + asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_act]) 1, + etac (temp_use INV1) 1, + rtac (temp_use prem3) 1, + asm_full_simp_tac (simpset() addsimps [split_box_conj,temp_use NotDmd,not_angle]) 1 ]); -qed_goalw "SF1" TLA.thy [leadsto] - "[| P .& N .-> P` .| Q`; \ -\ P .& N .& _v .-> Q`; \ -\ []P .& []N .& []F .-> <>$(Enabled(_v)) |] \ -\ ==> []N .& SF(A)_v .& []F .-> (P ~> Q)" - (fn [prem1,prem2,prem3] => - [auto_tac (temp_css addSDs2 [BoxSFI]), - eres_inst_tac [("F","F")] dup_boxE 1, - merge_temp_box_tac 1, - etac STL4Edup 1, atac 1, - Auto_tac, - subgoal_tac "sigmaa |= [](P .-> P` .| Q`)" 1, - auto_tac (temp_css addSDs2 [unless]), - etac (temp_conjimpE INV1) 1, atac 1, - merge_act_box_tac 1, - rtac STL2D 1, - rtac EnsuresInfinite 1, atac 2, - SELECT_GOAL (old_auto_tac (temp_css addsimps2 [SF_alt])) 1, - atac 2, - subgoal_tac "sigmaa |= []<>$(Enabled(_v))" 1, - force_tac (temp_css addsimps2 [dmd_def]) 1, - eres_inst_tac [("F","F")] dup_boxE 1, - etac STL4Edup 1, atac 1, - fast_tac (temp_cs addSEs [STL4E] addSIs [temp_mp prem3]) 1, - fast_tac (action_cs addSIs [action_mp prem2]) 1, - fast_tac (temp_cs addIs [temp_mp DmdPrime]) 1, - fast_tac (temp_cs addSEs [STL4E] addSIs [action_mp prem1]) 1 - ]); +qed_goal "SF1" TLA.thy + "[| |- $P & N --> P` | Q`; \ +\ |- ($P & N) & _v --> Q`; \ +\ |- []P & []N & []F --> <>Enabled(_v) |] \ +\ ==> |- []N & SF(A)_v & []F --> (P ~> Q)" + (fn prems => [ + clarsimp_tac (temp_css addSDs2 [BoxSFI]) 1, + rtac (temp_use ensures) 1, + TRYALL (ares_tac prems), + eres_inst_tac [("F","F")] dup_boxE 1, + merge_temp_box_tac 1, + etac STL4Edup 1, atac 1, + clarsimp_tac (temp_css addsimps2 [SF_def]) 1, + rtac (temp_use STL2) 1, etac mp 1, + resolve_tac (map temp_use (prems RL [STL4])) 1, + asm_simp_tac (simpset() addsimps [split_box_conj, STL3]) 1 + ]); qed_goal "WF2" TLA.thy - "[| N .& _f .-> _g; \ -\ P .& P` .& _f .-> B; \ -\ P .& $(Enabled(_g)) .-> $(Enabled(_f)); \ -\ [](N .& [.~B]_f) .& WF(A)_f .& []F .& <>[]($(Enabled(_g))) .-> <>[]P |] \ -\ ==> []N .& WF(A)_f .& []F .-> WF(M)_g" (fn [prem1,prem2,prem3,prem4] => [ - Auto_tac, - case_tac "sigma |= <>[]$Enabled(_g)" 1, - force_tac (temp_css addsimps2 [WF_def,dmd_def]) 2, - case_tac "sigma |= <>[][.~B]_f" 1, - subgoal_tac "sigma |= <>[]P" 1, - asm_full_simp_tac (simpset() addsimps [WF_def]) 1, - rtac (temp_mp (prem1 RS DmdImpl RS STL4)) 1, - eres_inst_tac [("V","sigma |= <>[][.~B]_f")] thin_rl 1, - etac (temp_conjimpE STL6) 1, atac 1, - subgoal_tac "sigma |= <>[]$Enabled(_f)" 1, - dtac mp 1, atac 1, - subgoal_tac "sigma |= <>([]N .& []P .& []<>_f)" 1, - rtac ((temp_unlift DmdBoxDmd) RS iffD1) 1, - eres_inst_tac [("F","[]N .& []P .& []<>_f")] DmdImplE 1, - SELECT_GOAL Auto_tac 1, - dres_inst_tac [("P","P")] (temp_mp BoxPrime) 1, - merge_act_box_tac 1, - etac InfImpl 1, atac 1, - force_tac (temp_css addsimps2 [angle_def]addSIs2[action_mp prem2]) 1, - etac BoxDmd 1, - dres_inst_tac [("F","<>_f"),("G","[]P")] BoxDmd 1, atac 1, - eres_inst_tac [("F","[]<>_f .& []P")] DmdImplE 1, - Force_tac 1, - rtac (temp_mp (prem3 RS STL4 RS DmdImpl)) 1, - fast_tac (temp_cs addIs [STL4E,DmdImplE]) 1, - etac (temp_conjimpE STL6) 1, atac 1, - eres_inst_tac [("V","sigma |= <>[][.~ B]_f")] thin_rl 1, - dtac BoxWFI 1, - eres_inst_tac [("F","N")] dup_boxE 1, - eres_inst_tac [("F","F")] dup_boxE 1, - merge_temp_box_tac 1, - dtac BoxDmd 1, atac 1, - eres_inst_tac [("V","sigma |= <>[]($Enabled (_g) .& [.~ B]_f)")] thin_rl 1, - rtac dup_dmdD 1, - rtac (temp_mp (prem4 RS DmdImpl)) 1, - etac DmdImplE 1, - force_tac (temp_css addsimps2 [symmetric(temp_rewrite STL5), temp_rewrite STL3] - addIs2 [(temp_unlift WF_Box) RS iffD1, temp_mp ImplDmd]) 1, - asm_full_simp_tac (simpset() addsimps (WF_def::more_temp_simps)) 1, - etac InfImpl 1, atac 1, - auto_tac (temp_css addSIs2 [action_mp prem1]), - ALLGOALS (asm_full_simp_tac (simpset() addsimps [square_def,angle_def])) + "[| |- N & _f --> _g; \ +\ |- $P & P` & _f --> B; \ +\ |- P & Enabled(_g) --> Enabled(_f); \ +\ |- [](N & [~B]_f) & WF(A)_f & []F & <>[]Enabled(_g) --> <>[]P |] \ +\ ==> |- []N & WF(A)_f & []F --> WF(M)_g" +(fn [prem1,prem2,prem3,prem4] => [ + clarsimp_tac (temp_css addSDs2 [BoxWFI, (temp_use BoxDmdBox) RS iffD2] + addsimps2 [read_instantiate [("A","M")] WF_def]) 1, + eres_inst_tac [("F","F")] dup_boxE 1, + merge_temp_box_tac 1, + etac STL4Edup 1, atac 1, + clarsimp_tac (temp_css addSIs2 [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1, + rtac classical 1, + subgoal_tac "sigmaa |= <>(($P & P` & N) & _f)" 1, + force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1, + rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1, + asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1, + merge_act_box_tac 1, + forward_tac [temp_use prem4] 1, TRYALL atac, + dtac (temp_use STL6) 1, atac 1, + eres_inst_tac [("V","sigmaa |= <>[]P")] thin_rl 1, + eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1, + dtac (temp_use BoxWFI) 1, + eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1, + merge_temp_box_tac 1, + etac DmdImpldup 1, atac 1, + auto_tac (temp_css addsimps2 [split_box_conj,STL3,WF_Box,box_stp_act]), + force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1, + rtac (temp_use STL2) 1, + force_tac (temp_css addsimps2 [WF_def,split_box_conj] addSEs2 [mp] + addSIs2 [InitDmd, prem3 RS STL4]) 1 ]); qed_goal "SF2" TLA.thy - "[| N .& _f .-> _g; \ -\ P .& P` .& _f .-> B; \ -\ P .& $(Enabled(_g)) .-> $(Enabled(_f)); \ -\ [](N .& [.~B]_f) .& SF(A)_f .& []F .& []<>($(Enabled(_g))) .-> <>[]P |] \ -\ ==> []N .& SF(A)_f .& []F .-> SF(M)_g" (fn [prem1,prem2,prem3,prem4] => [ - Auto_tac, - case_tac "sigma |= []<>$Enabled(_g)" 1, - force_tac (temp_css addsimps2 [SF_def,dmd_def]) 2, - case_tac "sigma |= <>[][.~B]_f" 1, - subgoal_tac "sigma |= <>[]P" 1, - asm_full_simp_tac (simpset() addsimps [SF_def]) 1, - rtac (temp_mp (prem1 RS DmdImpl RS STL4)) 1, - eres_inst_tac [("V","sigma |= <>[][.~B]_f")] thin_rl 1, - dtac BoxDmdDmdBox 1, atac 1, - subgoal_tac "sigma |= []<>$Enabled(_f)" 1, - dtac mp 1, atac 1, - subgoal_tac "sigma |= <>([]N .& []P .& []<>_f)" 1, - rtac ((temp_unlift DmdBoxDmd) RS iffD1) 1, - eres_inst_tac [("F","[]N .& []P .& []<>_f")] DmdImplE 1, - SELECT_GOAL Auto_tac 1, - dres_inst_tac [("P","P")] (temp_mp BoxPrime) 1, - merge_act_box_tac 1, - etac InfImpl 1, atac 1, - force_tac (temp_css addsimps2 [angle_def]addSIs2[action_mp prem2]) 1, - etac BoxDmd 1, - dres_inst_tac [("F","<>_f"),("G","[]P")] BoxDmd 1, atac 1, - eres_inst_tac [("F","[]<>_f .& []P")] DmdImplE 1, - Force_tac 1, - rtac (temp_mp (prem3 RS DmdImpl RS STL4)) 1, - fast_tac (temp_cs addEs [STL4E,DmdImplE]) 1, - dtac BoxSFI 1, - eres_inst_tac [("F","N")] dup_boxE 1, - eres_inst_tac [("F","F")] dup_boxE 1, - eres_inst_tac [("F","<>$Enabled (_g)")] dup_boxE 1, - merge_temp_box_tac 1, - dtac (temp_conjmp BoxDmdT2) 1 THEN atac 1, - rtac dup_dmdD 1, - rtac (temp_mp (prem4 RS DmdImpl)) 1, - force_tac (temp_css - addsimps2 [symmetric(temp_rewrite STL5), temp_rewrite STL3] - addSIs2 [(temp_unlift WF_Box) RS iffD1, temp_mp ImplDmd] - addSEs2 [DmdImplE]) 1, - asm_full_simp_tac (simpset() addsimps (SF_def::more_temp_simps)) 1, - eres_inst_tac [("F",".~ [.~ B]_f")] InfImpl 1 THEN atac 1, - auto_tac (temp_css addSIs2 [action_mp prem1]), - ALLGOALS (asm_full_simp_tac (simpset() addsimps [square_def,angle_def])) + "[| |- N & _f --> _g; \ +\ |- $P & P` & _f --> B; \ +\ |- P & Enabled(_g) --> Enabled(_f); \ +\ |- [](N & [~B]_f) & SF(A)_f & []F & []<>Enabled(_g) --> <>[]P |] \ +\ ==> |- []N & SF(A)_f & []F --> SF(M)_g" +(fn [prem1,prem2,prem3,prem4] => [ + clarsimp_tac (temp_css addSDs2 [BoxSFI] + addsimps2 [read_instantiate [("A","M")] SF_def]) 1, + eres_inst_tac [("F","F")] dup_boxE 1, + eres_inst_tac [("F","TEMP <>Enabled(_g)")] dup_boxE 1, + merge_temp_box_tac 1, + etac STL4Edup 1, atac 1, + clarsimp_tac (temp_css addSIs2 [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1, + rtac classical 1, + subgoal_tac "sigmaa |= <>(($P & P` & N) & _f)" 1, + force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1, + rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1, + asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1, + merge_act_box_tac 1, + forward_tac [temp_use prem4] 1, TRYALL atac, + eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1, + dtac (temp_use BoxSFI) 1, + eres_inst_tac [("F","TEMP <>Enabled(_g)")] dup_boxE 1, + eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1, + merge_temp_box_tac 1, + etac DmdImpldup 1, atac 1, + auto_tac (temp_css addsimps2 [split_box_conj,STL3,SF_Box,box_stp_act]), + force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1, + rtac (temp_use STL2) 1, + force_tac (temp_css addsimps2 [SF_def,split_box_conj] addSEs2 [mp,InfImpl] + addSIs2 [prem3]) 1 ]); (* ------------------------------------------------------------------------- *) @@ -897,108 +900,70 @@ (* ------------------------------------------------------------------------- *) section "Well-founded orderings"; -val prem1::prems = goal TLA.thy - "[| (wf r); \ -\ !!x. sigma |= [](F x .-> <>G .| <>(REX y. #((y,x):r) .& F y)) \ -\ |] ==> sigma |= [](F x .-> <>G)"; -by (cut_facts_tac [prem1] 1); -by (etac wf_induct 1); -by (subgoal_tac "sigma |= []((REX y. #((y,x):r) .& F y) .-> <>G)" 1); - by (cut_facts_tac prems 1); - by (etac STL4Edup 1 THEN atac 1); - by (Auto_tac THEN etac swap 1 THEN atac 1); - by (rtac dup_dmdD 1); - by (etac DmdImpl2 1 THEN atac 1); -by (subgoal_tac "sigma |= [](RALL y. #((y,x):r) .& F y .-> <>G)" 1); - by (fast_tac (temp_cs addSEs [STL4E]) 1); -by (auto_tac (temp_css addsimps2 [all_box])); -by (etac allE 1 THEN etac impCE 1); - by (rtac (temp_unlift necT) 1); -by Auto_tac; -by (auto_tac (temp_css addSEs2 [STL4E])); -qed "wf_dmd"; - - -(* Special case: leadsto via well-foundedness *) -qed_goalw "wf_leadsto" TLA.thy [leadsto] - "[| (wf r); \ -\ !!x. sigma |= P x ~> (Q .| (REX y. #((y,x):r) .& P y)) \ -\ |] ==> sigma |= P x ~> Q" - (fn prems => [REPEAT (resolve_tac (wf_dmd::prems) 1), - resolve_tac (prems RL [STL4E]) 1, - auto_tac (temp_css addsimps2 [temp_rewrite DmdOr]), - fast_tac temp_cs 1, - etac swap 1, - rewtac (temp_rewrite DmdAct), - auto_tac (temp_css addsimps2 [Init_def] addSEs2 [DmdImplE]) - ]); +qed_goal "wf_leadsto" TLA.thy + "[| wf r; \ +\ !!x. sigma |= F x ~> (G | (? y. #((y,x):r) & F y)) \ +\ |] ==> sigma |= F x ~> G" + (fn p1::prems => + [rtac (p1 RS wf_induct) 1, + rtac (temp_use LatticeTriangle) 1, + resolve_tac prems 1, + auto_tac (temp_css addsimps2 [leadsto_exists]), + case_tac "(y,x):r" 1, + Force_tac 1, + force_tac (temp_css addsimps2 leadsto_def::Init_simps addSIs2 [necT]) 1]); (* If r is well-founded, state function v cannot decrease forever *) qed_goal "wf_not_box_decrease" TLA.thy - "!!r. wf r ==> [][ {[v$, $v]} .: #r ]_v .-> <>[][#False]_v" - (fn _ => [Auto_tac, - subgoal_tac "ALL x. (sigma |= [](Init($v .= #x) .-> <>[][#False]_v))" 1, - etac allE 1, - dtac STL2D 1, - auto_tac (temp_css addsimps2 [Init_def]), - etac wf_dmd 1, - etac dup_boxE 1, - etac STL4E 1, - action_simp_tac (simpset() addsimps [con_abs]) [tempI] [] 1, - case_tac "sigma |= <>[][#False]_v" 1, - ALLGOALS Asm_full_simp_tac, - rewrite_goals_tac more_temp_simps, - dtac STL2D 1, - subgoal_tac "sigma |= <>(REX y. #((y, xa) : r) .& ($v .= #y))" 1, - force_tac (temp_css addsimps2 [DmdAct,Init_def] - addEs2 [DmdImplE]) 1, - subgoal_tac "sigma |= (stable ($v .= #xa) .| <>(REX y. #((y, xa) : r) .& $v .= #y)`)" 1, - case_tac "sigma |= stable ($v .= #xa)" 1, - force_tac (temp_css addIs2 [temp_mp DmdPrime]) 2, - SELECT_GOAL (rewrite_goals_tac ((symmetric (temp_rewrite NotBox))::action_rews)) 1, - etac swap 1, - subgoal_tac "sigma |= []($v .= #xa)" 1, - dres_inst_tac [("P", "$v .= #xa")] (temp_mp BoxPrime) 1, - force_tac (temp_css addEs2 [STL4E] addsimps2 [square_def]) 1, - force_tac (temp_css addSIs2 [INV1I] addsimps2 [Init_def]) 1, - old_auto_tac (temp_css addsimps2 [square_def] addSIs2 [unless] addSEs2 [STL4E]) + "!!r. wf r ==> |- [][ (v`, $v) : #r ]_v --> <>[][#False]_v" + (fn _ => [Clarsimp_tac 1, + rtac ccontr 1, + subgoal_tac "sigma |= (? x. v=#x) ~> #False" 1, + dtac ((temp_use leadsto_false) RS iffD1 RS (temp_use STL2_gen)) 1, + force_tac (temp_css addsimps2 Init_defs) 1, + clarsimp_tac (temp_css addsimps2 [leadsto_exists,not_square]@more_temp_simps) 1, + etac wf_leadsto 1, + rtac (temp_use ensures_simple) 1, TRYALL atac, + auto_tac (temp_css addsimps2 [square_def,angle_def]) ]); -(* "wf ?r ==> <>[][{[?v$, $?v]} .: #?r]_?v .-> <>[][#False]_?v" *) +(* "wf r ==> |- <>[][ (v`, $v) : #r ]_v --> <>[][#False]_v" *) bind_thm("wf_not_dmd_box_decrease", standard(rewrite_rule more_temp_simps (wf_not_box_decrease RS DmdImpl))); -(* If there are infinitely many steps where v decreases w.r.t. r, then there +(* If there are infinitely many steps where v decreases, then there have to be infinitely many non-stuttering steps where v doesn't decrease. *) qed_goal "wf_box_dmd_decrease" TLA.thy - "wf r ==> []<>({[v$, $v]} .: #r) .-> []<><.~({[v$, $v]} .: #r)>_v" - (fn [prem] => [Auto_tac, - rtac ccontr 1, - asm_full_simp_tac - (simpset() addsimps ([action_rewrite not_angle] @ more_temp_simps)) 1, - dtac (prem RS (temp_mp wf_not_dmd_box_decrease)) 1, - dtac BoxDmdDmdBox 1, atac 1, - subgoal_tac "sigma |= []<>((#False)::action)" 1, - Force_tac 1, - etac STL4E 1, - rtac DmdImpl 1, - auto_tac (action_css addsimps2 [square_def] addSEs2 [prem RS wf_irrefl]) - ]); + "wf r ==> |- []<>((v`, $v) : #r) --> []<><(v`, $v) ~: #r>_v" + (fn [prem] => [ + Clarsimp_tac 1, + rtac ccontr 1, + asm_full_simp_tac (simpset() addsimps not_angle::more_temp_simps) 1, + dtac (prem RS (temp_use wf_not_dmd_box_decrease)) 1, + dtac (temp_use BoxDmdDmdBox) 1, atac 1, + subgoal_tac "sigma |= []<>((#False)::action)" 1, + Force_tac 1, + etac STL4E 1, + rtac DmdImpl 1, + force_tac (temp_css addSEs2 [prem RS wf_irrefl]) 1 + ]); (* In particular, for natural numbers, if n decreases infinitely often then it has to increase infinitely often. *) -goal TLA.thy "!!n::nat stfun. []<>(n$ .< $n) .-> []<>($n .< n$)"; -by Auto_tac; -by (subgoal_tac "sigma |= []<><.~( {[n$,$n]} .: #less_than)>_n" 1); - by (etac thin_rl 1 THEN etac STL4E 1 THEN rtac DmdImpl 1); - by (SELECT_GOAL (auto_tac (claset(), simpset() addsimps [angle_def])) 1); - by (rtac nat_less_cases 1); - by (Auto_tac); -by (rtac (temp_mp wf_box_dmd_decrease) 1); - by (auto_tac (claset() addSEs [STL4E] addSIs [DmdImpl], simpset())); -qed "nat_box_dmd_decrease"; +qed_goal "nat_box_dmd_decrease" TLA.thy + "!!n::nat stfun. |- []<>(n` < $n) --> []<>($n < n`)" + (K [Clarsimp_tac 1, + subgoal_tac "sigma |= []<><~( (n`,$n) : #less_than )>_n" 1, + etac thin_rl 1, etac STL4E 1, rtac DmdImpl 1, + clarsimp_tac (temp_css addsimps2 [angle_def]) 1, + rtac nat_less_cases 1, + Auto_tac, + rtac (temp_use wf_box_dmd_decrease) 1, + auto_tac (temp_css addSEs2 [STL4E,DmdImplE]) + ]); + (* ------------------------------------------------------------------------- *) (*** Flexible quantification over state variables ***) @@ -1006,46 +971,55 @@ section "Flexible quantification"; qed_goal "aallI" TLA.thy - "(!!x. base_var x ==> sigma |= F x) ==> sigma |= (AALL x. F(x))" - (fn prems => [auto_tac (temp_css addsimps2 [aall_def] addSEs2 [eexE] addSDs2 prems)]); + "[| basevars vs; (!!x. basevars (x,vs) ==> sigma |= F x) |] ==> sigma |= (AALL x. F x)" + (fn prems => [auto_tac (temp_css addsimps2 [aall_def] addSEs2 [eexE] + addSIs2 prems addSDs2 prems)]); -qed_goal "aallE" TLA.thy - "[| sigma |= (AALL x. F(x)); (!!sigma. sigma |= F(x) ==> P sigma) |] \ -\ ==> (P sigma)::bool" - (fn prems => [cut_facts_tac prems 1, - resolve_tac prems 1, - rewrite_goals_tac (aall_def::intensional_rews), - etac swap 1, - auto_tac (temp_css addSIs2 [temp_mp eexI]) - ]); +qed_goalw "aallE" TLA.thy [aall_def] "|- (AALL x. F x) --> F x" + (K [Clarsimp_tac 1, etac swap 1, + force_tac (temp_css addSIs2 [eexI]) 1]); (* monotonicity of quantification *) qed_goal "eex_mono" TLA.thy - "[| sigma |= EEX x. F(x); !!x. F(x) .-> G(x) |] ==> sigma |= EEX x. G(x)" - (fn [min,maj] => [cut_facts_tac [min] 1, - etac eexE 1, - REPEAT (ares_tac (map temp_mp [eexI,maj]) 1) + "[| sigma |= EEX x. F x; !!x. sigma |= F x --> G x |] ==> sigma |= EEX x. G x" + (fn [min,maj] => [rtac (unit_base RS (min RS eexE)) 1, + rtac (temp_use eexI) 1, + etac ((rewrite_rule intensional_rews maj) RS mp) 1 ]); qed_goal "aall_mono" TLA.thy - "[| sigma |= AALL x. F(x); !!x. F(x) .-> G(x) |] ==> sigma |= AALL x. G(x)" - (fn [min,maj] => [cut_facts_tac [min] 1, - fast_tac (temp_cs addSIs [aallI, temp_mp maj] - addEs [aallE]) 1 + "[| sigma |= AALL x. F(x); !!x. sigma |= F(x) --> G(x) |] ==> sigma |= AALL x. G(x)" + (fn [min,maj] => [rtac (unit_base RS aallI) 1, + rtac ((rewrite_rule intensional_rews maj) RS mp) 1, + rtac (min RS (temp_use aallE)) 1 ]); +(* Derived history introduction rule *) +qed_goal "historyI" TLA.thy + "[| sigma |= Init I; sigma |= []N; basevars vs; \ +\ (!!h. basevars(h,vs) ==> |- I & h = ha --> HI h); \ +\ (!!h s t. [| basevars(h,vs); N (s,t); h t = hb (h s) (s,t) |] ==> HN h (s,t)) \ +\ |] ==> sigma |= EEX h. Init (HI h) & [](HN h)" + (fn [p1,p2,p3,p4,p5] + => [rtac ((temp_use history) RS eexE) 1, + rtac p3 1, + rtac (temp_use eexI) 1, + Clarsimp_tac 1, rtac conjI 1, + cut_facts_tac [p2] 2, + merge_box_tac 2, + force_tac (temp_css addSEs2 [STL4E,p5]) 2, + cut_facts_tac [p1] 1, + force_tac (temp_css addsimps2 Init_defs addSEs2 [p4]) 1 + ]); + (* ---------------------------------------------------------------------- example of a history variable: existence of a clock -Goal "(EEX h. Init($h .= #True) .& [](h$ .~= $h))"; +Goal "|- EEX h. Init(h = #True) & [](h` = (~$h))"; by (rtac tempI 1); by (rtac historyI 1); -by (rewrite_goals_tac action_rews); -by (TRYALL (rtac impI)); -by (TRYALL (etac conjE)); -by (assume_tac 3); -by (Asm_full_simp_tac 3); -by (auto_tac (temp_css addSIs2 [(temp_unlift Init_true) RS iffD2, temp_unlift BoxTrue])); +by (REPEAT (force_tac (temp_css addsimps2 Init_defs addIs2 [unit_base, necT]) 1)); (** solved **) ---------------------------------------------------------------------- *) + diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/TLA.thy --- a/src/HOL/TLA/TLA.thy Mon Feb 08 13:02:42 1999 +0100 +++ b/src/HOL/TLA/TLA.thy Mon Feb 08 13:02:56 1999 +0100 @@ -1,7 +1,7 @@ (* File: TLA/TLA.thy Author: Stephan Merz - Copyright: 1997 University of Munich + Copyright: 1998 University of Munich Theory Name: TLA Logic Image: HOL @@ -9,79 +9,88 @@ The temporal level of TLA. *) -TLA = Action + WF_Rel + - -types - behavior - temporal = "behavior form" - -arities - behavior :: world +TLA = Init + WF_Rel + consts - (* get first 2 states of behavior *) - fst_st :: "behavior => state" - snd_st :: "behavior => state" - - Init :: "action => temporal" - (* define Box and Dmd for both actions and temporals *) - Box :: "('w::world) form => temporal" ("([](_))" [40] 40) - Dmd :: "('w::world) form => temporal" ("(<>(_))" [40] 40) - "~>" :: "[action,action] => temporal" (infixr 22) - stable :: "action => temporal" - WF :: "[action,'a stfun] => temporal" ("(WF'(_')'_(_))" [0,60] 55) - SF :: "[action,'a stfun] => temporal" ("(SF'(_')'_(_))" [0,60] 55) + (** abstract syntax **) + Box :: ('w::world) form => temporal + Dmd :: ('w::world) form => temporal + leadsto :: ['w::world form, 'v::world form] => temporal + Stable :: stpred => temporal + WF :: [action, 'a stfun] => temporal + SF :: [action, 'a stfun] => temporal (* Quantification over (flexible) state variables *) - EEx :: "('a stfun => temporal) => temporal" (binder "EEX " 10) - AAll :: "('a stfun => temporal) => temporal" (binder "AALL " 10) + EEx :: ('a stfun => temporal) => temporal (binder "Eex " 10) + AAll :: ('a stfun => temporal) => temporal (binder "Aall " 10) + + (** concrete syntax **) +syntax + "_Box" :: lift => lift ("([]_)" [40] 40) + "_Dmd" :: lift => lift ("(<>_)" [40] 40) + "_leadsto" :: [lift,lift] => lift ("(_ ~> _)" [23,22] 22) + "_stable" :: lift => lift ("(stable/ _)") + "_WF" :: [lift,lift] => lift ("(WF'(_')'_(_))" [0,60] 55) + "_SF" :: [lift,lift] => lift ("(SF'(_')'_(_))" [0,60] 55) + + "_EEx" :: [idts, lift] => lift ("(3EEX _./ _)" [0,10] 10) + "_AAll" :: [idts, lift] => lift ("(3AALL _./ _)" [0,10] 10) translations - "sigma |= Init(A)" == "Init A sigma" - "sigma |= Box(F)" == "Box F sigma" - "sigma |= Dmd(F)" == "Dmd F sigma" - "sigma |= F ~> G" == "op ~> F G sigma" - "sigma |= stable(A)" == "stable A sigma" - "sigma |= WF(A)_v" == "WF A v sigma" - "sigma |= SF(A)_v" == "SF A v sigma" + "_Box" == "Box" + "_Dmd" == "Dmd" + "_leadsto" == "leadsto" + "_stable" == "Stable" + "_WF" == "WF" + "_SF" == "SF" + "_EEx v A" == "Eex v. A" + "_AAll v A" == "Aall v. A" + + "sigma |= []F" <= "_Box F sigma" + "sigma |= <>F" <= "_Dmd F sigma" + "sigma |= F ~> G" <= "_leadsto F G sigma" + "sigma |= stable P" <= "_stable P sigma" + "sigma |= WF(A)_v" <= "_WF A v sigma" + "sigma |= SF(A)_v" <= "_SF A v sigma" + "sigma |= EEX x. F" <= "_EEx x F sigma" + "sigma |= AALL x. F" <= "_AAll x F sigma" syntax (symbols) - Box :: "('w::world) form => temporal" ("(\\(_))" [40] 40) - Dmd :: "('w::world) form => temporal" ("(\\(_))" [40] 40) + "_Box" :: lift => lift ("(\\_)" [40] 40) + "_Dmd" :: lift => lift ("(\\_)" [40] 40) + "_leadsto" :: [lift,lift] => lift ("(_ \\ _)" [23,22] 22) + "_EEx" :: [idts, lift] => lift ("(3\\\\ _./ _)" [0,10] 10) + "_AAll" :: [idts, lift] => lift ("(3\\\\ _./ _)" [0,10] 10) rules - dmd_def "(<>F) == .~[].~F" - boxact_def "([](F::action)) == ([] Init F)" - leadsto "P ~> Q == [](Init(P) .-> <>Q)" - stable_def "stable P == [](P .-> P`)" - - WF_def "WF(A)_v == <>[] $(Enabled(_v)) .-> []<>_v" - SF_def "SF(A)_v == []<> $(Enabled(_v)) .-> []<>_v" - - Init_def "(sigma |= Init(F)) == ([[fst_st sigma, snd_st sigma]] |= F)" + (* Definitions of derived operators *) + dmd_def "TEMP <>F == TEMP ~[]~F" + boxInit "TEMP []F == TEMP []Init F" + leadsto_def "TEMP F ~> G == TEMP [](Init F --> <>G)" + stable_def "TEMP stable P == TEMP []($P --> P$)" + WF_def "TEMP WF(A)_v == TEMP <>[] Enabled(_v) --> []<>_v" + SF_def "TEMP SF(A)_v == TEMP []<> Enabled(_v) --> []<>_v" + aall_def "TEMP (AALL x. F x) == TEMP ~ (EEX x. ~ F x)" -(* The following axioms are written "polymorphically", not just for temporal formulas. *) - normalT "[](F .-> G) .-> ([]F .-> []G)" - reflT "[]F .-> F" (* F::temporal *) - transT "[]F .-> [][]F" - linT "(<>F) .& (<>G) .-> (<>(F .& <>G)) .| (<>(G .& <>F))" (* F,G::temporal *) - discT "[](F .-> <>(.~F .& <>F)) .-> (F .-> []<>F)" - primeI "[]P .-> Init(P`)" - primeE "[](Init(P) .-> []F) .-> Init(P`) .-> (F .-> []F)" - indT "[](Init(P) .& .~[]F .-> Init(P`) .& F) .-> Init(P) .-> []F" - allT "(RALL x. [](F(x))) .= ([](RALL x. F(x)))" +(* Base axioms for raw TLA. *) + normalT "|- [](F --> G) --> ([]F --> []G)" (* polymorphic *) + reflT "|- []F --> F" (* F::temporal *) + transT "|- []F --> [][]F" (* polymorphic *) + linT "|- <>F & <>G --> (<>(F & <>G)) | (<>(G & <>F))" + discT "|- [](F --> <>(~F & <>F)) --> (F --> []<>F)" + primeI "|- []P --> Init P`" + primeE "|- [](Init P --> []F) --> Init P` --> (F --> []F)" + indT "|- [](Init P & ~[]F --> Init P` & F) --> Init P --> []F" + allT "|- (! x. [](F x)) = ([](! x. F x))" - necT "F ==> []F" + necT "|- F ==> |- []F" (* polymorphic *) (* Flexible quantification: refinement mappings, history variables *) - aall_def "(AALL x. F(x)) == .~ (EEX x. .~ F(x))" - eexI "F x .-> (EEX x. F x)" - historyI "[| sigma |= Init(I); sigma |= []N; - (!!h s t. (h s = ha s t) & I [[s,t]] --> HI(h)[[s,t]]); - (!!h s t. (h t = hc s t (h s)) & N [[s,t]] --> HN(h) [[s,t]]) - |] ==> sigma |= (EEX h. Init(HI(h)) .& []HN(h))" - eexE "[| sigma |= (EEX x. F x); - (!!x. [| base_var x; (sigma |= F x) |] ==> (G sigma)::bool) - |] ==> G sigma" + eexI "|- F x --> (EEX x. F x)" + eexE "[| sigma |= (EEX x. F x); basevars vs; + (!!x. [| basevars (x, vs); sigma |= F x |] ==> (G sigma)::bool) + |] ==> G sigma" + history "|- EEX h. Init(h = ha) & [](!x. $h = #x --> h` = hb x)" +end -end +ML diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/cladata.ML --- a/src/HOL/TLA/cladata.ML Mon Feb 08 13:02:42 1999 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,57 +0,0 @@ -(* Title: TLA/cladata.ML - Author: Stephan Merz (mostly stolen from Isabelle/HOL) - -Setting up the classical reasoner for TLA. - -The classical prover for TLA uses a different hyp_subst_tac that substitutes -somewhat more liberally for state variables. Unfortunately, this requires -either generating a new prover or redefining the basic proof tactics. -We take the latter approach, because otherwise there would be a type conflict -between standard HOL and TLA classical sets, and we would have to redefine -even more things (e.g., blast_tac), and try to keep track of which rules -have been active in setting up a new default claset. - -*) - - -(* Generate a different hyp_subst_tac - that substitutes for x(s) if s is a bound variable of "world" type. - This is useful to solve equations that contain state variables. -*) - -use "hypsubst.ML"; (* local version! *) - -structure ActHypsubst_Data = - struct - structure Simplifier = Simplifier - (*Take apart an equality judgement; otherwise raise Match!*) - fun dest_eq (Const("Trueprop",_) $ (Const("op =",_) $ t $ u)) = (t,u); - val eq_reflection = eq_reflection - val imp_intr = impI - val rev_mp = rev_mp - val subst = subst - val sym = sym - end; - -structure ActHypsubst = ActHypsubstFun(ActHypsubst_Data); -open ActHypsubst; - - -(** - Define the basic classical set and clasimpset for the action part of TLA. - Add the new hyp_subst_tac to the wrapper (also for the default claset). -**) - -val action_css = (HOL_cs addSIs [actionI,intI] addSEs [exE_prop] - addDs [actionD,intD] - addSaltern ("action_hyp_subst_tac", action_hyp_subst_tac), - simpset()); -val action_cs = op addss action_css; - - -AddSIs [actionI,intI]; -AddDs [actionD,intD]; - - - - diff -r f6335d319e9f -r db63752140c7 src/HOL/TLA/hypsubst.ML --- a/src/HOL/TLA/hypsubst.ML Mon Feb 08 13:02:42 1999 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,238 +0,0 @@ -(* Title: ~/projects/isabelle/TLA/hypsubst.ML - Authors: Martin D Coen, Tobias Nipkow and Lawrence C Paulson - Copyright 1995 University of Cambridge - -Tactic to substitute using the assumption x=t in the rest of the subgoal, -and to delete that assumption. Original version due to Martin Coen. - -This version uses the simplifier, and requires it to be already present. - -Local changes for TLA (Stephan Merz): - Simplify equations like f(x) = g(y) if x,y are bound variables. - This is useful for TLA if f and g are state variables. f and g may be - free or bound variables, or even constants. (This may be unsafe, but - we do some type checking to restrict this to state variables!) - - - -Test data: - -Goal "[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)"; -Goal "[| Q(x,y,z); z=f(x); x=z |] ==> P(z)"; -Goal "[| ?x=y; P(?x) |] ==> y = a"; -Goal "[| ?x=y; P(?x) |] ==> y = a"; - -by (hyp_subst_tac 1); -by (bound_hyp_subst_tac 1); - -Here hyp_subst_tac goes wrong; harder still to prove P(f(f(a))) & P(f(a)) -Goal "P(a) --> (EX y. a=y --> P(f(a)))"; -*) - -(*** Signatures unchanged (but renamed) from the original hypsubst.ML ***) - -signature ACTHYPSUBST_DATA = - sig - structure Simplifier : SIMPLIFIER - val dest_eq : term -> term*term - val eq_reflection : thm (* a=b ==> a==b *) - val imp_intr : thm (* (P ==> Q) ==> P-->Q *) - val rev_mp : thm (* [| P; P-->Q |] ==> Q *) - val subst : thm (* [| a=b; P(a) |] ==> P(b) *) - val sym : thm (* a=b ==> b=a *) - end; - - -signature ACTHYPSUBST = - sig - val action_bound_hyp_subst_tac : int -> tactic - val action_hyp_subst_tac : int -> tactic - (*exported purely for debugging purposes*) - val gen_hyp_subst_tac : bool -> int -> tactic - val vars_gen_hyp_subst_tac : bool -> int -> tactic - val eq_var : bool -> bool -> term -> int * bool - val inspect_pair : bool -> bool -> term * term -> bool - val mk_eqs : thm -> thm list - val thin_leading_eqs_tac : bool -> int -> int -> tactic - end; - - -functor ActHypsubstFun(Data: ACTHYPSUBST_DATA): ACTHYPSUBST = -struct - -fun STATE tacfun st = tacfun st st; - - -local open Data - BasisLibrary (*for Int, List, ...*) -in - -exception EQ_VAR; - -fun loose (i,t) = 0 mem add_loose_bnos(t,i,[]); - -local val odot = ord"." -in -(*Simplifier turns Bound variables to dotted Free variables: - change it back (any Bound variable will do) -*) -fun contract t = - case Pattern.eta_contract t of - Free(a,T) => if (ord a = odot) then Bound 0 else Free(a,T) - | Free at $ Free(b,T) => Free at $ - (if ord b = odot then Bound 0 else Free(b,T)) - | t' => t' -end; - -fun has_vars t = maxidx_of_term t <> ~1; - -(* Added for TLA version. - Is type ty the type of a state variable? Only then do we substitute - in applications. This function either returns true or raises Match. -*) -fun is_stvar (Type("fun", Type("Stfun.state",[])::_)) = true; - - -(*If novars then we forbid Vars in the equality. - If bnd then we only look for Bound (or dotted Free) variables to eliminate. - When can we safely delete the equality? - Not if it equates two constants; consider 0=1. - Not if it resembles x=t[x], since substitution does not eliminate x. - Not if it resembles ?x=0; another goal could instantiate ?x to Suc(i) - Not if it involves a variable free in the premises, - but we can't check for this -- hence bnd and bound_hyp_subst_tac - Prefer to eliminate Bound variables if possible. - Result: true = use as is, false = reorient first *) -fun inspect_pair bnd novars (t,u) = - case (contract t, contract u) of - (Bound i, _) => if loose(i,u) orelse novars andalso has_vars u - then raise Match - else true (*eliminates t*) - | (_, Bound i) => if loose(i,t) orelse novars andalso has_vars t - then raise Match - else false (*eliminates u*) - | (Free _, _) => if bnd orelse Logic.occs(t,u) orelse - novars andalso has_vars u - then raise Match - else true (*eliminates t*) - | (_, Free _) => if bnd orelse Logic.occs(u,t) orelse - novars andalso has_vars t - then raise Match - else false (*eliminates u*) - | (Free(_,ty) $ (Bound _), _) => - if bnd orelse - novars andalso has_vars u - then raise Match - else is_stvar(ty) (* changed for TLA *) - | (_, Free(_,ty) $ (Bound _)) => - if bnd orelse - novars andalso has_vars t - then raise Match - else not(is_stvar(ty)) (* changed for TLA *) - | ((Bound _) $ (Bound _), _) => (* can't check for types here *) - if bnd orelse - novars andalso has_vars u - then raise Match - else true - | (_, (Bound _) $ (Bound _)) => (* can't check for types here *) - if bnd orelse - novars andalso has_vars t - then raise Match - else false - | (Const(_,ty) $ (Bound _), _) => - if bnd orelse - novars andalso has_vars u - then raise Match - else is_stvar(ty) (* changed for TLA *) - | (_, Const(_,ty) $ (Bound _)) => - if bnd orelse - novars andalso has_vars t - then raise Match - else not(is_stvar(ty)) (* changed for TLA *) - | _ => raise Match; - -(*Locates a substitutable variable on the left (resp. right) of an equality - assumption. Returns the number of intervening assumptions. *) -fun eq_var bnd novars = - let fun eq_var_aux k (Const("all",_) $ Abs(_,_,t)) = eq_var_aux k t - | eq_var_aux k (Const("==>",_) $ A $ B) = - ((k, inspect_pair bnd novars (dest_eq A)) - (*Exception comes from inspect_pair or dest_eq*) - handle Match => eq_var_aux (k+1) B) - | eq_var_aux k _ = raise EQ_VAR - in eq_var_aux 0 end; - -(*We do not try to delete ALL equality assumptions at once. But - it is easy to handle several consecutive equality assumptions in a row. - Note that we have to inspect the proof state after doing the rewriting, - since e.g. z=f(x); x=z changes to z=f(x); x=f(x) and the second equality - must NOT be deleted. Tactic must rotate or delete m assumptions. -*) -fun thin_leading_eqs_tac bnd m i = STATE(fn state => - let fun count [] = 0 - | count (A::Bs) = ((inspect_pair bnd true (dest_eq A); - 1 + count Bs) - handle Match => 0) - val (_,_,Bi,_) = dest_state(state,i) - val j = Int.min (m, count (Logic.strip_assums_hyp Bi)) - in REPEAT_DETERM_N j (etac thin_rl i) THEN - REPEAT_DETERM_N (m-j) (etac revcut_rl i) - end); - -(*For the simpset. Adds ALL suitable equalities, even if not first! - No vars are allowed here, as simpsets are built from meta-assumptions*) -fun mk_eqs th = - [ if inspect_pair false false (Data.dest_eq (#prop (rep_thm th))) - then th RS Data.eq_reflection - else symmetric(th RS Data.eq_reflection) (*reorient*) ] - handle Match => []; (*Exception comes from inspect_pair or dest_eq*) - -local open Simplifier -in - - val hyp_subst_ss = empty_ss setmksimps mk_eqs - - (*Select a suitable equality assumption and substitute throughout the subgoal - Replaces only Bound variables if bnd is true*) - fun gen_hyp_subst_tac bnd i = DETERM (STATE(fn state => - let val (_,_,Bi,_) = dest_state(state,i) - val n = length(Logic.strip_assums_hyp Bi) - 1 - val (k,_) = eq_var bnd true Bi - in - EVERY [REPEAT_DETERM_N k (etac revcut_rl i), - asm_full_simp_tac hyp_subst_ss i, - etac thin_rl i, - thin_leading_eqs_tac bnd (n-k) i] - end - handle THM _ => no_tac | EQ_VAR => no_tac)); - -end; - -val ssubst = standard (sym RS subst); - -(*Old version of the tactic above -- slower but the only way - to handle equalities containing Vars.*) -fun vars_gen_hyp_subst_tac bnd i = DETERM (STATE(fn state => - let val (_,_,Bi,_) = dest_state(state,i) - val n = length(Logic.strip_assums_hyp Bi) - 1 - val (k,symopt) = eq_var bnd false Bi - in - EVERY [REPEAT_DETERM_N k (etac rev_mp i), - etac revcut_rl i, - REPEAT_DETERM_N (n-k) (etac rev_mp i), - etac (if symopt then ssubst else subst) i, - REPEAT_DETERM_N n (rtac imp_intr i)] - end - handle THM _ => no_tac | EQ_VAR => no_tac)); - -(*Substitutes for Free or Bound variables*) -val action_hyp_subst_tac = - (* gen_hyp_subst_tac false ORELSE' *) vars_gen_hyp_subst_tac false; - -(*Substitutes for Bound variables only -- this is always safe*) -val action_bound_hyp_subst_tac = - (* gen_hyp_subst_tac true ORELSE' *) vars_gen_hyp_subst_tac true; - -end -end; -