# HG changeset patch # User wenzelm # Date 1165024322 -3600 # Node ID 6f79647cf53662cbe4dd1afb90ec21f85a505eb8 # Parent 17098171d46aefc6f1c2cf19c2c2c5c8ce99aa34 TLA: converted legacy ML scripts; diff -r 17098171d46a -r 6f79647cf536 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/IsaMakefile Sat Dec 02 02:52:02 2006 +0100 @@ -693,9 +693,8 @@ TLA: HOL $(OUT)/TLA -$(OUT)/TLA: $(OUT)/HOL TLA/Action.ML TLA/Action.thy TLA/Init.ML \ - TLA/Init.thy TLA/Intensional.ML TLA/Intensional.thy \ - TLA/ROOT.ML TLA/Stfun.ML TLA/Stfun.thy TLA/TLA.ML TLA/TLA.thy +$(OUT)/TLA: $(OUT)/HOL TLA/Action.thy TLA/Init.thy TLA/Intensional.thy \ + TLA/ROOT.ML TLA/Stfun.thy TLA/TLA.thy @cd TLA; $(ISATOOL) usedir -b $(OUT)/HOL TLA @@ -703,7 +702,7 @@ TLA-Inc: TLA $(LOG)/TLA-Inc.gz -$(LOG)/TLA-Inc.gz: $(OUT)/TLA TLA/Inc/Inc.thy TLA/Inc/Inc.ML +$(LOG)/TLA-Inc.gz: $(OUT)/TLA TLA/Inc/Inc.thy @cd TLA; $(ISATOOL) usedir $(OUT)/TLA Inc @@ -711,8 +710,7 @@ TLA-Buffer: TLA $(LOG)/TLA-Buffer.gz -$(LOG)/TLA-Buffer.gz: $(OUT)/TLA TLA/Buffer/Buffer.thy \ - TLA/Buffer/Buffer.ML TLA/Buffer/DBuffer.thy TLA/Buffer/DBuffer.ML +$(LOG)/TLA-Buffer.gz: $(OUT)/TLA TLA/Buffer/Buffer.thy TLA/Buffer/DBuffer.thy @cd TLA; $(ISATOOL) usedir $(OUT)/TLA Buffer @@ -721,15 +719,10 @@ TLA-Memory: TLA $(LOG)/TLA-Memory.gz $(LOG)/TLA-Memory.gz: $(OUT)/TLA TLA/Memory/MIParameters.thy \ - TLA/Memory/MIsafe.ML TLA/Memory/MemClerk.ML \ - TLA/Memory/MemClerk.thy TLA/Memory/MemClerkParameters.ML \ - TLA/Memory/MemClerkParameters.thy TLA/Memory/Memory.ML \ - TLA/Memory/Memory.thy TLA/Memory/MemoryImplementation.ML \ - TLA/Memory/MemoryImplementation.thy TLA/Memory/MemoryParameters.ML \ - TLA/Memory/MemoryParameters.thy TLA/Memory/ProcedureInterface.ML \ - TLA/Memory/ProcedureInterface.thy TLA/Memory/RPC.ML TLA/Memory/RPC.thy \ - TLA/Memory/RPCMemoryParams.thy TLA/Memory/RPCParameters.ML \ - TLA/Memory/RPCParameters.thy + TLA/Memory/MemClerk.thy TLA/Memory/MemClerkParameters.thy \ + TLA/Memory/Memory.thy TLA/Memory/MemoryImplementation.thy \ + TLA/Memory/MemoryParameters.thy TLA/Memory/ProcedureInterface.thy TLA/Memory/RPC.thy \ + TLA/Memory/RPCMemoryParams.thy TLA/Memory/RPCParameters.thy @cd TLA; $(ISATOOL) usedir $(OUT)/TLA Memory diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Action.ML --- a/src/HOL/TLA/Action.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,248 +0,0 @@ -(* - File: Action.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - -Lemmas and tactics for TLA actions. -*) - -(* The following assertion specializes "intI" for any world type - which is a pair, not just for "state * state". -*) -val [prem] = goal (the_context ()) "(!!s t. (s,t) |= A) ==> |- A"; -by (REPEAT (resolve_tac [prem,intI,prod_induct] 1)); -qed "actionI"; - -Goal "|- A ==> (s,t) |= A"; -by (etac intD 1); -qed "actionD"; - -local - fun prover s = prove_goal (the_context ()) 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; - -(* ================ Functions to "unlift" action theorems into HOL rules ================ *) - -(* 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)) - handle _ => int_unlift th; - -(* Turn |- A = B into meta-level rewrite rule A == B *) -val action_rewrite = int_rewrite; - -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 ============================= *) - -AddSIs [actionI]; -AddDs [actionD]; - -(* =========================== square / angle brackets =========================== *) - -Goalw [square_def] "(s,t) |= unchanged v ==> (s,t) |= [A]_v"; -by (Asm_full_simp_tac 1); -qed "idle_squareI"; - -Goalw [square_def] "(s,t) |= A ==> (s,t) |= [A]_v"; -by (Asm_simp_tac 1); -qed "busy_squareI"; - -val prems = goal (the_context ()) - "[| (s,t) |= [A]_v; A (s,t) ==> B (s,t); v t = v s ==> B (s,t) |] ==> B (s,t)"; -by (cut_facts_tac prems 1); -by (rewrite_goals_tac (square_def::action_rews)); -by (etac disjE 1); -by (REPEAT (eresolve_tac prems 1)); -qed "squareE"; - -val prems = goalw (the_context ()) (square_def::action_rews) - "[| v t ~= v s ==> A (s,t) |] ==> (s,t) |= [A]_v"; -by (rtac disjCI 1); -by (eresolve_tac prems 1); -qed "squareCI"; - -goalw (the_context ()) [angle_def] - "!!s t. [| A (s,t); v t ~= v s |] ==> (s,t) |= _v"; -by (Asm_simp_tac 1); -qed "angleI"; - -val prems = goalw (the_context ()) (angle_def::action_rews) - "[| (s,t) |= _v; [| A (s,t); v t ~= v s |] ==> R |] ==> R"; -by (cut_facts_tac prems 1); -by (etac conjE 1); -by (REPEAT (ares_tac prems 1)); -qed "angleE"; - -AddIs [angleI, squareCI]; -AddEs [angleE, squareE]; - -goal (the_context ()) - "!!f. [| |- unchanged f & ~B --> unchanged g; \ -\ |- A & ~unchanged g --> B \ -\ |] ==> |- [A]_f --> [B]_g"; -by (Clarsimp_tac 1); -by (etac squareE 1); -by (auto_tac (claset(), simpset() addsimps [square_def])); -qed "square_simulation"; - -goalw (the_context ()) [square_def,angle_def] - "|- (~ [A]_v) = <~A>_v"; -by Auto_tac; -qed "not_square"; - -goalw (the_context ()) [square_def,angle_def] - "|- (~ _v) = [~A]_v"; -by Auto_tac; -qed "not_angle"; - -(* ============================== Facts about ENABLED ============================== *) - -goal (the_context ()) "|- A --> $Enabled A"; -by (auto_tac (claset(), simpset() addsimps [enabled_def])); -qed "enabledI"; - -val prems = goalw (the_context ()) [enabled_def] - "[| s |= Enabled A; !!u. A (s,u) ==> Q |] ==> Q"; -by (cut_facts_tac prems 1); -by (etac exE 1); -by (resolve_tac prems 1); -by (atac 1); -qed "enabledE"; - -goal (the_context ()) "|- ~$Enabled G --> ~ G"; -by (auto_tac (claset(), simpset() addsimps [enabled_def])); -qed "notEnabledD"; - -(* Monotonicity *) -val [min,maj] = goal (the_context ()) - "[| s |= Enabled F; |- F --> G |] ==> s |= Enabled G"; -by (rtac (min RS enabledE) 1); -by (rtac (action_use enabledI) 1); -by (etac (action_use maj) 1); -qed "enabled_mono"; - -(* stronger variant *) -val [min,maj] = goal (the_context ()) - "[| s |= Enabled F; !!t. F (s,t) ==> G (s,t) |] ==> s |= Enabled G"; -by (rtac (min RS enabledE) 1); -by (rtac (action_use enabledI) 1); -by (etac maj 1); -qed "enabled_mono2"; - -goal (the_context ()) "|- Enabled F --> Enabled (F | G)"; -by (auto_tac (claset() addSEs [enabled_mono], simpset())); -qed "enabled_disj1"; - -goal (the_context ()) "|- Enabled G --> Enabled (F | G)"; -by (auto_tac (claset() addSEs [enabled_mono], simpset())); -qed "enabled_disj2"; - -goal (the_context ()) "|- Enabled (F & G) --> Enabled F"; -by (auto_tac (claset() addSEs [enabled_mono], simpset())); -qed "enabled_conj1"; - -goal (the_context ()) "|- Enabled (F & G) --> Enabled G"; -by (auto_tac (claset() addSEs [enabled_mono], simpset())); -qed "enabled_conj2"; - -val prems = goal (the_context ()) - "[| s |= Enabled (F & G); [| s |= Enabled F; s |= Enabled G |] ==> Q |] ==> Q"; -by (cut_facts_tac prems 1); -by (resolve_tac prems 1); -by (etac (action_use enabled_conj1) 1); -by (etac (action_use enabled_conj2) 1); -qed "enabled_conjE"; - -goal (the_context ()) "|- Enabled (F | G) --> Enabled F | Enabled G"; -by (auto_tac (claset(), simpset() addsimps [enabled_def])); -qed "enabled_disjD"; - -goal (the_context ()) "|- Enabled (F | G) = (Enabled F | Enabled G)"; -by (Clarsimp_tac 1); -by (rtac iffI 1); -by (etac (action_use enabled_disjD) 1); -by (REPEAT (eresolve_tac (disjE::map action_use [enabled_disj1,enabled_disj2]) 1)); -qed "enabled_disj"; - -goal (the_context ()) "|- Enabled (EX x. F x) = (EX x. Enabled (F x))"; -by (force_tac (claset(), simpset() addsimps [enabled_def]) 1); -qed "enabled_ex"; - - -(* A rule that combines enabledI and baseE, but generates fewer instantiations *) -val prems = goal (the_context ()) - "[| basevars vs; EX c. ! u. vs u = c --> A(s,u) |] ==> s |= Enabled A"; -by (cut_facts_tac prems 1); -by (etac exE 1); -by (etac baseE 1); -by (rtac (action_use enabledI) 1); -by (etac allE 1); -by (etac mp 1); -by (atac 1); -qed "base_enabled"; - -(* ======================= 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 "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. - - 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. -*) - -fun enabled_tac base_vars = - clarsimp_tac (claset() addSIs [base_vars RS base_enabled], simpset()); - -(* Example: - -val [prem] = goal (the_context ()) "basevars (x,y,z) ==> |- x --> Enabled ($x & (y$ = #False))"; -by (enabled_tac prem 1); -by Auto_tac; - -*) diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Action.thy --- a/src/HOL/TLA/Action.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Action.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1998 University of Munich +*) - Theory Name: Action - Logic Image: HOL - -Define the action level of TLA as an Isabelle theory. -*) +header {* The action level of TLA as an Isabelle theory *} theory Action imports Stfun @@ -75,6 +72,250 @@ enabled_def: "s |= Enabled A == EX u. (s,u) |= A" -ML {* use_legacy_bindings (the_context ()) *} + +(* The following assertion specializes "intI" for any world type + which is a pair, not just for "state * state". +*) + +lemma actionI [intro!]: + assumes "!!s t. (s,t) |= A" + shows "|- A" + apply (rule assms intI prod_induct)+ + done + +lemma actionD [dest]: "|- A ==> (s,t) |= A" + apply (erule intD) + done + +lemma pr_rews [int_rewrite]: + "|- (#c)` = #c" + "!!f. |- f` = f" + "!!f. |- f` = f" + "!!f. |- f` = f" + "|- (! x. P x)` = (! x. (P x)`)" + "|- (? x. P x)` = (? x. (P x)`)" + by (rule actionI, unfold unl_after intensional_rews, rule refl)+ + + +lemmas act_rews [simp] = unl_before unl_after unchanged_def pr_rews + +lemmas action_rews = act_rews intensional_rews + + +(* ================ Functions to "unlift" action theorems into HOL rules ================ *) + +ML {* +(* 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. +*) +local + val action_rews = thms "action_rews"; + val actionD = thm "actionD"; +in + +fun action_unlift th = + (rewrite_rule action_rews (th RS actionD)) + handle THM _ => int_unlift th; + +(* Turn |- A = B into meta-level rewrite rule A == B *) +val action_rewrite = int_rewrite + +fun action_use th = + case (concl_of th) of + Const _ $ (Const ("Intensional.Valid", _) $ _) => + (flatten (action_unlift th) handle THM _ => th) + | _ => th; end +*} + +setup {* + Attrib.add_attributes [ + ("action_unlift", Attrib.no_args (Thm.rule_attribute (K action_unlift)), ""), + ("action_rewrite", Attrib.no_args (Thm.rule_attribute (K action_rewrite)), ""), + ("action_use", Attrib.no_args (Thm.rule_attribute (K action_use)), "")] +*} + + +(* =========================== square / angle brackets =========================== *) + +lemma idle_squareI: "(s,t) |= unchanged v ==> (s,t) |= [A]_v" + by (simp add: square_def) + +lemma busy_squareI: "(s,t) |= A ==> (s,t) |= [A]_v" + by (simp add: square_def) + +lemma squareE [elim]: + "[| (s,t) |= [A]_v; A (s,t) ==> B (s,t); v t = v s ==> B (s,t) |] ==> B (s,t)" + apply (unfold square_def action_rews) + apply (erule disjE) + apply simp_all + done + +lemma squareCI [intro]: "[| v t ~= v s ==> A (s,t) |] ==> (s,t) |= [A]_v" + apply (unfold square_def action_rews) + apply (rule disjCI) + apply (erule (1) meta_mp) + done + +lemma angleI [intro]: "!!s t. [| A (s,t); v t ~= v s |] ==> (s,t) |= _v" + by (simp add: angle_def) + +lemma angleE [elim]: "[| (s,t) |= _v; [| A (s,t); v t ~= v s |] ==> R |] ==> R" + apply (unfold angle_def action_rews) + apply (erule conjE) + apply simp + done + +lemma square_simulation: + "!!f. [| |- unchanged f & ~B --> unchanged g; + |- A & ~unchanged g --> B + |] ==> |- [A]_f --> [B]_g" + apply clarsimp + apply (erule squareE) + apply (auto simp add: square_def) + done + +lemma not_square: "|- (~ [A]_v) = <~A>_v" + by (auto simp: square_def angle_def) + +lemma not_angle: "|- (~ _v) = [~A]_v" + by (auto simp: square_def angle_def) + + +(* ============================== Facts about ENABLED ============================== *) + +lemma enabledI: "|- A --> $Enabled A" + by (auto simp add: enabled_def) + +lemma enabledE: "[| s |= Enabled A; !!u. A (s,u) ==> Q |] ==> Q" + apply (unfold enabled_def) + apply (erule exE) + apply simp + done + +lemma notEnabledD: "|- ~$Enabled G --> ~ G" + by (auto simp add: enabled_def) + +(* Monotonicity *) +lemma enabled_mono: + assumes min: "s |= Enabled F" + and maj: "|- F --> G" + shows "s |= Enabled G" + apply (rule min [THEN enabledE]) + apply (rule enabledI [action_use]) + apply (erule maj [action_use]) + done + +(* stronger variant *) +lemma enabled_mono2: + assumes min: "s |= Enabled F" + and maj: "!!t. F (s,t) ==> G (s,t)" + shows "s |= Enabled G" + apply (rule min [THEN enabledE]) + apply (rule enabledI [action_use]) + apply (erule maj) + done + +lemma enabled_disj1: "|- Enabled F --> Enabled (F | G)" + by (auto elim!: enabled_mono) + +lemma enabled_disj2: "|- Enabled G --> Enabled (F | G)" + by (auto elim!: enabled_mono) + +lemma enabled_conj1: "|- Enabled (F & G) --> Enabled F" + by (auto elim!: enabled_mono) + +lemma enabled_conj2: "|- Enabled (F & G) --> Enabled G" + by (auto elim!: enabled_mono) + +lemma enabled_conjE: + "[| s |= Enabled (F & G); [| s |= Enabled F; s |= Enabled G |] ==> Q |] ==> Q" + apply (frule enabled_conj1 [action_use]) + apply (drule enabled_conj2 [action_use]) + apply simp + done + +lemma enabled_disjD: "|- Enabled (F | G) --> Enabled F | Enabled G" + by (auto simp add: enabled_def) + +lemma enabled_disj: "|- Enabled (F | G) = (Enabled F | Enabled G)" + apply clarsimp + apply (rule iffI) + apply (erule enabled_disjD [action_use]) + apply (erule disjE enabled_disj1 [action_use] enabled_disj2 [action_use])+ + done + +lemma enabled_ex: "|- Enabled (EX x. F x) = (EX x. Enabled (F x))" + by (force simp add: enabled_def) + + +(* A rule that combines enabledI and baseE, but generates fewer instantiations *) +lemma base_enabled: + "[| basevars vs; EX c. ! u. vs u = c --> A(s,u) |] ==> s |= Enabled A" + apply (erule exE) + apply (erule baseE) + apply (rule enabledI [action_use]) + apply (erule allE) + apply (erule mp) + apply assumption + done + +(* ======================= action_simp_tac ============================== *) + +ML {* +(* 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. +*) +local + val actionI = thm "actionI"; + val intI = thm "intI"; +in + +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()) [] [] + +end +*} + +(* ---------------- enabled_tac: tactic to prove (Enabled A) -------------------- *) + +ML {* +(* "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 "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. + - 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. +*) +local + val base_enabled = thm "base_enabled"; +in + +fun enabled_tac base_vars = + clarsimp_tac (claset() addSIs [base_vars RS base_enabled], simpset()); + +end +*} + +(* Example *) + +lemma + assumes "basevars (x,y,z)" + shows "|- x --> Enabled ($x & (y$ = #False))" + apply (tactic {* enabled_tac (thm "assms") 1 *}) + apply auto + done + +end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Buffer/Buffer.ML --- a/src/HOL/TLA/Buffer/Buffer.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,37 +0,0 @@ -(* - File: Buffer.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - Simple FIFO buffer (theorems and proofs) -*) - -(* ---------------------------- Data lemmas ---------------------------- *) - -(*FIXME: move to theory List? Maybe as (tl xs = xs) = (xs = [])"?*) -Goal "xs ~= [] --> tl xs ~= xs"; -by (auto_tac (claset(), simpset() addsimps [neq_Nil_conv])); -qed_spec_mp "tl_not_self"; - -Addsimps [tl_not_self]; - -(* ---------------------------- Action lemmas ---------------------------- *) - -(* Dequeue is visible *) -Goalw [angle_def,Deq_def] "|- _(ic,q,oc) = Deq ic q oc"; -by (REPEAT (Safe_tac THEN Asm_lr_simp_tac 1)); -qed "Deq_visible"; - -(* Enabling condition for dequeue -- NOT NEEDED *) -Goalw [temp_rewrite Deq_visible] - "!!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 (_(ic,q,oc)) --> (q ~= #[])"; -by (auto_tac (claset() addSEs [enabledE], simpset() addsimps [Deq_def])); -qed "Deq_enabledE"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Buffer/Buffer.thy --- a/src/HOL/TLA/Buffer/Buffer.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Buffer/Buffer.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,9 +3,6 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich - - Theory Name: Buffer - Logic Image: TLA *) header {* A simple FIFO buffer (synchronous communication, interleaving) *} @@ -40,6 +37,34 @@ & WF(Deq ic q oc)_(ic,q,oc)" Buffer_def: "Buffer ic oc == TEMP (EEX q. IBuffer ic q oc)" -ML {* use_legacy_bindings (the_context ()) *} + +(* ---------------------------- Data lemmas ---------------------------- *) + +(*FIXME: move to theory List? Maybe as (tl xs = xs) = (xs = [])"?*) +lemma tl_not_self [simp]: "xs ~= [] ==> tl xs ~= xs" + by (auto simp: neq_Nil_conv) + + +(* ---------------------------- Action lemmas ---------------------------- *) + +(* Dequeue is visible *) +lemma Deq_visible: "|- _(ic,q,oc) = Deq ic q oc" + apply (unfold angle_def Deq_def) + apply (safe, simp (asm_lr))+ + done + +(* Enabling condition for dequeue -- NOT NEEDED *) +lemma Deq_enabled: + "!!q. basevars (ic,q,oc) ==> |- Enabled (_(ic,q,oc)) = (q ~= #[])" + apply (unfold Deq_visible [temp_rewrite]) + apply (force elim!: base_enabled [temp_use] enabledE [temp_use] simp: Deq_def) + done + +(* For the left-to-right implication, we don't need the base variable stuff *) +lemma Deq_enabledE: + "|- Enabled (_(ic,q,oc)) --> (q ~= #[])" + apply (unfold Deq_visible [temp_rewrite]) + apply (auto elim!: enabledE simp add: Deq_def) + done end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Buffer/DBuffer.ML --- a/src/HOL/TLA/Buffer/DBuffer.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,119 +0,0 @@ -(* - File: DBuffer.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - Double FIFO buffer implements simple FIFO buffer. -*) - - -val db_css = (claset(), simpset() addsimps [qc_def]); -Addsimps [qc_def]; - -val db_defs = [BInit_def, Enq_def, Deq_def, Next_def, IBuffer_def, Buffer_def, - DBInit_def,DBEnq_def,DBDeq_def,DBPass_def,DBNext_def,DBuffer_def]; - - -(*** Proper initialization ***) -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]_(inp,mid,out,q1,q2) --> [Next inp qc out]_(inp,qc,out)"; -by (rtac square_simulation 1); -by (Clarsimp_tac 1); -by (action_simp_tac (simpset() addsimps hd_append::db_defs) [] [] 1); -qed "DB_step_simulation"; - - -(*** Simulation of fairness ***) - -(* Compute enabledness predicates for DBDeq and DBPass actions *) -Goalw [angle_def,DBDeq_def,Deq_def] "|- _(inp,mid,out,q1,q2) = DBDeq"; -by (REPEAT (Safe_tac THEN Asm_lr_simp_tac 1)); -qed "DBDeq_visible"; - -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 "|- _(inp,mid,out,q1,q2) = DBPass"; -by (auto_tac (db_css addsimps2 [angle_def,DBPass_def,Deq_def])); -qed "DBPass_visible"; - -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"; - - -(* The plan for proving weak fairness at the higher level is to prove - (0) DBuffer => (Enabled (Deq inp qc out) ~> (Deq inp qc out)) - 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 _(inp,qc,out) - (and therefore DBDeq ~> _(inp,qc,out) trivially holds). - - Condition (1) is reduced to - (1a) DBuffer => (qc ~= [] /\ q2 = [] ~> q2 ~= []) - by standard leadsto rules (leadsto_classical) and rule Deq_enabledE. - - Both (1a) and (2) are proved from DBuffer's WF conditions by standard - WF reasoning (Lamport's WF1 and WF_leadsto). - The condition WF(Deq inp qc out) follows from (0) by rule leadsto_WF. - - One could use Lamport's WF2 instead. -*) - -(* Condition (1a) *) -Goal "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2) \ -\ --> (qc ~= #[] & q2 = #[] ~> q2 ~= #[])"; -by (rtac WF1 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]_(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]_(inp,mid,out,q1,q2) & WF(DBDeq)_(inp,mid,out,q1,q2) \ -\ --> (q2 ~= #[] ~> DBDeq)"; -by (rtac WF_leadsto 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]_(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 (force_tac (temp_css addSIs2 [eexI,DBInit,DB_step_simulation RS STL4,DBFair]) 1); -qed "DBuffer_impl_Buffer"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Buffer/DBuffer.thy --- a/src/HOL/TLA/Buffer/DBuffer.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Buffer/DBuffer.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,9 +3,6 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich - - Theory Name: DBuffer - Logic Image: TLA *) header {* Two FIFO buffers in a row, with interleaving assumption *} @@ -48,6 +45,119 @@ & WF(DBDeq)_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2)" -ML {* use_legacy_bindings (the_context ()) *} + +declare qc_def [simp] + +lemmas db_defs = + BInit_def Enq_def Deq_def Next_def IBuffer_def Buffer_def + DBInit_def DBEnq_def DBDeq_def DBPass_def DBNext_def DBuffer_def + + +(*** Proper initialization ***) +lemma DBInit: "|- Init DBInit --> Init (BInit inp qc out)" + by (auto simp: Init_def DBInit_def BInit_def) + + +(*** Step simulation ***) +lemma DB_step_simulation: "|- [DBNext]_(inp,mid,out,q1,q2) --> [Next inp qc out]_(inp,qc,out)" + apply (rule square_simulation) + apply clarsimp + apply (tactic + {* action_simp_tac (simpset () addsimps (thm "hd_append" :: thms "db_defs")) [] [] 1 *}) + done + + +(*** Simulation of fairness ***) + +(* Compute enabledness predicates for DBDeq and DBPass actions *) +lemma DBDeq_visible: "|- _(inp,mid,out,q1,q2) = DBDeq" + apply (unfold angle_def DBDeq_def Deq_def) + apply (safe, simp (asm_lr))+ + done + +lemma DBDeq_enabled: + "|- Enabled (_(inp,mid,out,q1,q2)) = (q2 ~= #[])" + apply (unfold DBDeq_visible [action_rewrite]) + apply (force intro!: DB_base [THEN base_enabled, temp_use] + elim!: enabledE simp: angle_def DBDeq_def Deq_def) + done + +lemma DBPass_visible: "|- _(inp,mid,out,q1,q2) = DBPass" + by (auto simp: angle_def DBPass_def Deq_def) + +lemma DBPass_enabled: + "|- Enabled (_(inp,mid,out,q1,q2)) = (q1 ~= #[])" + apply (unfold DBPass_visible [action_rewrite]) + apply (force intro!: DB_base [THEN base_enabled, temp_use] + elim!: enabledE simp: angle_def DBPass_def Deq_def) + done + + +(* The plan for proving weak fairness at the higher level is to prove + (0) DBuffer => (Enabled (Deq inp qc out) ~> (Deq inp qc out)) + 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 _(inp,qc,out) + (and therefore DBDeq ~> _(inp,qc,out) trivially holds). -end \ No newline at end of file + Condition (1) is reduced to + (1a) DBuffer => (qc ~= [] /\ q2 = [] ~> q2 ~= []) + by standard leadsto rules (leadsto_classical) and rule Deq_enabledE. + + Both (1a) and (2) are proved from DBuffer's WF conditions by standard + WF reasoning (Lamport's WF1 and WF_leadsto). + The condition WF(Deq inp qc out) follows from (0) by rule leadsto_WF. + + One could use Lamport's WF2 instead. +*) + +(* Condition (1a) *) +lemma DBFair_1a: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2) + --> (qc ~= #[] & q2 = #[] ~> q2 ~= #[])" + apply (rule WF1) + apply (force simp: db_defs) + apply (force simp: angle_def DBPass_def) + apply (force simp: DBPass_enabled [temp_use]) + done + +(* Condition (1) *) +lemma DBFair_1: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBPass)_(inp,mid,out,q1,q2) + --> (Enabled (_(inp,qc,out)) ~> q2 ~= #[])" + apply clarsimp + apply (rule leadsto_classical [temp_use]) + apply (rule DBFair_1a [temp_use, THEN LatticeTransitivity [temp_use]]) + apply assumption+ + apply (rule ImplLeadsto_gen [temp_use]) + apply (force intro!: necT [temp_use] dest!: STL2_gen [temp_use] Deq_enabledE [temp_use] + simp add: Init_defs) + done + +(* Condition (2) *) +lemma DBFair_2: "|- [][DBNext]_(inp,mid,out,q1,q2) & WF(DBDeq)_(inp,mid,out,q1,q2) + --> (q2 ~= #[] ~> DBDeq)" + apply (rule WF_leadsto) + apply (force simp: DBDeq_enabled [temp_use]) + apply (force simp: angle_def) + apply (force simp: db_defs elim!: Stable [temp_use]) + done + +(* High-level fairness *) +lemma DBFair: "|- [][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)" + apply (auto simp del: qc_def intro!: leadsto_WF [temp_use] + DBFair_1 [temp_use, THEN [2] LatticeTransitivity [temp_use]] + DBFair_2 [temp_use, THEN [2] LatticeTransitivity [temp_use]]) + apply (auto intro!: ImplLeadsto_simple [temp_use] + simp: angle_def DBDeq_def Deq_def hd_append [try_rewrite]) + done + +(*** Main theorem ***) +lemma DBuffer_impl_Buffer: "|- DBuffer --> Buffer inp out" + apply (unfold DBuffer_def Buffer_def IBuffer_def) + apply (force intro!: eexI [temp_use] DBInit [temp_use] + DB_step_simulation [THEN STL4, temp_use] DBFair [temp_use]) + done + +end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Inc/Inc.ML --- a/src/HOL/TLA/Inc/Inc.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,244 +0,0 @@ -(* - File: TLA/ex/inc/Inc.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - -Proofs for the "increment" example from SRC79. -*) - -val PsiInv_defs = [PsiInv_def,PsiInv1_def,PsiInv2_def,PsiInv3_def]; -val Psi_defs = [Psi_def,InitPsi_def,N1_def,N2_def,alpha1_def,alpha2_def, - beta1_def,beta2_def,gamma1_def,gamma2_def]; - -val Inc_css = (claset(), simpset()); - -(*** Invariant proof for Psi: "manual" proof proves individual lemmas ***) - -Goal "|- InitPsi --> PsiInv"; -by (auto_tac (Inc_css addsimps2 InitPsi_def::PsiInv_defs)); -qed "PsiInv_Init"; - -Goal "|- alpha1 & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 alpha1_def::PsiInv_defs)); -qed "PsiInv_alpha1"; - -Goal "|- alpha2 & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 alpha2_def::PsiInv_defs)); -qed "PsiInv_alpha2"; - -Goal "|- beta1 & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 beta1_def::PsiInv_defs)); -qed "PsiInv_beta1"; - -Goal "|- beta2 & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 beta2_def::PsiInv_defs)); -qed "PsiInv_beta2"; - -Goal "|- gamma1 & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 gamma1_def::PsiInv_defs)); -qed "PsiInv_gamma1"; - -Goal "|- gamma2 & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 gamma2_def::PsiInv_defs)); -qed "PsiInv_gamma2"; - -Goal "|- unchanged (x,y,sem,pc1,pc2) & $PsiInv --> PsiInv$"; -by (auto_tac (Inc_css addsimps2 PsiInv_defs)); -qed "PsiInv_stutter"; - -Goal "|- Psi --> []PsiInv"; -by (inv_tac (Inc_css addsimps2 [Psi_def]) 1); - by (force_tac (Inc_css addsimps2 [PsiInv_Init, Init_def]) 1); -by (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])); -qed "PsiInv"; - -(* Automatic proof works too, but it make take a while on a slow machine. - More realistic examples require user guidance anyway. - -Goal "|- Psi --> []PsiInv"; -by (auto_inv_tac (simpset() addsimps PsiInv_defs @ Psi_defs) 1); - -*) - -(**** Step simulation ****) - -Goal "|- Psi --> Init InitPhi"; -by (auto_tac (Inc_css addsimps2 [InitPhi_def,Psi_def,InitPsi_def,Init_def])); -qed "Init_sim"; - -Goal "|- Psi --> [][M1 | M2]_(x,y)"; -by (auto_tac (Inc_css addsimps2 [square_def,M1_def,M2_def] @ Psi_defs - addSEs2 [STL4E])); -qed "Step_sim"; - -(**** Proof of fairness ****) - -(* - The goal is to prove Fair_M1 far below, which asserts - |- 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 - temporal premise needs two auxiliary lemmas: - 1. Stuck_at_b: control can only proceed at pc1 = b by executing beta1 - 2. N1_live: the first component will eventually reach b - - Lemma 1 is easy, lemma 2 relies on the invariant, the strong fairness - of the semaphore, and needs auxiliary lemmas that ensure that the second - component will eventually release the semaphore. Most of the proofs of - the auxiliary lemmas are very similar. -*) - -Goal "|- [][(N1 | N2) & ~ beta1]_(x,y,sem,pc1,pc2) --> stable(pc1 = #b)"; -by (auto_tac (Inc_css addSEs2 [Stable,squareE] addsimps2 Psi_defs)); -qed "Stuck_at_b"; - -Goal "|- pc1 = #g --> Enabled (_(x,y,sem,pc1,pc2))"; -by (Clarsimp_tac 1); -by (res_inst_tac [("F","gamma1")] enabled_mono 1); -by (enabled_tac Inc_base 1); - by (force_tac (Inc_css addsimps2 [gamma1_def]) 1); -by (force_tac (Inc_css addsimps2 [angle_def,gamma1_def,N1_def]) 1); -qed "N1_enabled_at_g"; - -Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N1)_(x,y,sem,pc1,pc2) & []#True \ -\ --> (pc1 = #g ~> pc1 = #a)"; -by (rtac SF1 1); -by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1); -by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1); -(* reduce |- []A --> <>Enabled B to |- A --> Enabled B *) -by (auto_tac (Inc_css addSIs2 [InitDmd_gen, N1_enabled_at_g] - addSDs2 [STL2_gen] - addsimps2 [Init_def])); -qed "g1_leadsto_a1"; - -(* symmetrical for N2, and similar for beta2 *) -Goal "|- pc2 = #g --> Enabled (_(x,y,sem,pc1,pc2))"; -by (Clarsimp_tac 1); -by (res_inst_tac [("F","gamma2")] enabled_mono 1); -by (enabled_tac Inc_base 1); - by (force_tac (Inc_css addsimps2 [gamma2_def]) 1); -by (force_tac (Inc_css addsimps2 [angle_def,gamma2_def,N2_def]) 1); -qed "N2_enabled_at_g"; - -Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \ -\ --> (pc2 = #g ~> pc2 = #a)"; -by (rtac SF1 1); -by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1); -by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1); -by (auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_g] - addSDs2 [STL2_gen] - addsimps2 [Init_def])); -qed "g2_leadsto_a2"; - -Goal "|- pc2 = #b --> Enabled (_(x,y,sem,pc1,pc2))"; -by (Clarsimp_tac 1); -by (res_inst_tac [("F","beta2")] enabled_mono 1); -by (enabled_tac Inc_base 1); - by (force_tac (Inc_css addsimps2 [beta2_def]) 1); -by (force_tac (Inc_css addsimps2 [angle_def,beta2_def,N2_def]) 1); -qed "N2_enabled_at_b"; - -Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \ -\ --> (pc2 = #b ~> pc2 = #g)"; -by (rtac SF1 1); -by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1); -by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1); -by (auto_tac (Inc_css addSIs2 [InitDmd_gen, N2_enabled_at_b] - addSDs2 [STL2_gen] - addsimps2 [Init_def])); -qed "b2_leadsto_g2"; - -(* Combine above lemmas: the second component will eventually reach pc2 = a *) -Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True \ -\ --> (pc2 = #a | pc2 = #b | pc2 = #g ~> pc2 = #a)"; -by (auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro])); -by (rtac (temp_use LatticeReflexivity) 1); -by (rtac (temp_use LatticeTransitivity) 1); -by (auto_tac (Inc_css addSIs2 [b2_leadsto_g2,g2_leadsto_a2])); -qed "N2_leadsto_a"; - -(* Get rid of disjunction on the left-hand side of ~> above. *) -Goal "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) \ -\ --> <>(pc2 = #a)"; -by (auto_tac (Inc_css addsimps2 Init_defs - addSIs2 [(temp_use N2_leadsto_a) - RSN(2, (temp_use leadsto_init))])); -by (case_tac "pc2 (st1 sigma)" 1); -by Auto_tac; -qed "N2_live"; - -(* Now prove that the first component will eventually reach pc1 = b from pc1 = a *) - -Goal "|- pc2 = #a & (PsiInv & pc1 = #a) --> Enabled (_(x,y,sem,pc1,pc2))"; -by (Clarsimp_tac 1); -by (res_inst_tac [("F","alpha1")] enabled_mono 1); -by (enabled_tac Inc_base 1); - by (force_tac (Inc_css addsimps2 (alpha1_def::PsiInv_defs)) 1); -by (force_tac (Inc_css addsimps2 [angle_def,alpha1_def,N1_def]) 1); -qed "N1_enabled_at_both_a"; - -Goal "|- []($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)"; -by (rtac SF1 1); -by (action_simp_tac (simpset() addsimps Psi_defs) [] [squareE] 1); -by (action_simp_tac (simpset() addsimps angle_def::Psi_defs) [] [] 1); -by (clarsimp_tac (Inc_css addSIs2 [N1_enabled_at_both_a RS (temp_use DmdImpl)]) 1); -by (auto_tac (Inc_css addSIs2 [BoxDmd2_simple, N2_live] - addsimps2 split_box_conj::more_temp_simps)); -qed "a1_leadsto_b1"; - -(* Combine the leadsto properties for N1: it will arrive at pc1 = b *) - -Goal "|- []($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)"; -by (auto_tac (Inc_css addSIs2 [LatticeDisjunctionIntro])); -by (rtac (temp_use LatticeReflexivity) 1); -by (rtac (temp_use LatticeTransitivity) 1); -by (auto_tac (Inc_css addSIs2 [a1_leadsto_b1,g1_leadsto_a1] - addsimps2 [split_box_conj])); -qed "N1_leadsto_b"; - -Goal "|- []($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)"; -by (auto_tac (Inc_css addsimps2 Init_defs - addSIs2 [(temp_use N1_leadsto_b) - RSN(2, temp_use leadsto_init)])); -by (case_tac "pc1 (st1 sigma)" 1); -by Auto_tac; -qed "N1_live"; - -Goal "|- pc1 = #b --> Enabled (_(x,y,sem,pc1,pc2))"; -by (Clarsimp_tac 1); -by (res_inst_tac [("F","beta1")] enabled_mono 1); -by (enabled_tac Inc_base 1); - by (force_tac (Inc_css addsimps2 [beta1_def]) 1); -by (force_tac (Inc_css addsimps2 [angle_def,beta1_def,N1_def]) 1); -qed "N1_enabled_at_b"; - -(* Now assemble the bits and pieces to prove that Psi is fair. *) - -Goal "|- []($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)"; -by (res_inst_tac [("B","beta1"),("P","PRED pc1 = #b")] SF2 1); - (* action premises *) -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 [N1_enabled_at_b]) 1); - (* temporal premise: use previous lemmas and simple TL *) -by (force_tac (Inc_css addSIs2 [DmdStable, N1_live,Stuck_at_b] - addEs2 [STL4E] addsimps2 [square_def]) 1); -qed "Fair_M1_lemma"; - -Goal "|- Psi --> WF(M1)_(x,y)"; -by (auto_tac (Inc_css addSIs2 [SFImplWF, Fair_M1_lemma, PsiInv] - addsimps2 [Psi_def,split_box_conj]@more_temp_simps)); -qed "Fair_M1"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Inc/Inc.thy --- a/src/HOL/TLA/Inc/Inc.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Inc/Inc.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,9 +3,6 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich - - Theory Name: Inc - Logic Image: TLA *) header {* Lamport's "increment" example *} @@ -85,6 +82,234 @@ PsiInv3_def: "PsiInv3 == PRED sem = # 0 & pc2 = #a & (pc1 = #b | pc1 = #g)" PsiInv_def: "PsiInv == PRED (PsiInv1 | PsiInv2 | PsiInv3)" -ML {* use_legacy_bindings (the_context ()) *} + +lemmas PsiInv_defs = PsiInv_def PsiInv1_def PsiInv2_def PsiInv3_def +lemmas Psi_defs = Psi_def InitPsi_def N1_def N2_def alpha1_def alpha2_def + beta1_def beta2_def gamma1_def gamma2_def + + +(*** Invariant proof for Psi: "manual" proof proves individual lemmas ***) + +lemma PsiInv_Init: "|- InitPsi --> PsiInv" + by (auto simp: InitPsi_def PsiInv_defs) + +lemma PsiInv_alpha1: "|- alpha1 & $PsiInv --> PsiInv$" + by (auto simp: alpha1_def PsiInv_defs) + +lemma PsiInv_alpha2: "|- alpha2 & $PsiInv --> PsiInv$" + by (auto simp: alpha2_def PsiInv_defs) + +lemma PsiInv_beta1: "|- beta1 & $PsiInv --> PsiInv$" + by (auto simp: beta1_def PsiInv_defs) + +lemma PsiInv_beta2: "|- beta2 & $PsiInv --> PsiInv$" + by (auto simp: beta2_def PsiInv_defs) + +lemma PsiInv_gamma1: "|- gamma1 & $PsiInv --> PsiInv$" + by (auto simp: gamma1_def PsiInv_defs) + +lemma PsiInv_gamma2: "|- gamma2 & $PsiInv --> PsiInv$" + by (auto simp: gamma2_def PsiInv_defs) + +lemma PsiInv_stutter: "|- unchanged (x,y,sem,pc1,pc2) & $PsiInv --> PsiInv$" + by (auto simp: PsiInv_defs) + +lemma PsiInv: "|- Psi --> []PsiInv" + apply (tactic {* inv_tac (clasimpset () addsimps2 [thm "Psi_def"]) 1 *}) + apply (force simp: PsiInv_Init [try_rewrite] Init_def) + apply (auto intro: PsiInv_alpha1 [try_rewrite] PsiInv_alpha2 [try_rewrite] + PsiInv_beta1 [try_rewrite] PsiInv_beta2 [try_rewrite] PsiInv_gamma1 [try_rewrite] + PsiInv_gamma2 [try_rewrite] PsiInv_stutter [try_rewrite] + simp add: square_def N1_def N2_def) + done + +(* Automatic proof works too, but it make take a while on a slow machine. + More realistic examples require user guidance anyway. +*) + +lemma "|- Psi --> []PsiInv" + by (tactic {* auto_inv_tac (simpset() addsimps (thms "PsiInv_defs" @ thms "Psi_defs")) 1 *}) + + +(**** Step simulation ****) + +lemma Init_sim: "|- Psi --> Init InitPhi" + by (auto simp: InitPhi_def Psi_def InitPsi_def Init_def) + +lemma Step_sim: "|- Psi --> [][M1 | M2]_(x,y)" + by (auto simp: square_def M1_def M2_def Psi_defs elim!: STL4E [temp_use]) + + +(**** Proof of fairness ****) + +(* + The goal is to prove Fair_M1 far below, which asserts + |- 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 + temporal premise needs two auxiliary lemmas: + 1. Stuck_at_b: control can only proceed at pc1 = b by executing beta1 + 2. N1_live: the first component will eventually reach b + + Lemma 1 is easy, lemma 2 relies on the invariant, the strong fairness + of the semaphore, and needs auxiliary lemmas that ensure that the second + component will eventually release the semaphore. Most of the proofs of + the auxiliary lemmas are very similar. +*) + +lemma Stuck_at_b: "|- [][(N1 | N2) & ~ beta1]_(x,y,sem,pc1,pc2) --> stable(pc1 = #b)" + by (auto elim!: Stable squareE simp: Psi_defs) + +lemma N1_enabled_at_g: "|- pc1 = #g --> Enabled (_(x,y,sem,pc1,pc2))" + apply clarsimp + apply (rule_tac F = gamma1 in enabled_mono) + apply (tactic {* enabled_tac (thm "Inc_base") 1 *}) + apply (force simp: gamma1_def) + apply (force simp: angle_def gamma1_def N1_def) + done + +lemma g1_leadsto_a1: + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N1)_(x,y,sem,pc1,pc2) & []#True + --> (pc1 = #g ~> pc1 = #a)" + apply (rule SF1) + apply (tactic + {* action_simp_tac (simpset () addsimps (thms "Psi_defs")) [] [thm "squareE"] 1 *}) + apply (tactic + {* action_simp_tac (simpset () addsimps (thm "angle_def" :: thms "Psi_defs")) [] [] 1 *}) + (* reduce |- []A --> <>Enabled B to |- A --> Enabled B *) + apply (auto intro!: InitDmd_gen [temp_use] N1_enabled_at_g [temp_use] + dest!: STL2_gen [temp_use] simp: Init_def) + done + +(* symmetrical for N2, and similar for beta2 *) +lemma N2_enabled_at_g: "|- pc2 = #g --> Enabled (_(x,y,sem,pc1,pc2))" + apply clarsimp + apply (rule_tac F = gamma2 in enabled_mono) + apply (tactic {* enabled_tac (thm "Inc_base") 1 *}) + apply (force simp: gamma2_def) + apply (force simp: angle_def gamma2_def N2_def) + done + +lemma g2_leadsto_a2: + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True + --> (pc2 = #g ~> pc2 = #a)" + apply (rule SF1) + apply (tactic {* action_simp_tac (simpset () addsimps (thms "Psi_defs")) [] [thm "squareE"] 1 *}) + apply (tactic {* action_simp_tac (simpset () addsimps (thm "angle_def" :: thms "Psi_defs")) + [] [] 1 *}) + apply (auto intro!: InitDmd_gen [temp_use] N2_enabled_at_g [temp_use] + dest!: STL2_gen [temp_use] simp add: Init_def) + done + +lemma N2_enabled_at_b: "|- pc2 = #b --> Enabled (_(x,y,sem,pc1,pc2))" + apply clarsimp + apply (rule_tac F = beta2 in enabled_mono) + apply (tactic {* enabled_tac (thm "Inc_base") 1 *}) + apply (force simp: beta2_def) + apply (force simp: angle_def beta2_def N2_def) + done + +lemma b2_leadsto_g2: + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) & []#True + --> (pc2 = #b ~> pc2 = #g)" + apply (rule SF1) + apply (tactic + {* action_simp_tac (simpset () addsimps (thms "Psi_defs")) [] [thm "squareE"] 1 *}) + apply (tactic + {* action_simp_tac (simpset () addsimps (thm "angle_def" :: thms "Psi_defs")) [] [] 1 *}) + apply (auto intro!: InitDmd_gen [temp_use] N2_enabled_at_b [temp_use] + dest!: STL2_gen [temp_use] simp: Init_def) + done + +(* Combine above lemmas: the second component will eventually reach pc2 = a *) +lemma N2_leadsto_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)" + apply (auto intro!: LatticeDisjunctionIntro [temp_use]) + apply (rule LatticeReflexivity [temp_use]) + apply (rule LatticeTransitivity [temp_use]) + apply (auto intro!: b2_leadsto_g2 [temp_use] g2_leadsto_a2 [temp_use]) + done + +(* Get rid of disjunction on the left-hand side of ~> above. *) +lemma N2_live: + "|- [][(N1 | N2) & ~beta1]_(x,y,sem,pc1,pc2) & SF(N2)_(x,y,sem,pc1,pc2) + --> <>(pc2 = #a)" + apply (auto simp: Init_defs intro!: N2_leadsto_a [temp_use, THEN [2] leadsto_init [temp_use]]) + apply (case_tac "pc2 (st1 sigma)") + apply auto + done + +(* Now prove that the first component will eventually reach pc1 = b from pc1 = a *) + +lemma N1_enabled_at_both_a: + "|- pc2 = #a & (PsiInv & pc1 = #a) --> Enabled (_(x,y,sem,pc1,pc2))" + apply clarsimp + apply (rule_tac F = alpha1 in enabled_mono) + apply (tactic {* enabled_tac (thm "Inc_base") 1 *}) + apply (force simp: alpha1_def PsiInv_defs) + apply (force simp: angle_def alpha1_def N1_def) + done + +lemma a1_leadsto_b1: + "|- []($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)" + apply (rule SF1) + apply (tactic {* action_simp_tac (simpset () addsimps thms "Psi_defs") [] [thm "squareE"] 1 *}) + apply (tactic + {* action_simp_tac (simpset () addsimps (thm "angle_def" :: thms "Psi_defs")) [] [] 1 *}) + apply (clarsimp intro!: N1_enabled_at_both_a [THEN DmdImpl [temp_use]]) + apply (auto intro!: BoxDmd2_simple [temp_use] N2_live [temp_use] + simp: split_box_conj more_temp_simps) + done + +(* Combine the leadsto properties for N1: it will arrive at pc1 = b *) + +lemma N1_leadsto_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)" + apply (auto intro!: LatticeDisjunctionIntro [temp_use]) + apply (rule LatticeReflexivity [temp_use]) + apply (rule LatticeTransitivity [temp_use]) + apply (auto intro!: a1_leadsto_b1 [temp_use] g1_leadsto_a1 [temp_use] + simp: split_box_conj) + done + +lemma N1_live: "|- []($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)" + apply (auto simp: Init_defs intro!: N1_leadsto_b [temp_use, THEN [2] leadsto_init [temp_use]]) + apply (case_tac "pc1 (st1 sigma)") + apply auto + done + +lemma N1_enabled_at_b: "|- pc1 = #b --> Enabled (_(x,y,sem,pc1,pc2))" + apply clarsimp + apply (rule_tac F = beta1 in enabled_mono) + apply (tactic {* enabled_tac (thm "Inc_base") 1 *}) + apply (force simp: beta1_def) + apply (force simp: angle_def beta1_def N1_def) + done + +(* Now assemble the bits and pieces to prove that Psi is fair. *) + +lemma Fair_M1_lemma: "|- []($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)" + apply (rule_tac B = beta1 and P = "PRED pc1 = #b" in SF2) + (* action premises *) + apply (force simp: angle_def M1_def beta1_def) + apply (force simp: angle_def Psi_defs) + apply (force elim!: N1_enabled_at_b [temp_use]) + (* temporal premise: use previous lemmas and simple TL *) + apply (force intro!: DmdStable [temp_use] N1_live [temp_use] Stuck_at_b [temp_use] + elim: STL4E [temp_use] simp: square_def) + done + +lemma Fair_M1: "|- Psi --> WF(M1)_(x,y)" + by (auto intro!: SFImplWF [temp_use] Fair_M1_lemma [temp_use] PsiInv [temp_use] + simp: Psi_def split_box_conj [temp_use] more_temp_simps) end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Init.ML --- a/src/HOL/TLA/Init.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,45 +0,0 @@ - -(* $Id$ *) - -local - fun prover s = prove_goal (the_context ()) 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 17098171d46a -r 6f79647cf536 src/HOL/TLA/Init.thy --- a/src/HOL/TLA/Init.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Init.thy Sat Dec 02 02:52:02 2006 +0100 @@ -4,9 +4,6 @@ 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). *) @@ -43,6 +40,40 @@ fw_stp_def: "first_world == st1" fw_act_def: "first_world == %sigma. (st1 sigma, st2 sigma)" -ML {* use_legacy_bindings (the_context ()) *} +lemma const_simps [int_rewrite, simp]: + "|- (Init #True) = #True" + "|- (Init #False) = #False" + by (auto simp: Init_def) + +lemma Init_simps [int_rewrite]: + "!!F. |- (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))" + by (auto simp: Init_def) + +lemma Init_stp_act: "|- (Init $P) = (Init P)" + by (auto simp add: Init_def fw_act_def fw_stp_def) + +lemmas Init_simps = Init_stp_act [int_rewrite] Init_simps +lemmas Init_stp_act_rev = Init_stp_act [int_rewrite, symmetric] + +lemma Init_temp: "|- (Init F) = F" + by (auto simp add: Init_def fw_temp_def) + +lemmas Init_simps = Init_temp [int_rewrite] Init_simps + +(* Trivial instances of the definitions that avoid introducing lambda expressions. *) +lemma Init_stp: "(sigma |= Init P) = P (st1 sigma)" + by (simp add: Init_def fw_stp_def) + +lemma Init_act: "(sigma |= Init A) = A (st1 sigma, st2 sigma)" + by (simp add: Init_def fw_act_def) + +lemmas Init_defs = Init_stp Init_act Init_temp [int_use] end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Intensional.ML --- a/src/HOL/TLA/Intensional.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,125 +0,0 @@ -(* - File: Intensional.ML - ID: $Id$ - Author: Stephan Merz - 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,unl_Rex1]; - -Goalw [Valid_def,unl_lift2] "|- x=y ==> (x==y)"; -by (rtac eq_reflection 1); -by (rtac ext 1); -by (etac spec 1); -qed "inteq_reflection"; - -val [prem] = goalw (the_context ()) [Valid_def] "(!!w. w |= A) ==> |- A"; -by (REPEAT (resolve_tac [allI,prem] 1)); -qed "intI"; - -Goalw [Valid_def] "|- A ==> w |= A"; -by (etac spec 1); -qed "intD"; - -(** Lift usual HOL simplifications to "intensional" level. **) -local - -fun prover s = (prove_goal (the_context ()) s - (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", - "|- (! x. P) = P", "|- (? x. P) = P", - "|- (~Q --> ~P) = (P --> Q)", - "|- (P|Q --> R) = ((P-->R)&(Q-->R))" ] -end; - -Goal "|- #True"; -by (simp_tac (simpset() addsimps [Valid_def,unl_con]) 1); -qed "TrueW"; - -Addsimps (TrueW::intensional_rews); -Addsimps int_simps; -AddSIs [intI]; -AddDs [intD]; - - -(* ======== Functions to "unlift" intensional implications into HOL rules ====== *) - -(* Basic unlifting introduces a parameter "w" and applies basic rewrites, e.g. - |- F = G becomes F w = G w - |- F --> G becomes F w --> G w -*) - -fun int_unlift th = - rewrite_rule intensional_rews ((th RS intD) handle _ => th); - -(* 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)); - -(* flattening turns "-->" into "==>" and eliminates conjunctions in the - antecedent. For example, - - P & Q --> (R | S --> T) becomes [| P; Q; R | S |] ==> T - - 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]) - - (* 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 - - 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; - -fun int_use th = - case (concl_of th) of - Const _ $ (Const ("Intensional.Valid", _) $ _) => - ((flatten (int_unlift th)) handle _ => th) - | _ => th; - -(* ========================================================================= *) - -Goalw [Valid_def] "|- (~(! x. F x)) = (? x. ~F x)"; -by (Simp_tac 1); -qed "Not_Rall"; - -Goalw [Valid_def] "|- (~ (? x. F x)) = (! x. ~ F x)"; -by (Simp_tac 1); -qed "Not_Rex"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Intensional.thy --- a/src/HOL/TLA/Intensional.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Intensional.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,13 +3,10 @@ ID: $Id$ Author: Stephan Merz Copyright: 1998 University of Munich - - Theory Name: Intensional - Logic Image: HOL +*) -Define a framework for "intensional" (possible-world based) logics -on top of HOL, with lifting of constants and functions. -*) +header {* A framework for "intensional" (possible-world based) logics + on top of HOL, with lifting of constants and functions *} theory Intensional imports Main @@ -188,6 +185,136 @@ unl_Rex: "w |= EX x. A x == EX x. (w |= A x)" unl_Rex1: "w |= EX! x. A x == EX! x. (w |= A x)" -ML {* use_legacy_bindings (the_context ()) *} + +subsection {* Lemmas and tactics for "intensional" logics. *} + +lemmas intensional_rews [simp] = + unl_con unl_lift unl_lift2 unl_lift3 unl_Rall unl_Rex unl_Rex1 + +lemma inteq_reflection: "|- x=y ==> (x==y)" + apply (unfold Valid_def unl_lift2) + apply (rule eq_reflection) + apply (rule ext) + apply (erule spec) + done + +lemma intI [intro!]: "(!!w. w |= A) ==> |- A" + apply (unfold Valid_def) + apply (rule allI) + apply (erule meta_spec) + done + +lemma intD [dest]: "|- A ==> w |= A" + apply (unfold Valid_def) + apply (erule spec) + done + +(** Lift usual HOL simplifications to "intensional" level. **) + +lemma int_simps: + "|- (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))" + apply (unfold Valid_def intensional_rews) + apply blast+ + done + +declare int_simps [THEN inteq_reflection, simp] + +lemma TrueW [simp]: "|- #True" + by (simp add: Valid_def unl_con) + + + +(* ======== Functions to "unlift" intensional implications into HOL rules ====== *) + +ML {* + +local + val intD = thm "intD"; + val inteq_reflection = thm "inteq_reflection"; + val intensional_rews = thms "intensional_rews"; +in + +(* Basic unlifting introduces a parameter "w" and applies basic rewrites, e.g. + |- F = G becomes F w = G w + |- F --> G becomes F w --> G w +*) + +fun int_unlift th = + rewrite_rule intensional_rews (th RS intD handle THM _ => th); + +(* 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)) + +(* flattening turns "-->" into "==>" and eliminates conjunctions in the + antecedent. For example, + + P & Q --> (R | S --> T) becomes [| P; Q; R | S |] ==> T + + 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]) + + (* 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 THM _ => hmatch (n-1) + in hmatch (nprems_of thb) end + + fun hflatten t = + case (concl_of t) of + Const _ $ (Const ("op -->", _) $ _ $ _) => hflatten (t RS mp) + | _ => (hflatten (matchsome conjI t)) handle THM _ => zero_var_indexes t + in + hflatten t + end + +fun int_use th = + case (concl_of th) of + Const _ $ (Const ("Intensional.Valid", _) $ _) => + (flatten (int_unlift th) handle THM _ => th) + | _ => th end +*} + +setup {* + Attrib.add_attributes [ + ("int_unlift", Attrib.no_args (Thm.rule_attribute (K int_unlift)), ""), + ("int_rewrite", Attrib.no_args (Thm.rule_attribute (K int_rewrite)), ""), + ("flatten", Attrib.no_args (Thm.rule_attribute (K flatten)), ""), + ("int_use", Attrib.no_args (Thm.rule_attribute (K int_use)), "")] +*} + +lemma Not_Rall: "|- (~(! x. F x)) = (? x. ~F x)" + by (simp add: Valid_def) + +lemma Not_Rex: "|- (~ (? x. F x)) = (! x. ~ F x)" + by (simp add: Valid_def) + +end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MIsafe.ML --- a/src/HOL/TLA/Memory/MIsafe.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,341 +0,0 @@ -(* - File: MIsafe.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: Lower-level lemmas about memory implementation (safety) -*) - -(* ========================= Lemmas about values ========================= *) - -(* RPCFailure notin MemVals U {OK,BadArg} *) - -Goalw [MVOKBA_def] "MVOKBA x ==> x ~= RPCFailure"; -by Auto_tac; -qed "MVOKBAnotRF"; - -(* NotAResult notin MemVals U {OK,BadArg,RPCFailure} *) - -Goalw [MVOKBARF_def] "MVOKBARF x ==> x ~= NotAResult"; -by Auto_tac; -qed "MVOKBARFnotNR"; - -(* ================ Si's are mutually exclusive ================================ *) -(* Si and Sj are mutually exclusive for i # j. This helps to simplify the big - conditional in the definition of resbar when doing the step-simulation proof. - We prove a weaker result, which suffices for our purposes: - Si implies (not Sj), for j S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p & \ -\ ~S4 rmhist p & ~S5 rmhist p & ~S6 rmhist p"; -by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, - S3_def, S4_def, S5_def, S6_def])); -qed "S1_excl"; -*) - -Goal "|- S2 rmhist p --> S2 rmhist p & ~S1 rmhist p"; -by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def])); -qed "S2_excl"; - -Goal "|- S3 rmhist p --> S3 rmhist p & ~S1 rmhist p & ~S2 rmhist p"; -by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, S3_def])); -qed "S3_excl"; - -Goal "|- S4 rmhist p --> S4 rmhist p & ~S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p"; -by (auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def])); -qed "S4_excl"; - -Goal "|- S5 rmhist p --> S5 rmhist p & ~S1 rmhist p & ~S2 rmhist p \ -\ & ~S3 rmhist p & ~S4 rmhist p"; -by (auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def])); -qed "S5_excl"; - -Goal "|- S6 rmhist p --> S6 rmhist p & ~S1 rmhist p & ~S2 rmhist p \ -\ & ~S3 rmhist p & ~S4 rmhist p & ~S5 rmhist p"; -by (auto_tac (MI_css addsimps2 [S_def,S1_def,S2_def,S3_def,S4_def,S5_def,S6_def])); -qed "S6_excl"; - - -(* ==================== Lemmas about the environment ============================== *) - -Goal "|- $(Calling memCh p) --> ~ENext p"; -by (auto_tac (MI_css addsimps2 [ENext_def,Call_def])); -qed "Envbusy"; - -(* ==================== 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 - determine which component actions are possible and what state they result in. -*) - -(* ------------------------------ State S1 ---------------------------------------- *) - -Goal "|- ENext p & $(S1 rmhist p) & unchanged (c p, r p, m p, rmhist!p) \ -\ --> (S2 rmhist p)$"; -by (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); -qed "S1Env"; - -Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S1 rmhist p) --> unchanged (c p)"; -by (auto_tac (MI_fast_css addSDs2 [MClkidle] addsimps2 [S_def,S1_def])); -qed "S1ClerkUnch"; - -Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $(S1 rmhist p) --> unchanged (r p)"; -by (auto_tac (MI_fast_css addSDs2 [RPCidle] addsimps2 [S_def,S1_def])); -qed "S1RPCUnch"; - -Goal "|- [RNext rmCh mm ires p]_(m p) & $(S1 rmhist p) --> unchanged (m p)"; -by (auto_tac (MI_fast_css addSDs2 [Memoryidle] addsimps2 [S_def,S1_def])); -qed "S1MemUnch"; - -Goal "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S1 rmhist p)\ -\ --> unchanged (rmhist!p)"; -by (action_simp_tac (simpset() addsimps [HNext_def, S_def, S1_def, MemReturn_def, - RPCFail_def,MClkReply_def,Return_def]) - [] [squareE] 1); -qed "S1Hist"; - -(* ------------------------------ State S2 ---------------------------------------- *) - -Goal "|- [ENext p]_(e p) & $(S2 rmhist p) --> unchanged (e p)"; -by (auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S2_def])); -qed "S2EnvUnch"; - -Goal "|- MClkNext memCh crCh cst p & $(S2 rmhist p) --> MClkFwd memCh crCh cst p"; -by (auto_tac (MI_css addsimps2 [MClkNext_def,MClkRetry_def,MClkReply_def, - S_def,S2_def])); -qed "S2Clerk"; - -Goal "|- $(S2 rmhist p) & MClkFwd memCh crCh cst p\ -\ & unchanged (e p, r p, m p, rmhist!p) \ -\ --> (S3 rmhist p)$"; -by (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); -qed "S2Forward"; - -Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $(S2 rmhist p) --> unchanged (r p)"; -by (auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [RPCidle])); -qed "S2RPCUnch"; - -Goal "|- [RNext rmCh mm ires p]_(m p) & $(S2 rmhist p) --> unchanged (m p)"; -by (auto_tac (MI_css addsimps2 [S_def,S2_def] addSDs2 [Memoryidle])); -qed "S2MemUnch"; - -Goal "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S2 rmhist p)\ -\ --> unchanged (rmhist!p)"; -by (auto_tac (MI_fast_css - addsimps2 [HNext_def,MemReturn_def, - RPCFail_def,MClkReply_def,Return_def,S_def,S2_def])); -qed "S2Hist"; - -(* ------------------------------ State S3 ---------------------------------------- *) - -Goal "|- [ENext p]_(e p) & $(S3 rmhist p) --> unchanged (e p)"; -by (auto_tac (MI_css addSDs2 [Envbusy] addsimps2 [S_def,S3_def])); -qed "S3EnvUnch"; - -Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S3 rmhist p) --> unchanged (c p)"; -by (auto_tac (MI_css addSDs2 [MClkbusy] addsimps2 [square_def,S_def,S3_def])); -qed "S3ClerkUnch"; - -Goal "|- S3 rmhist p --> IsLegalRcvArg>"; -by (auto_tac (MI_css addsimps2 [IsLegalRcvArg_def,MClkRelayArg_def,S_def,S3_def])); -qed "S3LegalRcvArg"; - -Goal "|- RPCNext crCh rmCh rst p & $(S3 rmhist p) \ -\ --> RPCFwd crCh rmCh rst p | RPCFail crCh rmCh rst p"; -by (Clarsimp_tac 1); -by (forward_tac [action_use S3LegalRcvArg] 1); -by (auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCReply_def,S_def,S3_def])); -qed "S3RPC"; - -Goal "|- RPCFwd crCh rmCh rst p & HNext rmhist p & $(S3 rmhist p)\ -\ & unchanged (e p, c p, m p) \ -\ --> (S4 rmhist p)$ & unchanged (rmhist!p)"; -by (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); -qed "S3Forward"; - -Goal "|- RPCFail crCh rmCh rst p & $(S3 rmhist p) & HNext rmhist p\ -\ & unchanged (e p, c p, m p) \ -\ --> (S6 rmhist p)$"; -by (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); -qed "S3Fail"; - -Goal "|- [RNext rmCh mm ires p]_(m p) & $(S3 rmhist p) --> unchanged (m p)"; -by (auto_tac (MI_css addsimps2 [S_def,S3_def] addSDs2 [Memoryidle])); -qed "S3MemUnch"; - -Goal "|- HNext rmhist p & $(S3 rmhist p) & unchanged (r p) --> unchanged (rmhist!p)"; -by (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])); -qed "S3Hist"; - -(* ------------------------------ State S4 ---------------------------------------- *) - -Goal "|- [ENext p]_(e p) & $(S4 rmhist p) --> unchanged (e p)"; -by (auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [Envbusy])); -qed "S4EnvUnch"; - -Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S4 rmhist p) --> unchanged (c p)"; -by (auto_tac (MI_css addsimps2 [S_def,S4_def] addSDs2 [MClkbusy])); -qed "S4ClerkUnch"; - -Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $(S4 rmhist p) --> unchanged (r p)"; -by (auto_tac (MI_fast_css addsimps2 [S_def,S4_def] addSDs2 [RPCbusy])); -qed "S4RPCUnch"; - -Goal "|- 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)"; -by (action_simp_tac - (simpset() addsimps [ReadInner_def,GoodRead_def, BadRead_def,HNext_def, - MemReturn_def, RPCFail_def,MClkReply_def,Return_def, - e_def,c_def,r_def,rtrner_def,caller_def,MVNROKBA_def, - S_def,S4_def,RdRequest_def,Calling_def,MemInv_def]) - [] [] 1); -qed "S4ReadInner"; - -Goal "|- 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)"; -by (auto_tac (MI_css addsimps2 [Read_def] addSDs2 [S4ReadInner])); -qed "S4Read"; - -Goal "|- 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)"; -by (action_simp_tac - (simpset() addsimps [WriteInner_def,GoodWrite_def, BadWrite_def,HNext_def, - MemReturn_def,RPCFail_def,MClkReply_def,Return_def, - e_def,c_def,r_def,rtrner_def,caller_def,MVNROKBA_def, - S_def,S4_def,WrRequest_def,Calling_def]) - [] [] 1); -qed "S4WriteInner"; - -Goal "|- 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)"; -by (auto_tac (MI_css addsimps2 [Write_def] addSDs2 [S4WriteInner])); -qed "S4Write"; - -Goal "|- $ImpInv rmhist p & Write rmCh mm ires p l --> $S4 rmhist p"; -by (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])); -qed "WriteS4"; - -Goal "|- MemReturn rmCh ires p & $S4 rmhist p & unchanged (e p, c p, r p)\ -\ & HNext rmhist p \ -\ --> (S5 rmhist p)$"; -by (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])); -qed "S4Return"; - -Goal "|- HNext rmhist p & $S4 rmhist p & (m p)$ = $(m p) --> (rmhist!p)$ = $(rmhist!p)"; -by (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])); -qed "S4Hist"; - -(* ------------------------------ State S5 ---------------------------------------- *) - -Goal "|- [ENext p]_(e p) & $(S5 rmhist p) --> unchanged (e p)"; -by (auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Envbusy])); -qed "S5EnvUnch"; - -Goal "|- [MClkNext memCh crCh cst p]_(c p) & $(S5 rmhist p) --> unchanged (c p)"; -by (auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [MClkbusy])); -qed "S5ClerkUnch"; - -Goal "|- RPCNext crCh rmCh rst p & $(S5 rmhist p) \ -\ --> RPCReply crCh rmCh rst p | RPCFail crCh rmCh rst p"; -by (auto_tac (MI_css addsimps2 [RPCNext_def,RPCReject_def,RPCFwd_def,S_def,S5_def])); -qed "S5RPC"; - -Goal "|- RPCReply crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p)\ -\ --> (S6 rmhist p)$"; -by (action_simp_tac - (simpset() addsimps [RPCReply_def,Return_def,e_def,c_def,m_def, - MVOKBA_def,MVOKBARF_def,caller_def,rtrner_def, - S_def,S5_def,S6_def,Calling_def]) - [] [] 1); -qed "S5Reply"; - -Goal "|- RPCFail crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) \ -\ --> (S6 rmhist p)$"; -by (action_simp_tac - (simpset() addsimps [RPCFail_def,Return_def,e_def,c_def,m_def, - MVOKBARF_def,caller_def,rtrner_def, - S_def,S5_def,S6_def,Calling_def]) - [] [] 1); -qed "S5Fail"; - -Goal "|- [RNext rmCh mm ires p]_(m p) & $(S5 rmhist p) --> unchanged (m p)"; -by (auto_tac (MI_css addsimps2 [S_def,S5_def] addSDs2 [Memoryidle])); -qed "S5MemUnch"; - -Goal "|- [HNext rmhist p]_(c p, r p, m p, rmhist!p) & $(S5 rmhist p)\ -\ --> (rmhist!p)$ = $(rmhist!p)"; -by (auto_tac (MI_fast_css - addsimps2 [HNext_def,MemReturn_def, - RPCFail_def,MClkReply_def,Return_def,S_def,S5_def])); -qed "S5Hist"; - -(* ------------------------------ State S6 ---------------------------------------- *) - -Goal "|- [ENext p]_(e p) & $(S6 rmhist p) --> unchanged (e p)"; -by (auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Envbusy])); -qed "S6EnvUnch"; - -Goal "|- MClkNext memCh crCh cst p & $(S6 rmhist p) \ -\ --> MClkRetry memCh crCh cst p | MClkReply memCh crCh cst p"; -by (auto_tac (MI_css addsimps2 [MClkNext_def,MClkFwd_def,S_def,S6_def])); -qed "S6Clerk"; - -Goal "|- MClkRetry memCh crCh cst p & HNext rmhist p & $S6 rmhist p\ -\ & unchanged (e p,r p,m p) \ -\ --> (S3 rmhist p)$ & unchanged (rmhist!p)"; -by (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); -qed "S6Retry"; - -Goal "|- MClkReply memCh crCh cst p & HNext rmhist p & $S6 rmhist p\ -\ & unchanged (e p,r p,m p) \ -\ --> (S1 rmhist p)$"; -by (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); -qed "S6Reply"; - -Goal "|- [RPCNext crCh rmCh rst p]_(r p) & $S6 rmhist p --> unchanged (r p)"; -by (auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [RPCidle])); -qed "S6RPCUnch"; - -Goal "|- [RNext rmCh mm ires p]_(m p) & $(S6 rmhist p) --> unchanged (m p)"; -by (auto_tac (MI_css addsimps2 [S_def,S6_def] addSDs2 [Memoryidle])); -qed "S6MemUnch"; - -Goal "|- HNext rmhist p & $S6 rmhist p & (c p)$ = $(c p) --> (rmhist!p)$ = $(rmhist!p)"; -by (auto_tac (MI_css addsimps2 [HNext_def,MClkReply_def,Return_def,c_def,rtrner_def, - S_def,S6_def,Calling_def])); -qed "S6Hist"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemClerk.ML --- a/src/HOL/TLA/Memory/MemClerk.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,58 +0,0 @@ -(* - File: MemClerk.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: Memory clerk specification (theorems and proofs) -*) - -val MC_action_defs = - [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. -*) -Goal "|- ~$Calling send p & $(cst!p) = #clkA --> ~MClkNext send rcv cst p"; -by (auto_tac (mem_css addsimps2 (Return_def::MC_action_defs))); -qed "MClkidle"; - -Goal "|- $Calling rcv p --> ~MClkNext send rcv cst p"; -by (auto_tac (mem_css addsimps2 (Call_def::MC_action_defs))); -qed "MClkbusy"; - -(* Enabledness of actions *) - -Goal "!!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)"; -by (action_simp_tac (simpset() addsimps [MClkFwd_def,Call_def,caller_def,rtrner_def]) - [exI] [base_enabled,Pair_inject] 1); -qed "MClkFwd_enabled"; - -Goal "|- Enabled (MClkFwd send rcv cst p) --> \ -\ Enabled (_(cst!p, rtrner send!p, caller rcv!p))"; -by (auto_tac (mem_css addSEs2 [enabled_mono] addsimps2 [angle_def,MClkFwd_def])); -qed "MClkFwd_ch_enabled"; - -Goal "|- MClkReply send rcv cst p --> \ -\ _(cst!p, rtrner send!p, caller rcv!p)"; -by (auto_tac (mem_css addsimps2 [angle_def,MClkReply_def] - addEs2 [Return_changed])); -qed "MClkReply_change"; - -Goal "!!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))"; -by (action_simp_tac (simpset()) [MClkReply_change RSN (2,enabled_mono)] [] 1); -by (action_simp_tac (simpset() addsimps [MClkReply_def,Return_def,caller_def,rtrner_def]) - [exI] [base_enabled,Pair_inject] 1); -qed "MClkReply_enabled"; - -Goal "|- MClkReply send rcv cst p --> ~MClkRetry send rcv cst p"; -by (auto_tac (mem_css addsimps2 [MClkReply_def,MClkRetry_def])); -qed "MClkReplyNotRetry"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemClerk.thy --- a/src/HOL/TLA/Memory/MemClerk.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/MemClerk.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: MemClerk - Logic Image: TLA - - RPC-Memory example: specification of the memory clerk. -*) +header {* RPC-Memory example: specification of the memory clerk *} theory MemClerk imports Memory RPC MemClerkParameters @@ -67,6 +64,49 @@ MClkISpec :: "mClkSndChType => mClkRcvChType => mClkStType => temporal" "MClkISpec send rcv cst == TEMP (ALL p. MClkIPSpec send rcv cst p)" -ML {* use_legacy_bindings (the_context ()) *} +lemmas MC_action_defs = + MClkInit_def MClkFwd_def MClkRetry_def MClkReply_def MClkNext_def + +lemmas MC_temp_defs = MClkIPSpec_def MClkISpec_def + +(* The Clerk engages in an action for process p only if there is an outstanding, + unanswered call for that process. +*) +lemma MClkidle: "|- ~$Calling send p & $(cst!p) = #clkA --> ~MClkNext send rcv cst p" + by (auto simp: Return_def MC_action_defs) + +lemma MClkbusy: "|- $Calling rcv p --> ~MClkNext send rcv cst p" + by (auto simp: Call_def MC_action_defs) + + +(* Enabledness of actions *) + +lemma MClkFwd_enabled: "!!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)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "MClkFwd_def", + thm "Call_def", thm "caller_def", thm "rtrner_def"]) [exI] + [thm "base_enabled", Pair_inject] 1 *}) + +lemma MClkFwd_ch_enabled: "|- Enabled (MClkFwd send rcv cst p) --> + Enabled (_(cst!p, rtrner send!p, caller rcv!p))" + by (auto elim!: enabled_mono simp: angle_def MClkFwd_def) + +lemma MClkReply_change: "|- MClkReply send rcv cst p --> + _(cst!p, rtrner send!p, caller rcv!p)" + by (auto simp: angle_def MClkReply_def elim: Return_changed [temp_use]) + +lemma MClkReply_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))" + apply (tactic {* action_simp_tac (simpset ()) + [thm "MClkReply_change" RSN (2, thm "enabled_mono") ] [] 1 *}) + apply (tactic {* action_simp_tac (simpset () addsimps + [thm "MClkReply_def", thm "Return_def", thm "caller_def", thm "rtrner_def"]) + [exI] [thm "base_enabled", Pair_inject] 1 *}) + done + +lemma MClkReplyNotRetry: "|- MClkReply send rcv cst p --> ~MClkRetry send rcv cst p" + by (auto simp: MClkReply_def MClkRetry_def) end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemClerkParameters.ML --- a/src/HOL/TLA/Memory/MemClerkParameters.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,12 +0,0 @@ -(* - File: MemClerkParameters.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: Memory clerk parameters (ML file) -*) - -(* -val CP_simps = RP_simps @ mClkState.simps; -*) diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemClerkParameters.thy --- a/src/HOL/TLA/Memory/MemClerkParameters.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/MemClerkParameters.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: MemClerkParameters - Logic Image: TLA - - RPC-Memory example: Parameters of the memory clerk. -*) +header {* RPC-Memory example: Parameters of the memory clerk *} theory MemClerkParameters imports RPCParameters @@ -30,6 +27,4 @@ MClkReplyVal :: "Vals => Vals" "MClkReplyVal v == if v = RPCFailure then MemFailure else v" -ML {* use_legacy_bindings (the_context ()) *} - end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/Memory.ML --- a/src/HOL/TLA/Memory/Memory.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,126 +0,0 @@ -(* - File: Memory.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: Memory specification (theorems and proofs) -*) - -val RM_action_defs = - [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]; - -val mem_css = (claset(), simpset()); - -(* -------------------- Proofs ---------------------------------------------- *) - -(* The reliable memory is an implementation of the unreliable one *) -Goal "|- IRSpec ch mm rs --> IUSpec ch mm rs"; -by (force_tac (temp_css addsimps2 ([UNext_def,UPSpec_def,IUSpec_def] @ RM_temp_defs) - addSEs2 [STL4E,squareE]) 1); -qed "ReliableImplementsUnReliable"; - -(* The memory spec implies the memory invariant *) -Goal "|- MSpec ch mm rs l --> [](MemInv mm l)"; -by (auto_inv_tac (simpset() addsimps RM_temp_defs @ RM_action_defs) 1); -qed "MemoryInvariant"; - -(* The invariant is trivial for non-locations *) -Goal "|- #l ~: #MemLoc --> [](MemInv mm l)"; -by (auto_tac (temp_css addsimps2 [MemInv_def] addSIs2 [necT])); -qed "NonMemLocInvariant"; - -Goal "|- (ALL l. #l : #MemLoc --> MSpec ch mm rs l) --> (ALL l. [](MemInv mm l))"; -by (step_tac temp_cs 1); -by (case_tac "l : MemLoc" 1); -by (auto_tac (temp_css addSEs2 [MemoryInvariant,NonMemLocInvariant])); -qed "MemoryInvariantAll"; - -(* The memory engages in an action for process p only if there is an - unanswered call from p. - We need this only for the reliable memory. -*) - -Goal "|- ~$(Calling ch p) --> ~ RNext ch mm rs p"; -by (auto_tac (mem_css addsimps2 (Return_def::RM_action_defs))); -qed "Memoryidle"; - -(* Enabledness conditions *) - -Goal "|- MemReturn ch rs p --> _(rtrner ch ! p, rs!p)"; -by (force_tac (mem_css addsimps2 [MemReturn_def,angle_def]) 1); -qed "MemReturn_change"; - -Goal "!!p. basevars (rtrner ch ! p, rs!p) ==> \ -\ |- Calling ch p & (rs!p ~= #NotAResult) \ -\ --> Enabled (_(rtrner ch ! p, rs!p))"; -by (action_simp_tac (simpset()) [MemReturn_change RSN (2,enabled_mono)] [] 1); -by (action_simp_tac (simpset() addsimps [MemReturn_def,Return_def,rtrner_def]) - [exI] [base_enabled,Pair_inject] 1); -qed "MemReturn_enabled"; - -Goal "!!p. basevars (rtrner ch ! p, rs!p) ==> \ -\ |- Calling ch p & (arg = #(read l)) --> Enabled (ReadInner ch mm rs p l)"; -by (case_tac "l : MemLoc" 1); -by (ALLGOALS - (force_tac (mem_css addsimps2 [ReadInner_def,GoodRead_def, - BadRead_def,RdRequest_def] - addSIs2 [exI] addSEs2 [base_enabled]))); -qed "ReadInner_enabled"; - -Goal "!!p. basevars (mm!l, rtrner ch ! p, rs!p) ==> \ -\ |- Calling ch p & (arg = #(write l v)) \ -\ --> Enabled (WriteInner ch mm rs p l v)"; -by (case_tac "l:MemLoc & v:MemVal" 1); -by (ALLGOALS - (force_tac (mem_css addsimps2 [WriteInner_def,GoodWrite_def, - BadWrite_def,WrRequest_def] - addSIs2 [exI] addSEs2 [base_enabled]))); -qed "WriteInner_enabled"; - -Goal "|- Read ch mm rs p & (!l. $(MemInv mm l)) --> (rs!p)` ~= #NotAResult"; -by (force_tac (mem_css addsimps2 - [Read_def,ReadInner_def,GoodRead_def,BadRead_def,MemInv_def]) 1); -qed "ReadResult"; - -Goal "|- Write ch mm rs p l --> (rs!p)` ~= #NotAResult"; -by (auto_tac (mem_css addsimps2 ([Write_def,WriteInner_def,GoodWrite_def,BadWrite_def]))); -qed "WriteResult"; - -Goal "|- (ALL l. $MemInv mm l) & MemReturn ch rs p \ -\ --> ~ Read ch mm rs p & (ALL l. ~ Write ch mm rs p l)"; -by (auto_tac (mem_css addsimps2 [MemReturn_def] addSDs2 [WriteResult, ReadResult])); -qed "ReturnNotReadWrite"; - -Goal "|- (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))"; -by (force_tac (mem_css addsimps2 [RNext_def,angle_def] - addSEs2 [enabled_mono2] - addDs2 [ReadResult, WriteResult]) 1); -qed "RWRNext_enabled"; - - -(* Combine previous lemmas: the memory can make a visible step if there is an - outstanding call for which no result has been produced. -*) -Goal "!!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))"; -by (auto_tac (mem_css addsimps2 [enabled_disj] addSIs2 [RWRNext_enabled])); -by (case_tac "arg(ch w p)" 1); - by (action_simp_tac (simpset()addsimps[Read_def,enabled_ex]) - [ReadInner_enabled,exI] [] 1); - by (force_tac (mem_css addDs2 [base_pair]) 1); -by (etac contrapos_np 1); -by (action_simp_tac (simpset() addsimps [Write_def,enabled_ex]) - [WriteInner_enabled,exI] [] 1); -qed "RNext_enabled"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/Memory.thy --- a/src/HOL/TLA/Memory/Memory.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/Memory.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: Memory - Logic Image: TLA - - RPC-Memory example: Memory specification -*) +header {* RPC-Memory example: Memory specification *} theory Memory imports MemoryParameters ProcedureInterface @@ -136,6 +133,108 @@ MemInv_def: "MemInv mm l == PRED #l : #MemLoc --> mm!l : #MemVal" -ML {* use_legacy_bindings (the_context ()) *} +lemmas RM_action_defs = + 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 + +lemmas UM_action_defs = RM_action_defs MemFail_def UNext_def + +lemmas RM_temp_defs = RPSpec_def MSpec_def IRSpec_def +lemmas UM_temp_defs = UPSpec_def MSpec_def IUSpec_def + + +(* The reliable memory is an implementation of the unreliable one *) +lemma ReliableImplementsUnReliable: "|- IRSpec ch mm rs --> IUSpec ch mm rs" + by (force simp: UNext_def UPSpec_def IUSpec_def RM_temp_defs elim!: STL4E [temp_use] squareE) + +(* The memory spec implies the memory invariant *) +lemma MemoryInvariant: "|- MSpec ch mm rs l --> [](MemInv mm l)" + by (tactic {* auto_inv_tac + (simpset () addsimps (thms "RM_temp_defs" @ thms "RM_action_defs")) 1 *}) + +(* The invariant is trivial for non-locations *) +lemma NonMemLocInvariant: "|- #l ~: #MemLoc --> [](MemInv mm l)" + by (auto simp: MemInv_def intro!: necT [temp_use]) + +lemma MemoryInvariantAll: + "|- (ALL l. #l : #MemLoc --> MSpec ch mm rs l) --> (ALL l. [](MemInv mm l))" + apply clarify + apply (case_tac "l : MemLoc") + apply (auto elim!: MemoryInvariant [temp_use] NonMemLocInvariant [temp_use]) + done + +(* The memory engages in an action for process p only if there is an + unanswered call from p. + We need this only for the reliable memory. +*) + +lemma Memoryidle: "|- ~$(Calling ch p) --> ~ RNext ch mm rs p" + by (auto simp: Return_def RM_action_defs) + +(* Enabledness conditions *) + +lemma MemReturn_change: "|- MemReturn ch rs p --> _(rtrner ch ! p, rs!p)" + by (force simp: MemReturn_def angle_def) + +lemma MemReturn_enabled: "!!p. basevars (rtrner ch ! p, rs!p) ==> + |- Calling ch p & (rs!p ~= #NotAResult) + --> Enabled (_(rtrner ch ! p, rs!p))" + apply (tactic + {* action_simp_tac (simpset ()) [thm "MemReturn_change" RSN (2, thm "enabled_mono") ] [] 1 *}) + apply (tactic + {* action_simp_tac (simpset () addsimps [thm "MemReturn_def", thm "Return_def", + thm "rtrner_def"]) [exI] [thm "base_enabled", thm "Pair_inject"] 1 *}) + done + +lemma ReadInner_enabled: "!!p. basevars (rtrner ch ! p, rs!p) ==> + |- Calling ch p & (arg = #(read l)) --> Enabled (ReadInner ch mm rs p l)" + apply (case_tac "l : MemLoc") + apply (force simp: ReadInner_def GoodRead_def BadRead_def RdRequest_def + intro!: exI elim!: base_enabled [temp_use])+ + done + +lemma WriteInner_enabled: "!!p. basevars (mm!l, rtrner ch ! p, rs!p) ==> + |- Calling ch p & (arg = #(write l v)) + --> Enabled (WriteInner ch mm rs p l v)" + apply (case_tac "l:MemLoc & v:MemVal") + apply (force simp: WriteInner_def GoodWrite_def BadWrite_def WrRequest_def + intro!: exI elim!: base_enabled [temp_use])+ + done + +lemma ReadResult: "|- Read ch mm rs p & (!l. $(MemInv mm l)) --> (rs!p)` ~= #NotAResult" + by (force simp: Read_def ReadInner_def GoodRead_def BadRead_def MemInv_def) + +lemma WriteResult: "|- Write ch mm rs p l --> (rs!p)` ~= #NotAResult" + by (auto simp: Write_def WriteInner_def GoodWrite_def BadWrite_def) + +lemma ReturnNotReadWrite: "|- (ALL l. $MemInv mm l) & MemReturn ch rs p + --> ~ Read ch mm rs p & (ALL l. ~ Write ch mm rs p l)" + by (auto simp: MemReturn_def dest!: WriteResult [temp_use] ReadResult [temp_use]) + +lemma RWRNext_enabled: "|- (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))" + by (force simp: RNext_def angle_def elim!: enabled_mono2 + dest: ReadResult [temp_use] WriteResult [temp_use]) + + +(* Combine previous lemmas: the memory can make a visible step if there is an + outstanding call for which no result has been produced. +*) +lemma RNext_enabled: "!!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))" + apply (auto simp: enabled_disj [try_rewrite] intro!: RWRNext_enabled [temp_use]) + apply (case_tac "arg (ch w p)") + apply (tactic {* action_simp_tac (simpset () addsimps [thm "Read_def", + temp_rewrite (thm "enabled_ex")]) [thm "ReadInner_enabled", exI] [] 1 *}) + apply (force dest: base_pair [temp_use]) + apply (erule contrapos_np) + apply (tactic {* action_simp_tac (simpset () addsimps [thm "Write_def", + temp_rewrite (thm "enabled_ex")]) + [thm "WriteInner_enabled", exI] [] 1 *}) + done end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemoryImplementation.ML --- a/src/HOL/TLA/Memory/MemoryImplementation.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,904 +0,0 @@ -(* - File: MemoryImplementation.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: Memory implementation (ML file) - - The main theorem is theorem "Implementation" at the end of this file, - which shows that the composition of a reliable memory, an RPC component, and - a memory clerk implements an unreliable memory. The files "MIsafe.ML" and - "MIlive.ML" contain lower-level lemmas for the safety and liveness parts. - - Steps are (roughly) numbered as in the hand proof. -*) - -(* --------------------------- automatic prover --------------------------- *) - -Delcongs [if_weak_cong]; - -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 it can be a lot faster than MI_css) -*) -val MI_fast_css = - let - val (cs,ss) = MI_css - in - (cs addSEs [squareE], ss addSSolver (mk_solver "" (fn thms => assume_tac ORELSE' (etac notE)))) -end; - -val temp_elim = make_elim o temp_use; - -(****************************** The history variable ******************************) -section "History variable"; - -Goal "|- Init(ALL p. ImpInit p) & [](ALL p. ImpNext p) \ -\ --> (EEX rmhist. Init(ALL p. HInit rmhist p) \ -\ & [](ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)))"; -by (Clarsimp_tac 1); -by (rtac historyI 1); -by (TRYALL atac); -by (rtac MI_base 1); -by (action_simp_tac (simpset() addsimps [HInit_def]) [] [] 1); -by (etac fun_cong 1); -by (action_simp_tac (simpset() addsimps [HNext_def]) [busy_squareI] [] 1); -by (etac fun_cong 1); -qed "HistoryLemma"; - -Goal "|- Implementation --> (EEX rmhist. Hist rmhist)"; -by (Clarsimp_tac 1); -by (rtac ((temp_use HistoryLemma) RS eex_mono) 1); -by (force_tac (MI_css - addsimps2 [Hist_def,HistP_def,Init_def,all_box,split_box_conj]) 3); -by (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])); -qed "History"; - -(******************************** The safety part *********************************) - -section "The safety part"; - -(* ------------------------- Include lower-level lemmas ------------------------- *) -use "MIsafe.ML"; - -section "Correctness of predicate-action diagram"; - - -(* ========== Step 1.1 ================================================= *) -(* The implementation's initial condition implies the state predicate S1 *) - -Goal "|- ImpInit p & HInit rmhist p --> S1 rmhist p"; -by (auto_tac (MI_fast_css - addsimps2 [MVNROKBA_def,MClkInit_def,RPCInit_def,PInit_def, - HInit_def,ImpInit_def,S_def,S1_def])); -qed "Step1_1"; - -(* ========== Step 1.2 ================================================== *) -(* Figure 16 is a predicate-action diagram for the implementation. *) - -Goal "|- [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)"; -by (action_simp_tac (simpset() addsimps [ImpNext_def]) [] - (map temp_elim [S1ClerkUnch,S1RPCUnch,S1MemUnch,S1Hist]) 1); -by (auto_tac (MI_fast_css addSIs2 [S1Env])); -qed "Step1_2_1"; - -Goal "|- [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)"; -by (action_simp_tac (simpset() addsimps [ImpNext_def]) [] - (map temp_elim [S2EnvUnch,S2RPCUnch,S2MemUnch,S2Hist]) 1); -by (auto_tac (MI_fast_css addSIs2 [S2Clerk,S2Forward])); -qed "Step1_2_2"; - -Goal "|- [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))"; -by (action_simp_tac (simpset() addsimps [ImpNext_def]) [] - (map temp_elim [S3EnvUnch,S3ClerkUnch,S3MemUnch]) 1); -by (action_simp_tac (simpset()) [] - (squareE::map temp_elim [S3RPC,S3Forward,S3Fail]) 1); -by (auto_tac (MI_css addDs2 [S3Hist])); -qed "Step1_2_3"; - -Goal "|- [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))"; -by (action_simp_tac (simpset() addsimps [ImpNext_def]) [] - (map temp_elim [S4EnvUnch,S4ClerkUnch,S4RPCUnch]) 1); -by (action_simp_tac (simpset() addsimps [RNext_def]) [] - (squareE::map temp_elim [S4Read,S4Write,S4Return]) 1); -by (auto_tac (MI_css addDs2 [S4Hist])); -qed "Step1_2_4"; - -Goal "|- [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))"; -by (action_simp_tac (simpset() addsimps [ImpNext_def]) [] - (map temp_elim [S5EnvUnch,S5ClerkUnch,S5MemUnch,S5Hist]) 1); -by (action_simp_tac (simpset()) [] [squareE, temp_elim S5RPC] 1); -by (auto_tac (MI_fast_css addSDs2 [S5Reply,S5Fail])); -qed "Step1_2_5"; - -Goal "|- [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))"; -by (action_simp_tac (simpset() addsimps [ImpNext_def]) [] - (map temp_elim [S6EnvUnch,S6RPCUnch,S6MemUnch]) 1); -by (action_simp_tac (simpset()) [] - (squareE::map temp_elim [S6Clerk,S6Retry,S6Reply]) 1); -by (auto_tac (MI_css addDs2 [S6Hist])); -qed "Step1_2_6"; - -(* -------------------------------------------------------------------------- - Step 1.3: S1 implies the barred initial condition. -*) - -section "Initialization (Step 1.3)"; - -Goal "|- S1 rmhist p --> PInit (resbar rmhist) p"; -by (action_simp_tac (simpset() addsimps [resbar_def,PInit_def,S_def,S1_def]) - [] [] 1); -qed "Step1_3"; - -(* ---------------------------------------------------------------------- - Step 1.4: Implementation's next-state relation simulates specification's - next-state relation (with appropriate substitutions) -*) - -section "Step simulation (Step 1.4)"; - -Goal "|- ENext p & $S1 rmhist p & (S2 rmhist p)$ & unchanged (c p, r p, m p) \ -\ --> unchanged (rtrner memCh!p, resbar rmhist!p)"; -by (auto_tac (MI_fast_css addsimps2 [c_def,r_def,m_def,resbar_def])); -qed "Step1_4_1"; - -Goal "|- 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)"; -by (action_simp_tac - (simpset() addsimps [MClkFwd_def, e_def, r_def, m_def, resbar_def, - S_def, S2_def, S3_def]) [] [] 1); -qed "Step1_4_2"; - -Goal "|- 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)"; -by (Clarsimp_tac 1); -by (REPEAT (dresolve_tac (map temp_use [S3_excl,S4_excl]) 1)); -by (action_simp_tac - (simpset() addsimps [e_def,c_def,m_def,resbar_def,S_def, S3_def]) [] [] 1); -qed "Step1_4_3a"; - -Goal "|- RPCFail crCh rmCh rst p & $S3 rmhist p & (S6 rmhist p)$\ -\ & unchanged (e p, c p, m p) \ -\ --> MemFail memCh (resbar rmhist) p"; -by (Clarsimp_tac 1); -by (dtac (temp_use S6_excl) 1); -by (auto_tac (MI_css addsimps2 [RPCFail_def,MemFail_def,e_def,c_def,m_def, - resbar_def])); -by (force_tac (MI_css addsimps2 [S3_def,S_def]) 1); -by (auto_tac (MI_css addsimps2 [Return_def])); -qed "Step1_4_3b"; - -Goal "|- $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"; -by (Clarsimp_tac 1); -by (REPEAT (dtac (temp_use S4_excl) 1)); -by (action_simp_tac - (simpset() addsimps [ReadInner_def,GoodRead_def,BadRead_def,e_def,c_def,m_def]) - [] [] 1); -by (auto_tac (MI_css addsimps2 [resbar_def])); -by (ALLGOALS (action_simp_tac - (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def, - S_def,S4_def,RdRequest_def,MemInv_def]) - [] [impE,MemValNotAResultE])); -qed "Step1_4_4a1"; - -Goal "|- 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"; -by (force_tac (MI_css addsimps2 [Read_def] addSEs2 [Step1_4_4a1]) 1); -qed "Step1_4_4a"; - -Goal "|- $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"; -by (Clarsimp_tac 1); -by (REPEAT (dtac (temp_use S4_excl) 1)); -by (action_simp_tac - (simpset() addsimps [WriteInner_def, GoodWrite_def, BadWrite_def, - e_def, c_def, m_def]) - [] [] 1); -by (auto_tac (MI_css addsimps2 [resbar_def])); -by (ALLGOALS (action_simp_tac - (simpset() addsimps [RPCRelayArg_def,MClkRelayArg_def, - S_def,S4_def,WrRequest_def]) - [] [])); -qed "Step1_4_4b1"; - -Goal "|- 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"; -by (force_tac (MI_css addsimps2 [Write_def] addSEs2 [Step1_4_4b1]) 1); -qed "Step1_4_4b"; - -Goal "|- MemReturn rmCh ires p & $S4 rmhist p & (S5 rmhist p)$\ -\ & unchanged (e p, c p, r p) \ -\ --> unchanged (rtrner memCh!p, resbar rmhist!p)"; -by (action_simp_tac - (simpset() addsimps [e_def,c_def,r_def,resbar_def]) [] [] 1); -by (REPEAT (dresolve_tac [temp_use S4_excl, temp_use S5_excl] 1)); -by (auto_tac (MI_fast_css addsimps2 [MemReturn_def,Return_def])); -qed "Step1_4_4c"; - -Goal "|- 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)"; -by (Clarsimp_tac 1); -by (REPEAT (dresolve_tac [temp_use S5_excl, temp_use S6_excl] 1)); -by (auto_tac (MI_css addsimps2 [e_def,c_def,m_def, resbar_def])); -by (auto_tac (MI_css addsimps2 [RPCReply_def,Return_def,S5_def,S_def] - addSDs2 [MVOKBAnotRF])); -qed "Step1_4_5a"; - -Goal "|- RPCFail crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$\ -\ & unchanged (e p, c p, m p) \ -\ --> MemFail memCh (resbar rmhist) p"; -by (Clarsimp_tac 1); -by (dtac (temp_use S6_excl) 1); -by (auto_tac (MI_css addsimps2 [e_def, c_def, m_def, RPCFail_def, Return_def, - MemFail_def, resbar_def])); -by (auto_tac (MI_css addsimps2 [S5_def,S_def])); -qed "Step1_4_5b"; - -Goal "|- MClkReply memCh crCh cst p & $S6 rmhist p & (S1 rmhist p)$\ -\ & unchanged (e p, r p, m p) \ -\ --> MemReturn memCh (resbar rmhist) p"; -by (Clarsimp_tac 1); -by (dtac (temp_use S6_excl) 1); -by (action_simp_tac - (simpset() addsimps [e_def,r_def,m_def,MClkReply_def,MemReturn_def, - Return_def,resbar_def]) [] [] 1); -by (ALLGOALS Asm_full_simp_tac); (* simplify if-then-else *) -by (ALLGOALS (action_simp_tac - (simpset() addsimps [MClkReplyVal_def,S6_def,S_def]) - [] [MVOKBARFnotNR])); -qed "Step1_4_6a"; - -Goal "|- 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"; -by (Clarsimp_tac 1); -by (dtac (temp_use S3_excl) 1); -by (action_simp_tac - (simpset() addsimps [e_def, r_def, m_def, MClkRetry_def, MemFail_def, resbar_def]) - [] [] 1); -by (auto_tac (MI_css addsimps2 [S6_def,S_def])); -qed "Step1_4_6b"; - -Goal "|- unchanged (e p, c p, r p, m p, rmhist!p) \ -\ --> unchanged (S rmhist ec cc rc cs rs hs1 hs2 p)"; -by (auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,caller_def,rtrner_def, - S_def,Calling_def])); -qed "S_lemma"; - -Goal "|- 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)"; -by (Clarsimp_tac 1); -by (rtac conjI 1); -by (force_tac (MI_css addsimps2 [c_def]) 1); -by (force_tac (MI_css addsimps2 [S1_def,S2_def,S3_def,S4_def,S5_def,S6_def] - addSIs2 [S_lemma]) 1); -qed "Step1_4_7H"; - -Goal "|- 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)"; -by (rtac actionI 1); -by (rewrite_goals_tac action_rews); -by (rtac impI 1); -by (forward_tac [temp_use Step1_4_7H] 1); -by (auto_tac (MI_css addsimps2 [e_def,c_def,r_def,m_def,rtrner_def,resbar_def])); -qed "Step1_4_7"; - -(* Frequently needed abbreviation: distinguish between idling and non-idling - steps of the implementation, and try to solve the idling case by simplification -*) -fun split_idle_tac simps 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, - forward_tac [temp_use Step1_4_7] i, - asm_full_simp_tac (simpset() addsimps simps) i - ]; - -(* ---------------------------------------------------------------------- - Combine steps 1.2 and 1.4 to prove that the implementation satisfies - the specification's next-state relation. -*) - -(* Steps that leave all variables unchanged are safe, so I may assume - that some variable changes in the proof that a step is safe. *) -Goal "|- (~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)"; -by (split_idle_tac [square_def] 1); -by (Force_tac 1); -qed "unchanged_safe"; -(* turn into (unsafe, looping!) introduction rule *) -bind_thm("unchanged_safeI", impI RS (action_use unchanged_safe)); - -Goal "|- $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)"; -by (Clarsimp_tac 1); -by (rtac unchanged_safeI 1); -by (rtac idle_squareI 1); -by (auto_tac (MI_css addSDs2 [Step1_2_1,Step1_4_1])); -qed "S1safe"; - -Goal "|- $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)"; -by (Clarsimp_tac 1); -by (rtac unchanged_safeI 1); -by (rtac idle_squareI 1); -by (auto_tac (MI_css addSDs2 [Step1_2_2,Step1_4_2])); -qed "S2safe"; - -Goal "|- $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)"; -by (Clarsimp_tac 1); -by (rtac unchanged_safeI 1); -by (auto_tac (MI_css addSDs2 [Step1_2_3])); -by (auto_tac (MI_css addsimps2 [square_def,UNext_def] addSDs2 [Step1_4_3a,Step1_4_3b])); -qed "S3safe"; - -Goal "|- $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)"; -by (Clarsimp_tac 1); -by (rtac unchanged_safeI 1); -by (auto_tac (MI_css addSDs2 [Step1_2_4])); -by (auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def] - addSDs2 [Step1_4_4a,Step1_4_4b,Step1_4_4c])); -qed "S4safe"; - -Goal "|- $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)"; -by (Clarsimp_tac 1); -by (rtac unchanged_safeI 1); -by (auto_tac (MI_css addSDs2 [Step1_2_5])); -by (auto_tac (MI_css addsimps2 [square_def,UNext_def] - addSDs2 [Step1_4_5a,Step1_4_5b])); -qed "S5safe"; - -Goal "|- $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)"; -by (Clarsimp_tac 1); -by (rtac unchanged_safeI 1); -by (auto_tac (MI_css addSDs2 [Step1_2_6])); -by (auto_tac (MI_css addsimps2 [square_def,UNext_def,RNext_def] - addSDs2 [Step1_4_6a,Step1_4_6b])); -qed "S6safe"; - -(* ---------------------------------------------------------------------- - Step 1.5: Temporal refinement proof, based on previous steps. -*) - -section "The liveness part"; - -(* Liveness assertions for the different implementation states, based on the - fairness conditions. Prove subgoals of WF1 / SF1 rules as separate lemmas - for readability. Reuse action proofs from safety part. -*) - -(* ------------------------------ State S1 ------------------------------ *) - -Goal "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ -\ --> (S1 rmhist p)$ | (S2 rmhist p)$"; -by (split_idle_tac [] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_1])); -qed "S1_successors"; - -(* Show that the implementation can satisfy the high-level fairness requirements - by entering the state S1 infinitely often. -*) - -Goal "|- S1 rmhist p --> \ -\ ~Enabled (_(rtrner memCh!p, resbar rmhist!p))"; -by (action_simp_tac (simpset() addsimps [angle_def,S_def,S1_def]) - [notI] [enabledE,temp_elim Memoryidle] 1); -by (Force_tac 1); -qed "S1_RNextdisabled"; - -Goal "|- S1 rmhist p --> \ -\ ~Enabled (_(rtrner memCh!p, resbar rmhist!p))"; -by (action_simp_tac - (simpset() addsimps [angle_def,MemReturn_def,Return_def,S_def,S1_def]) - [notI] [enabledE] 1); -qed "S1_Returndisabled"; - -Goal "|- []<>S1 rmhist p \ -\ --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"; -by (auto_tac (MI_css addsimps2 [WF_alt] - addSIs2 [S1_RNextdisabled] addSEs2 [STL4E,DmdImplE])); -qed "RNext_fair"; - -Goal "|- []<>S1 rmhist p \ -\ --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"; -by (auto_tac (MI_css addsimps2 [WF_alt] - addSIs2 [S1_Returndisabled] addSEs2 [STL4E,DmdImplE])); -qed "Return_fair"; - -(* ------------------------------ State S2 ------------------------------ *) - -Goal "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ -\ --> (S2 rmhist p)$ | (S3 rmhist p)$"; -by (split_idle_tac [] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_2])); -qed "S2_successors"; - -Goal "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\ -\ & _(c p) \ -\ --> (S3 rmhist p)$"; -by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_2])); -qed "S2MClkFwd_successors"; - -Goal "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\ -\ --> $Enabled (_(c p))"; -by (auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkFwd_ch_enabled,MClkFwd_enabled])); -by (cut_facts_tac [MI_base] 1); -by (blast_tac (claset() addDs [base_pair]) 1); -by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S2_def]))); -qed "S2MClkFwd_enabled"; - -Goal "|- [](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)"; -by (REPEAT (resolve_tac [WF1,S2_successors, - S2MClkFwd_successors,S2MClkFwd_enabled] 1)); -qed "S2_live"; - -(* ------------------------------ State S3 ------------------------------ *) - -Goal "|- $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)$"; -by (split_idle_tac [] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_3])); -qed "S3_successors"; - -Goal "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p))\ -\ & _(r p) \ -\ --> (S4 rmhist p | S6 rmhist p)$"; -by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_3])); -qed "S3RPC_successors"; - -Goal "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\ -\ --> $Enabled (_(r p))"; -by (auto_tac (MI_css addsimps2 [r_def] - addSIs2 [RPCFail_Next_enabled,RPCFail_enabled])); -by (cut_facts_tac [MI_base] 1); -by (blast_tac (claset() addDs [base_pair]) 1); -by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S3_def]))); -qed "S3RPC_enabled"; - -Goal "|- [](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)"; -by (REPEAT (resolve_tac [WF1,S3_successors,S3RPC_successors,S3RPC_enabled] 1)); -qed "S3_live"; - -(* ------------- State S4 -------------------------------------------------- *) - -Goal"|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ -\ & (ALL l. $MemInv mm l) \ -\ --> (S4 rmhist p)$ | (S5 rmhist p)$"; -by (split_idle_tac [] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_4])); -qed "S4_successors"; - -(* --------- State S4a: S4 /\ (ires p = NotAResult) ------------------------ *) - -Goal "|- $(S4 rmhist p & ires!p = #NotAResult) \ -\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l)\ -\ --> (S4 rmhist p & ires!p = #NotAResult)$ \ -\ | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$"; -by (split_idle_tac [m_def] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_4])); -qed "S4a_successors"; - -Goal "|- ($(S4 rmhist p & ires!p = #NotAResult) \ -\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l))\ -\ & _(m p) \ -\ --> ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$"; -by (auto_tac (MI_css addsimps2 [angle_def] - addSDs2 [Step1_2_4, ReadResult, WriteResult])); -qed "S4aRNext_successors"; - -Goal "|- $(S4 rmhist p & ires!p = #NotAResult) \ -\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l)\ -\ --> $Enabled (_(m p))"; -by (auto_tac (MI_css addsimps2 [m_def] addSIs2 [RNext_enabled])); -by (cut_facts_tac [MI_base] 1); -by (blast_tac (claset() addDs [base_pair]) 1); -by (asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1); -qed "S4aRNext_enabled"; - -Goal "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\ -\ & (ALL 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)"; -by (REPEAT (resolve_tac [WF1, S4a_successors, S4aRNext_successors, S4aRNext_enabled] 1)); -qed "S4a_live"; - -(* ---------- State S4b: S4 /\ (ires p # NotAResult) --------------------------- *) - -Goal "|- $(S4 rmhist p & ires!p ~= #NotAResult) \ -\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l)\ -\ --> (S4 rmhist p & ires!p ~= #NotAResult)$ | (S5 rmhist p)$"; -by (split_idle_tac [m_def] 1); -by (auto_tac (MI_css addSDs2 [WriteResult,Step1_2_4,ReadResult])); -qed "S4b_successors"; - -Goal "|- ($(S4 rmhist p & ires!p ~= #NotAResult) \ -\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ -\ & (ALL l. $MemInv mm l)) & _(m p) \ -\ --> (S5 rmhist p)$"; -by (force_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_4] - addDs2 [ReturnNotReadWrite]) 1); -qed "S4bReturn_successors"; - -Goal "|- $(S4 rmhist p & ires!p ~= #NotAResult) \ -\ & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)\ -\ & (ALL l. $MemInv mm l) \ -\ --> $Enabled (_(m p))"; -by (auto_tac (MI_css addsimps2 [m_def] addSIs2 [MemReturn_enabled])); -by (cut_facts_tac [MI_base] 1); -by (blast_tac (claset() addDs [base_pair]) 1); -by (asm_full_simp_tac (simpset() addsimps [S_def,S4_def]) 1); -qed "S4bReturn_enabled"; - -Goal "|- [](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)"; -by (REPEAT (resolve_tac [WF1, S4b_successors,S4bReturn_successors, S4bReturn_enabled] 1)); -qed "S4b_live"; - -(* ------------------------------ State S5 ------------------------------ *) - -Goal "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ -\ --> (S5 rmhist p)$ | (S6 rmhist p)$"; -by (split_idle_tac [] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_5])); -qed "S5_successors"; - -Goal "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ -\ & _(r p) \ -\ --> (S6 rmhist p)$"; -by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_5])); -qed "S5RPC_successors"; - -Goal "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) \ -\ --> $Enabled (_(r p))"; -by (auto_tac (MI_css addsimps2 [r_def] - addSIs2 [RPCFail_Next_enabled, RPCFail_enabled])); -by (cut_facts_tac [MI_base] 1); -by (blast_tac (claset() addDs [base_pair]) 1); -by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [S_def,S5_def]))); -qed "S5RPC_enabled"; - -Goal "|- [](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)"; -by (REPEAT (resolve_tac [WF1,S5_successors,S5RPC_successors,S5RPC_enabled] 1)); -qed "S5_live"; - -(* ------------------------------ State S6 ------------------------------ *) - -Goal "|- $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)$"; -by (split_idle_tac [] 1); -by (auto_tac (MI_css addSDs2 [Step1_2_6])); -qed "S6_successors"; - -Goal "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) \ -\ & _(c p) \ -\ --> (S1 rmhist p)$"; -by (auto_tac (MI_css addsimps2 [angle_def] addSDs2 [Step1_2_6, MClkReplyNotRetry])); -qed "S6MClkReply_successors"; - -Goal "|- $ImpInv rmhist p & _(c p) --> $S6 rmhist p"; -by (action_simp_tac - (simpset() addsimps [angle_def,MClkReply_def,Return_def, - ImpInv_def,S_def,S1_def,S2_def,S3_def,S4_def,S5_def]) - [] [] 1); -qed "MClkReplyS6"; - -Goal "|- S6 rmhist p --> Enabled (_(c p))"; -by (auto_tac (MI_css addsimps2 [c_def] addSIs2 [MClkReply_enabled])); -by (cut_facts_tac [MI_base] 1); -by (blast_tac (claset() addDs [base_pair]) 1); -by (ALLGOALS (action_simp_tac (simpset() addsimps [S_def,S6_def]) [] [])); -qed "S6MClkReply_enabled"; - -Goal "|- [](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)"; -by (Clarsimp_tac 1); -by (subgoal_tac "sigma |= []<>(_(c p))" 1); -by (etac InfiniteEnsures 1); -by (atac 1); -by (action_simp_tac (simpset()) [] - (map temp_elim [MClkReplyS6,S6MClkReply_successors]) 1); -by (auto_tac (MI_css addsimps2 [SF_def])); -by (etac contrapos_np 1); -by (auto_tac (MI_css addSIs2 [S6MClkReply_enabled] addSEs2 [STL4E, DmdImplE])); -qed "S6_live"; - -(* --------------- aggregate leadsto properties----------------------------- *) - -Goal "sigma |= S5 rmhist p ~> S6 rmhist p \ -\ ==> sigma |= (S5 rmhist p | S6 rmhist p) ~> S6 rmhist p"; -by (auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro, LatticeReflexivity])); -qed "S5S6LeadstoS6"; - -Goal "[| 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"; -by (auto_tac (MI_css addSIs2 [LatticeDisjunctionIntro,S5S6LeadstoS6] - addIs2 [LatticeTransitivity])); -qed "S4bS5S6LeadstoS6"; - -Goal "[| 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"; -by (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); - by (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); - by (force_tac (MI_css addsimps2 Init_defs addSIs2 [ImplLeadsto_gen, necT]) 1); -by (rtac (temp_use LatticeDisjunctionIntro) 1); -by (etac (temp_use LatticeTransitivity) 1); -by (etac (temp_use LatticeTriangle2) 1); -by (atac 1); -by (auto_tac (MI_css addSIs2 [S4bS5S6LeadstoS6])); -qed "S4S5S6LeadstoS6"; - -Goal "[| 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"; -by (rtac (temp_use LatticeDisjunctionIntro) 1); -by (etac (temp_use LatticeTriangle2) 1); -by (rtac (S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1); -by (auto_tac (MI_css addSIs2 [S4S5S6LeadstoS6,necT] - addIs2 [ImplLeadsto_gen] addsimps2 Init_defs)); -qed "S3S4S5S6LeadstoS6"; - -Goal "[| 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"; -by (rtac (temp_use LatticeDisjunctionIntro) 1); -by (rtac (temp_use LatticeTransitivity) 1); -by (atac 2); -by (rtac (S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1); -by (auto_tac (MI_css addSIs2 [S3S4S5S6LeadstoS6,necT] - addIs2 [ImplLeadsto_gen] addsimps2 Init_defs)); -qed "S2S3S4S5S6LeadstoS6"; - -Goal "[| 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"; -by (rtac (S2S3S4S5S6LeadstoS6 RS (temp_use LatticeTransitivity)) 1); -by (TRYALL atac); -by (etac (temp_use INV_leadsto) 1); -by (rtac (temp_use ImplLeadsto_gen) 1); -by (rtac (temp_use necT) 1); -by (auto_tac (MI_css addsimps2 ImpInv_def::Init_defs addSIs2 [necT])); -qed "NotS1LeadstoS6"; - -Goal "[| sigma |= ~S1 rmhist p ~> S6 rmhist p; \ -\ sigma |= []<>S6 rmhist p --> []<>S1 rmhist p |] \ -\ ==> sigma |= []<>S1 rmhist p"; -by (rtac classical 1); -by (asm_lr_simp_tac (simpset() addsimps [temp_use NotBox, NotDmd]) 1); -by (auto_tac (MI_css addSEs2 [mp,leadsto_infinite] addSDs2 [DBImplBD])); -qed "S1Infinite"; - -section "Refinement proof (step 1.5)"; - -(* Prove invariants of the implementation: - a. memory invariant - b. "implementation invariant": always in states S1,...,S6 -*) -Goal "|- IPImp p --> (ALL l. []$MemInv mm l)"; -by (auto_tac (MI_css addsimps2 [IPImp_def,box_stp_act] - addSIs2 [MemoryInvariantAll])); -qed "Step1_5_1a"; - -Goal "|- Init(ImpInit p & HInit rmhist p) & [](ImpNext p) \ -\ & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) & [](ALL l. $MemInv mm l) \ -\ --> []ImpInv rmhist p"; -by (inv_tac MI_css 1); -by (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])); -qed "Step1_5_1b"; - -(*** Initialization ***) -Goal "|- Init(ImpInit p & HInit rmhist p) --> Init(PInit (resbar rmhist) p)"; -by (auto_tac (MI_css addsimps2 [Init_def] addSIs2 [Step1_1,Step1_3])); -qed "Step1_5_2a"; - -(*** step simulation ***) -Goal "|- [](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)"; -by (auto_tac (MI_css addsimps2 [ImpInv_def] addSEs2 [STL4E] - addSDs2 [S1safe,S2safe,S3safe,S4safe,S5safe,S6safe])); -qed "Step1_5_2b"; - -(*** Liveness ***) -Goal "|- IPImp p & HistP rmhist p \ -\ --> Init(ImpInit p & HInit rmhist p) \ -\ & [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ -\ & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) \ -\ & ImpLive p"; -by (Clarsimp_tac 1); -by (subgoal_tac - "sigma |= Init(ImpInit p & HInit rmhist p) \ -\ & [](ImpNext p) \ -\ & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) \ -\ & [](ALL l. $MemInv mm l)" 1); -by (auto_tac (MI_css addsimps2 [split_box_conj,box_stp_act] addSDs2 [Step1_5_1b])); -by (force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, - ImpLive_def,c_def,r_def,m_def]) 1); -by (force_tac (MI_css addsimps2 [IPImp_def,MClkIPSpec_def,RPCIPSpec_def,RPSpec_def, - HistP_def,Init_def,ImpInit_def]) 1); -by (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); -by (force_tac (MI_css addsimps2 [HistP_def]) 1); -by (force_tac (MI_css addsimps2 [temp_use allT] addSDs2 [Step1_5_1a]) 1); -qed "GoodImpl"; - -(* The implementation is infinitely often in state S1... *) -Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ -\ & [](ALL l. $MemInv mm l) \ -\ & []($ImpInv rmhist p) & ImpLive p \ -\ --> []<>S1 rmhist p"; -by (clarsimp_tac (MI_css addsimps2 [ImpLive_def]) 1); -by (rtac S1Infinite 1); -by (force_tac - (MI_css addsimps2 [split_box_conj,box_stp_act] - addSIs2 [NotS1LeadstoS6,S2_live,S3_live,S4a_live,S4b_live,S5_live]) 1); -by (auto_tac (MI_css addsimps2 [split_box_conj] addSIs2 [S6_live])); -qed "Step1_5_3a"; - -(* ... and therefore satisfies the fairness requirements of the specification *) -Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ -\ & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p \ -\ --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"; -by (auto_tac (MI_css addSIs2 [RNext_fair,Step1_5_3a])); -qed "Step1_5_3b"; - -Goal "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ -\ & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p \ -\ --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)"; -by (auto_tac (MI_css addSIs2 [Return_fair,Step1_5_3a])); -qed "Step1_5_3c"; - -(* QED step of step 1 *) -Goal "|- IPImp p & HistP rmhist p --> UPSpec memCh mm (resbar rmhist) p"; -by (auto_tac (MI_css addsimps2 [UPSpec_def,split_box_conj] - addSDs2 [GoodImpl] - addSIs2 [Step1_5_2a,Step1_5_2b,Step1_5_3b,Step1_5_3c])); -qed "Step1"; - -(* ------------------------------ Step 2 ------------------------------ *) -section "Step 2"; - -Goal "|- 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)"; -by (Clarsimp_tac 1); -by (dtac (action_use WriteS4) 1); -by (atac 1); -by (split_idle_tac [] 1); -by (auto_tac (MI_css addsimps2 [ImpNext_def] - addSDs2 [S4EnvUnch,S4ClerkUnch,S4RPCUnch])); -by (auto_tac (MI_css addsimps2 [square_def] addDs2 [S4Write])); -qed "Step2_2a"; - -Goal "|- (ALL p. ImpNext p) \ -\ & (ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ -\ & (ALL p. $ImpInv rmhist p) \ -\ & [EX q. Write rmCh mm ires q l]_(mm!l) \ -\ --> [EX q. Write memCh mm (resbar rmhist) q l]_(mm!l)"; -by (auto_tac (MI_css addSIs2 [squareCI] addSEs2 [squareE])); -by (REPEAT (ares_tac [exI, action_use Step1_4_4b] 1)); -by (force_tac (MI_css addSIs2 [WriteS4]) 1); -by (auto_tac (MI_css addSDs2 [Step2_2a])); -qed "Step2_2"; - -Goal "|- []( (ALL p. ImpNext p) \ -\ & (ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) \ -\ & (ALL p. $ImpInv rmhist p) \ -\ & [EX q. Write rmCh mm ires q l]_(mm!l)) \ -\ --> [][EX q. Write memCh mm (resbar rmhist) q l]_(mm!l)"; -by (force_tac (MI_css addSEs2 [STL4E] addSDs2 [Step2_2]) 1); -qed "Step2_lemma"; - -Goal "|- #l : #MemLoc & (ALL p. IPImp p & HistP rmhist p) \ -\ --> MSpec memCh mm (resbar rmhist) l"; -by (auto_tac (MI_css addsimps2 [MSpec_def])); -by (force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 1); -by (auto_tac (MI_css addSIs2 [Step2_lemma] - addsimps2 [split_box_conj,all_box])); -by (force_tac (MI_css addsimps2 [IPImp_def,MSpec_def]) 4); -by (auto_tac (MI_css addsimps2 [split_box_conj] addSEs2 [allE] addSDs2 [GoodImpl])); -qed "Step2"; - -(* ----------------------------- Main theorem --------------------------------- *) -section "Memory implementation"; - -(* The combination of a legal caller, the memory clerk, the RPC component, - and a reliable memory implement the unreliable memory. -*) - -(* Implementation of internal specification by combination of implementation - and history variable with explicit refinement mapping -*) -Goal "|- Implementation & Hist rmhist --> IUSpec memCh mm (resbar rmhist)"; -by (auto_tac (MI_css addsimps2 [IUSpec_def,Implementation_def,IPImp_def, - MClkISpec_def,RPCISpec_def,IRSpec_def,Hist_def] - addSIs2 [Step1,Step2])); -qed "Impl_IUSpec"; - -(* The main theorem: introduce hiding and eliminate history variable. *) -Goal "|- Implementation --> USpec memCh"; -by (Clarsimp_tac 1); -by (forward_tac [temp_use History] 1); -by (auto_tac (MI_css addsimps2 [USpec_def] - addIs2 [eexI, Impl_IUSpec, MI_base] - addSEs2 [eexE])); -qed "Implementation"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemoryImplementation.thy --- a/src/HOL/TLA/Memory/MemoryImplementation.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/MemoryImplementation.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: MemoryImplementation - Logic Image: TLA - - RPC-Memory example: Memory implementation -*) +header {* RPC-Memory example: Memory implementation *} theory MemoryImplementation imports Memory RPC MemClerk @@ -177,6 +174,1143 @@ (rtrner crCh!p, caller rmCh!p, rst!p), (mm!l, rtrner rmCh!p, ires!p))" -ML {* use_legacy_bindings (the_context ()) *} +(* + The main theorem is theorem "Implementation" at the end of this file, + which shows that the composition of a reliable memory, an RPC component, and + a memory clerk implements an unreliable memory. The files "MIsafe.ML" and + "MIlive.ML" contain lower-level lemmas for the safety and liveness parts. + + Steps are (roughly) numbered as in the hand proof. +*) + +(* --------------------------- automatic prover --------------------------- *) + +declare if_weak_cong [cong del] + +ML {* 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 it can be a lot faster than MI_css) +*) + +ML {* +val MI_fast_css = + let + val (cs,ss) = MI_css + in + (cs addSEs [temp_use (thm "squareE")], + ss addSSolver (mk_solver "" (fn thms => assume_tac ORELSE' (etac notE)))) + end; + +val temp_elim = make_elim o temp_use; +*} + + + +(****************************** The history variable ******************************) + +section "History variable" + +lemma HistoryLemma: "|- Init(ALL p. ImpInit p) & [](ALL p. ImpNext p) + --> (EEX rmhist. Init(ALL p. HInit rmhist p) + & [](ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)))" + apply clarsimp + apply (rule historyI) + apply assumption+ + apply (rule MI_base) + apply (tactic {* action_simp_tac (simpset () addsimps [thm "HInit_def"]) [] [] 1 *}) + apply (erule fun_cong) + apply (tactic {* action_simp_tac (simpset () addsimps [thm "HNext_def"]) + [thm "busy_squareI"] [] 1 *}) + apply (erule fun_cong) + done + +lemma History: "|- Implementation --> (EEX rmhist. Hist rmhist)" + apply clarsimp + apply (rule HistoryLemma [temp_use, THEN eex_mono]) + prefer 3 + apply (force simp: Hist_def HistP_def Init_def all_box [try_rewrite] + split_box_conj [try_rewrite]) + apply (auto simp: 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 [temp_use] split_box_conj [temp_use]) + done + +(******************************** The safety part *********************************) + +section "The safety part" + +(* ------------------------- Include lower-level lemmas ------------------------- *) + +(* RPCFailure notin MemVals U {OK,BadArg} *) + +lemma MVOKBAnotRF: "MVOKBA x ==> x ~= RPCFailure" + apply (unfold MVOKBA_def) + apply auto + done + +(* NotAResult notin MemVals U {OK,BadArg,RPCFailure} *) + +lemma MVOKBARFnotNR: "MVOKBARF x ==> x ~= NotAResult" + apply (unfold MVOKBARF_def) + apply auto + done + +(* ================ Si's are mutually exclusive ================================ *) +(* Si and Sj are mutually exclusive for i # j. This helps to simplify the big + conditional in the definition of resbar when doing the step-simulation proof. + We prove a weaker result, which suffices for our purposes: + Si implies (not Sj), for j S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p & + ~S4 rmhist p & ~S5 rmhist p & ~S6 rmhist p" +by (auto_tac (MI_css addsimps2 [S_def, S1_def, S2_def, + S3_def, S4_def, S5_def, S6_def])); +qed "S1_excl"; +*) + +lemma S2_excl: "|- S2 rmhist p --> S2 rmhist p & ~S1 rmhist p" + by (auto simp: S_def S1_def S2_def) + +lemma S3_excl: "|- S3 rmhist p --> S3 rmhist p & ~S1 rmhist p & ~S2 rmhist p" + by (auto simp: S_def S1_def S2_def S3_def) + +lemma S4_excl: "|- S4 rmhist p --> S4 rmhist p & ~S1 rmhist p & ~S2 rmhist p & ~S3 rmhist p" + by (auto simp: S_def S1_def S2_def S3_def S4_def) + +lemma S5_excl: "|- S5 rmhist p --> S5 rmhist p & ~S1 rmhist p & ~S2 rmhist p + & ~S3 rmhist p & ~S4 rmhist p" + by (auto simp: S_def S1_def S2_def S3_def S4_def S5_def) + +lemma S6_excl: "|- S6 rmhist p --> S6 rmhist p & ~S1 rmhist p & ~S2 rmhist p + & ~S3 rmhist p & ~S4 rmhist p & ~S5 rmhist p" + by (auto simp: S_def S1_def S2_def S3_def S4_def S5_def S6_def) + + +(* ==================== Lemmas about the environment ============================== *) + +lemma Envbusy: "|- $(Calling memCh p) --> ~ENext p" + by (auto simp: 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 + determine which component actions are possible and what state they result in. +*) + +(* ------------------------------ State S1 ---------------------------------------- *) + +lemma S1Env: "|- ENext p & $(S1 rmhist p) & unchanged (c p, r p, m p, rmhist!p) + --> (S2 rmhist p)$" + by (force simp: ENext_def Call_def c_def r_def m_def + caller_def rtrner_def MVNROKBA_def S_def S1_def S2_def Calling_def) + +lemma S1ClerkUnch: "|- [MClkNext memCh crCh cst p]_(c p) & $(S1 rmhist p) --> unchanged (c p)" + by (tactic {* auto_tac (MI_fast_css addSDs2 [temp_use (thm "MClkidle")] + addsimps2 [thm "S_def", thm "S1_def"]) *}) + +lemma S1RPCUnch: "|- [RPCNext crCh rmCh rst p]_(r p) & $(S1 rmhist p) --> unchanged (r p)" + by (tactic {* auto_tac (MI_fast_css addSDs2 [temp_use (thm "RPCidle")] + addsimps2 [thm "S_def", thm "S1_def"]) *}) + +lemma S1MemUnch: "|- [RNext rmCh mm ires p]_(m p) & $(S1 rmhist p) --> unchanged (m p)" + by (tactic {* auto_tac (MI_fast_css addSDs2 [temp_use (thm "Memoryidle")] + addsimps2 [thm "S_def", thm "S1_def"]) *}) + +lemma S1Hist: "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S1 rmhist p) + --> unchanged (rmhist!p)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "HNext_def", thm "S_def", + thm "S1_def", thm "MemReturn_def", thm "RPCFail_def", thm "MClkReply_def", + thm "Return_def"]) [] [temp_use (thm "squareE")] 1 *}) + + +(* ------------------------------ State S2 ---------------------------------------- *) + +lemma S2EnvUnch: "|- [ENext p]_(e p) & $(S2 rmhist p) --> unchanged (e p)" + by (auto dest!: Envbusy [temp_use] simp: S_def S2_def) + +lemma S2Clerk: "|- MClkNext memCh crCh cst p & $(S2 rmhist p) --> MClkFwd memCh crCh cst p" + by (auto simp: MClkNext_def MClkRetry_def MClkReply_def S_def S2_def) + +lemma S2Forward: "|- $(S2 rmhist p) & MClkFwd memCh crCh cst p + & unchanged (e p, r p, m p, rmhist!p) + --> (S3 rmhist p)$" + by (tactic {* action_simp_tac (simpset () addsimps [thm "MClkFwd_def", + thm "Call_def", thm "e_def", thm "r_def", thm "m_def", thm "caller_def", + thm "rtrner_def", thm "S_def", thm "S2_def", thm "S3_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S2RPCUnch: "|- [RPCNext crCh rmCh rst p]_(r p) & $(S2 rmhist p) --> unchanged (r p)" + by (auto simp: S_def S2_def dest!: RPCidle [temp_use]) + +lemma S2MemUnch: "|- [RNext rmCh mm ires p]_(m p) & $(S2 rmhist p) --> unchanged (m p)" + by (auto simp: S_def S2_def dest!: Memoryidle [temp_use]) + +lemma S2Hist: "|- [HNext rmhist p]_(c p,r p,m p,rmhist!p) & $(S2 rmhist p) + --> unchanged (rmhist!p)" + by (tactic {* auto_tac (MI_fast_css addsimps2 [thm "HNext_def", thm "MemReturn_def", + thm "RPCFail_def", thm "MClkReply_def", thm "Return_def", thm "S_def", thm "S2_def"]) *}) + +(* ------------------------------ State S3 ---------------------------------------- *) + +lemma S3EnvUnch: "|- [ENext p]_(e p) & $(S3 rmhist p) --> unchanged (e p)" + by (auto dest!: Envbusy [temp_use] simp: S_def S3_def) + +lemma S3ClerkUnch: "|- [MClkNext memCh crCh cst p]_(c p) & $(S3 rmhist p) --> unchanged (c p)" + by (auto dest!: MClkbusy [temp_use] simp: square_def S_def S3_def) + +lemma S3LegalRcvArg: "|- S3 rmhist p --> IsLegalRcvArg>" + by (auto simp: IsLegalRcvArg_def MClkRelayArg_def S_def S3_def) + +lemma S3RPC: "|- RPCNext crCh rmCh rst p & $(S3 rmhist p) + --> RPCFwd crCh rmCh rst p | RPCFail crCh rmCh rst p" + apply clarsimp + apply (frule S3LegalRcvArg [action_use]) + apply (auto simp: RPCNext_def RPCReject_def RPCReply_def S_def S3_def) + done + +lemma S3Forward: "|- RPCFwd crCh rmCh rst p & HNext rmhist p & $(S3 rmhist p) + & unchanged (e p, c p, m p) + --> (S4 rmhist p)$ & unchanged (rmhist!p)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "RPCFwd_def", + thm "HNext_def", thm "MemReturn_def", thm "RPCFail_def", + thm "MClkReply_def", thm "Return_def", thm "Call_def", thm "e_def", + thm "c_def", thm "m_def", thm "caller_def", thm "rtrner_def", thm "S_def", + thm "S3_def", thm "S4_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S3Fail: "|- RPCFail crCh rmCh rst p & $(S3 rmhist p) & HNext rmhist p + & unchanged (e p, c p, m p) + --> (S6 rmhist p)$" + by (tactic {* action_simp_tac (simpset () addsimps [thm "HNext_def", + thm "RPCFail_def", thm "Return_def", thm "e_def", thm "c_def", + thm "m_def", thm "caller_def", thm "rtrner_def", thm "MVOKBARF_def", + thm "S_def", thm "S3_def", thm "S6_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S3MemUnch: "|- [RNext rmCh mm ires p]_(m p) & $(S3 rmhist p) --> unchanged (m p)" + by (auto simp: S_def S3_def dest!: Memoryidle [temp_use]) + +lemma S3Hist: "|- HNext rmhist p & $(S3 rmhist p) & unchanged (r p) --> unchanged (rmhist!p)" + by (auto simp: HNext_def MemReturn_def RPCFail_def MClkReply_def + Return_def r_def rtrner_def S_def S3_def Calling_def) + +(* ------------------------------ State S4 ---------------------------------------- *) + +lemma S4EnvUnch: "|- [ENext p]_(e p) & $(S4 rmhist p) --> unchanged (e p)" + by (auto simp: S_def S4_def dest!: Envbusy [temp_use]) + +lemma S4ClerkUnch: "|- [MClkNext memCh crCh cst p]_(c p) & $(S4 rmhist p) --> unchanged (c p)" + by (auto simp: S_def S4_def dest!: MClkbusy [temp_use]) + +lemma S4RPCUnch: "|- [RPCNext crCh rmCh rst p]_(r p) & $(S4 rmhist p) --> unchanged (r p)" + by (tactic {* auto_tac (MI_fast_css addsimps2 [thm "S_def", thm "S4_def"] + addSDs2 [temp_use (thm "RPCbusy")]) *}) + +lemma S4ReadInner: "|- 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)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "ReadInner_def", + thm "GoodRead_def", thm "BadRead_def", thm "HNext_def", thm "MemReturn_def", + thm "RPCFail_def", thm "MClkReply_def", thm "Return_def", thm "e_def", + thm "c_def", thm "r_def", thm "rtrner_def", thm "caller_def", + thm "MVNROKBA_def", thm "S_def", thm "S4_def", thm "RdRequest_def", + thm "Calling_def", thm "MemInv_def"]) [] [] 1 *}) + +lemma 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)" + by (auto simp: Read_def dest!: S4ReadInner [temp_use]) + +lemma S4WriteInner: "|- 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)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "WriteInner_def", + thm "GoodWrite_def", thm "BadWrite_def", thm "HNext_def", thm "MemReturn_def", + thm "RPCFail_def", thm "MClkReply_def", thm "Return_def", thm "e_def", + thm "c_def", thm "r_def", thm "rtrner_def", thm "caller_def", thm "MVNROKBA_def", + thm "S_def", thm "S4_def", thm "WrRequest_def", thm "Calling_def"]) [] [] 1 *}) + +lemma 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)" + by (auto simp: Write_def dest!: S4WriteInner [temp_use]) + +lemma WriteS4: "|- $ImpInv rmhist p & Write rmCh mm ires p l --> $S4 rmhist p" + by (auto simp: Write_def WriteInner_def ImpInv_def + WrRequest_def S_def S1_def S2_def S3_def S4_def S5_def S6_def) + +lemma S4Return: "|- MemReturn rmCh ires p & $S4 rmhist p & unchanged (e p, c p, r p) + & HNext rmhist p + --> (S5 rmhist p)$" + by (auto simp: 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) + +lemma S4Hist: "|- HNext rmhist p & $S4 rmhist p & (m p)$ = $(m p) --> (rmhist!p)$ = $(rmhist!p)" + by (auto simp: HNext_def MemReturn_def RPCFail_def MClkReply_def + Return_def m_def rtrner_def S_def S4_def Calling_def) + +(* ------------------------------ State S5 ---------------------------------------- *) + +lemma S5EnvUnch: "|- [ENext p]_(e p) & $(S5 rmhist p) --> unchanged (e p)" + by (auto simp: S_def S5_def dest!: Envbusy [temp_use]) + +lemma S5ClerkUnch: "|- [MClkNext memCh crCh cst p]_(c p) & $(S5 rmhist p) --> unchanged (c p)" + by (auto simp: S_def S5_def dest!: MClkbusy [temp_use]) + +lemma S5RPC: "|- RPCNext crCh rmCh rst p & $(S5 rmhist p) + --> RPCReply crCh rmCh rst p | RPCFail crCh rmCh rst p" + by (auto simp: RPCNext_def RPCReject_def RPCFwd_def S_def S5_def) + +lemma S5Reply: "|- RPCReply crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) + --> (S6 rmhist p)$" + by (tactic {* action_simp_tac (simpset () addsimps [thm "RPCReply_def", + thm "Return_def", thm "e_def", thm "c_def", thm "m_def", thm "MVOKBA_def", + thm "MVOKBARF_def", thm "caller_def", thm "rtrner_def", thm "S_def", + thm "S5_def", thm "S6_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S5Fail: "|- RPCFail crCh rmCh rst p & $(S5 rmhist p) & unchanged (e p, c p, m p,rmhist!p) + --> (S6 rmhist p)$" + by (tactic {* action_simp_tac (simpset () addsimps [thm "RPCFail_def", + thm "Return_def", thm "e_def", thm "c_def", thm "m_def", + thm "MVOKBARF_def", thm "caller_def", thm "rtrner_def", + thm "S_def", thm "S5_def", thm "S6_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S5MemUnch: "|- [RNext rmCh mm ires p]_(m p) & $(S5 rmhist p) --> unchanged (m p)" + by (auto simp: S_def S5_def dest!: Memoryidle [temp_use]) + +lemma S5Hist: "|- [HNext rmhist p]_(c p, r p, m p, rmhist!p) & $(S5 rmhist p) + --> (rmhist!p)$ = $(rmhist!p)" + by (tactic {* auto_tac (MI_fast_css addsimps2 [thm "HNext_def", + thm "MemReturn_def", thm "RPCFail_def", thm "MClkReply_def", thm "Return_def", + thm "S_def", thm "S5_def"]) *}) + +(* ------------------------------ State S6 ---------------------------------------- *) + +lemma S6EnvUnch: "|- [ENext p]_(e p) & $(S6 rmhist p) --> unchanged (e p)" + by (auto simp: S_def S6_def dest!: Envbusy [temp_use]) + +lemma S6Clerk: "|- MClkNext memCh crCh cst p & $(S6 rmhist p) + --> MClkRetry memCh crCh cst p | MClkReply memCh crCh cst p" + by (auto simp: MClkNext_def MClkFwd_def S_def S6_def) + +lemma S6Retry: "|- MClkRetry memCh crCh cst p & HNext rmhist p & $S6 rmhist p + & unchanged (e p,r p,m p) + --> (S3 rmhist p)$ & unchanged (rmhist!p)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "HNext_def", + thm "MClkReply_def", thm "MClkRetry_def", thm "Call_def", thm "Return_def", + thm "e_def", thm "r_def", thm "m_def", thm "caller_def", thm "rtrner_def", + thm "S_def", thm "S6_def", thm "S3_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S6Reply: "|- MClkReply memCh crCh cst p & HNext rmhist p & $S6 rmhist p + & unchanged (e p,r p,m p) + --> (S1 rmhist p)$" + by (tactic {* action_simp_tac (simpset () addsimps [thm "HNext_def", + thm "MemReturn_def", thm "RPCFail_def", thm "Return_def", thm "MClkReply_def", + thm "e_def", thm "r_def", thm "m_def", thm "caller_def", thm "rtrner_def", + thm "S_def", thm "S6_def", thm "S1_def", thm "Calling_def"]) [] [] 1 *}) + +lemma S6RPCUnch: "|- [RPCNext crCh rmCh rst p]_(r p) & $S6 rmhist p --> unchanged (r p)" + by (auto simp: S_def S6_def dest!: RPCidle [temp_use]) + +lemma S6MemUnch: "|- [RNext rmCh mm ires p]_(m p) & $(S6 rmhist p) --> unchanged (m p)" + by (auto simp: S_def S6_def dest!: Memoryidle [temp_use]) + +lemma S6Hist: "|- HNext rmhist p & $S6 rmhist p & (c p)$ = $(c p) --> (rmhist!p)$ = $(rmhist!p)" + by (auto simp: HNext_def MClkReply_def Return_def c_def rtrner_def S_def S6_def Calling_def) + + +section "Correctness of predicate-action diagram" + + +(* ========== Step 1.1 ================================================= *) +(* The implementation's initial condition implies the state predicate S1 *) + +lemma Step1_1: "|- ImpInit p & HInit rmhist p --> S1 rmhist p" + by (tactic {* auto_tac (MI_fast_css addsimps2 [thm "MVNROKBA_def", + thm "MClkInit_def", thm "RPCInit_def", thm "PInit_def", thm "HInit_def", + thm "ImpInit_def", thm "S_def", thm "S1_def"]) *}) + +(* ========== Step 1.2 ================================================== *) +(* Figure 16 is a predicate-action diagram for the implementation. *) + +lemma Step1_2_1: "|- [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)" + apply (tactic {* action_simp_tac (simpset() addsimps [thm "ImpNext_def"]) [] + (map temp_elim [thm "S1ClerkUnch", thm "S1RPCUnch", thm "S1MemUnch", thm "S1Hist"]) 1 *}) + apply (tactic {* auto_tac (MI_fast_css addSIs2 [temp_use (thm "S1Env")]) *}) + done + +lemma Step1_2_2: "|- [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)" + apply (tactic {* action_simp_tac (simpset() addsimps [thm "ImpNext_def"]) [] + (map temp_elim [thm "S2EnvUnch", thm "S2RPCUnch", thm "S2MemUnch", thm "S2Hist"]) 1 *}) + apply (tactic {* auto_tac (MI_fast_css addSIs2 [temp_use (thm "S2Clerk"), + temp_use (thm "S2Forward")]) *}) + done + +lemma Step1_2_3: "|- [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))" + apply (tactic {* action_simp_tac (simpset() addsimps [thm "ImpNext_def"]) [] + (map temp_elim [thm "S3EnvUnch", thm "S3ClerkUnch", thm "S3MemUnch"]) 1 *}) + apply (tactic {* action_simp_tac (simpset()) [] + (thm "squareE" :: map temp_elim [thm "S3RPC", thm "S3Forward", thm "S3Fail"]) 1 *}) + apply (auto dest!: S3Hist [temp_use]) + done + +lemma Step1_2_4: "|- [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))" + apply (tactic {* action_simp_tac (simpset() addsimps [thm "ImpNext_def"]) [] + (map temp_elim [thm "S4EnvUnch", thm "S4ClerkUnch", thm "S4RPCUnch"]) 1 *}) + apply (tactic {* action_simp_tac (simpset() addsimps [thm "RNext_def"]) [] + (thm "squareE" :: map temp_elim [thm "S4Read", thm "S4Write", thm "S4Return"]) 1 *}) + apply (auto dest!: S4Hist [temp_use]) + done + +lemma Step1_2_5: "|- [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))" + apply (tactic {* action_simp_tac (simpset() addsimps [thm "ImpNext_def"]) [] + (map temp_elim [thm "S5EnvUnch", thm "S5ClerkUnch", thm "S5MemUnch", thm "S5Hist"]) 1 *}) + apply (tactic {* action_simp_tac (simpset()) [] [thm "squareE", temp_elim (thm "S5RPC")] 1 *}) + apply (tactic {* auto_tac (MI_fast_css addSDs2 + [temp_use (thm "S5Reply"), temp_use (thm "S5Fail")]) *}) + done + +lemma Step1_2_6: "|- [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))" + apply (tactic {* action_simp_tac (simpset() addsimps [thm "ImpNext_def"]) [] + (map temp_elim [thm "S6EnvUnch", thm "S6RPCUnch", thm "S6MemUnch"]) 1 *}) + apply (tactic {* action_simp_tac (simpset()) [] + (thm "squareE" :: map temp_elim [thm "S6Clerk", thm "S6Retry", thm "S6Reply"]) 1 *}) + apply (auto dest: S6Hist [temp_use]) + done + +(* -------------------------------------------------------------------------- + Step 1.3: S1 implies the barred initial condition. +*) + +section "Initialization (Step 1.3)" + +lemma Step1_3: "|- S1 rmhist p --> PInit (resbar rmhist) p" + by (tactic {* action_simp_tac (simpset () addsimps [thm "resbar_def", + thm "PInit_def", thm "S_def", thm "S1_def"]) [] [] 1 *}) + +(* ---------------------------------------------------------------------- + Step 1.4: Implementation's next-state relation simulates specification's + next-state relation (with appropriate substitutions) +*) + +section "Step simulation (Step 1.4)" + +lemma Step1_4_1: "|- ENext p & $S1 rmhist p & (S2 rmhist p)$ & unchanged (c p, r p, m p) + --> unchanged (rtrner memCh!p, resbar rmhist!p)" + by (tactic {* auto_tac (MI_fast_css addsimps2 [thm "c_def", thm "r_def", + thm "m_def", thm "resbar_def"]) *}) + +lemma Step1_4_2: "|- 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)" + by (tactic {* action_simp_tac + (simpset() addsimps [thm "MClkFwd_def", thm "e_def", thm "r_def", thm "m_def", + thm "resbar_def", thm "S_def", thm "S2_def", thm "S3_def"]) [] [] 1 *}) + +lemma Step1_4_3a: "|- 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)" + apply clarsimp + apply (drule S3_excl [temp_use] S4_excl [temp_use])+ + apply (tactic {* action_simp_tac (simpset () addsimps [thm "e_def", + thm "c_def", thm "m_def", thm "resbar_def", thm "S_def", thm "S3_def"]) [] [] 1 *}) + done + +lemma Step1_4_3b: "|- RPCFail crCh rmCh rst p & $S3 rmhist p & (S6 rmhist p)$ + & unchanged (e p, c p, m p) + --> MemFail memCh (resbar rmhist) p" + apply clarsimp + apply (drule S6_excl [temp_use]) + apply (auto simp: RPCFail_def MemFail_def e_def c_def m_def resbar_def) + apply (force simp: S3_def S_def) + apply (auto simp: Return_def) + done + +lemma Step1_4_4a1: "|- $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" + apply clarsimp + apply (drule S4_excl [temp_use])+ + apply (tactic {* action_simp_tac (simpset () addsimps [thm "ReadInner_def", + thm "GoodRead_def", thm "BadRead_def", thm "e_def", thm "c_def", thm "m_def"]) [] [] 1 *}) + apply (auto simp: resbar_def) + apply (tactic {* ALLGOALS (action_simp_tac + (simpset() addsimps [thm "RPCRelayArg_def", thm "MClkRelayArg_def", + thm "S_def", thm "S4_def", thm "RdRequest_def", thm "MemInv_def"]) + [] [thm "impE", thm "MemValNotAResultE"]) *}) + done + +lemma Step1_4_4a: "|- 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" + by (force simp: Read_def elim!: Step1_4_4a1 [temp_use]) + +lemma Step1_4_4b1: "|- $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" + apply clarsimp + apply (drule S4_excl [temp_use])+ + apply (tactic {* action_simp_tac (simpset () addsimps + [thm "WriteInner_def", thm "GoodWrite_def", thm "BadWrite_def", thm "e_def", + thm "c_def", thm "m_def"]) [] [] 1 *}) + apply (auto simp: resbar_def) + apply (tactic {* ALLGOALS (action_simp_tac (simpset () addsimps + [thm "RPCRelayArg_def", thm "MClkRelayArg_def", thm "S_def", + thm "S4_def", thm "WrRequest_def"]) [] []) *}) + done + +lemma Step1_4_4b: "|- 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" + by (force simp: Write_def elim!: Step1_4_4b1 [temp_use]) + +lemma Step1_4_4c: "|- MemReturn rmCh ires p & $S4 rmhist p & (S5 rmhist p)$ + & unchanged (e p, c p, r p) + --> unchanged (rtrner memCh!p, resbar rmhist!p)" + apply (tactic {* action_simp_tac (simpset () addsimps [thm "e_def", + thm "c_def", thm "r_def", thm "resbar_def"]) [] [] 1 *}) + apply (drule S4_excl [temp_use] S5_excl [temp_use])+ + apply (tactic {* auto_tac (MI_fast_css addsimps2 [thm "MemReturn_def", thm "Return_def"]) *}) + done + +lemma Step1_4_5a: "|- 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)" + apply clarsimp + apply (drule S5_excl [temp_use] S6_excl [temp_use])+ + apply (auto simp: e_def c_def m_def resbar_def) + apply (auto simp: RPCReply_def Return_def S5_def S_def dest!: MVOKBAnotRF [temp_use]) + done + +lemma Step1_4_5b: "|- RPCFail crCh rmCh rst p & $S5 rmhist p & (S6 rmhist p)$ + & unchanged (e p, c p, m p) + --> MemFail memCh (resbar rmhist) p" + apply clarsimp + apply (drule S6_excl [temp_use]) + apply (auto simp: e_def c_def m_def RPCFail_def Return_def MemFail_def resbar_def) + apply (auto simp: S5_def S_def) + done + +lemma Step1_4_6a: "|- MClkReply memCh crCh cst p & $S6 rmhist p & (S1 rmhist p)$ + & unchanged (e p, r p, m p) + --> MemReturn memCh (resbar rmhist) p" + apply clarsimp + apply (drule S6_excl [temp_use])+ + apply (tactic {* action_simp_tac (simpset () addsimps [thm "e_def", + thm "r_def", thm "m_def", thm "MClkReply_def", thm "MemReturn_def", + thm "Return_def", thm "resbar_def"]) [] [] 1 *}) + apply simp_all (* simplify if-then-else *) + apply (tactic {* ALLGOALS (action_simp_tac (simpset () addsimps + [thm "MClkReplyVal_def", thm "S6_def", thm "S_def"]) [] [thm "MVOKBARFnotNR"]) *}) + done + +lemma Step1_4_6b: "|- 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" + apply clarsimp + apply (drule S3_excl [temp_use])+ + apply (tactic {* action_simp_tac (simpset () addsimps [thm "e_def", thm "r_def", + thm "m_def", thm "MClkRetry_def", thm "MemFail_def", thm "resbar_def"]) [] [] 1 *}) + apply (auto simp: S6_def S_def) + done + +lemma S_lemma: "|- unchanged (e p, c p, r p, m p, rmhist!p) + --> unchanged (S rmhist ec cc rc cs rs hs1 hs2 p)" + by (auto simp: e_def c_def r_def m_def caller_def rtrner_def S_def Calling_def) + +lemma Step1_4_7H: "|- 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)" + apply clarsimp + apply (rule conjI) + apply (force simp: c_def) + apply (force simp: S1_def S2_def S3_def S4_def S5_def S6_def intro!: S_lemma [temp_use]) + done + +lemma Step1_4_7: "|- 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)" + apply (rule actionI) + apply (unfold action_rews) + apply (rule impI) + apply (frule Step1_4_7H [temp_use]) + apply (auto simp: e_def c_def r_def m_def rtrner_def resbar_def) + done + +(* Frequently needed abbreviation: distinguish between idling and non-idling + steps of the implementation, and try to solve the idling case by simplification +*) +ML {* +local + val actionI = thm "actionI"; + val action_rews = thms "action_rews"; + val Step1_4_7 = thm "Step1_4_7"; +in +fun split_idle_tac simps 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, + forward_tac [temp_use Step1_4_7] i, + asm_full_simp_tac (simpset() addsimps simps) i + ] +end +*} +(* ---------------------------------------------------------------------- + Combine steps 1.2 and 1.4 to prove that the implementation satisfies + the specification's next-state relation. +*) + +(* Steps that leave all variables unchanged are safe, so I may assume + that some variable changes in the proof that a step is safe. *) +lemma unchanged_safe: "|- (~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)" + apply (tactic {* split_idle_tac [thm "square_def"] 1 *}) + apply force + done +(* turn into (unsafe, looping!) introduction rule *) +lemmas unchanged_safeI = impI [THEN unchanged_safe [action_use], standard] + +lemma S1safe: "|- $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)" + apply clarsimp + apply (rule unchanged_safeI) + apply (rule idle_squareI) + apply (auto dest!: Step1_2_1 [temp_use] Step1_4_1 [temp_use]) + done + +lemma S2safe: "|- $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)" + apply clarsimp + apply (rule unchanged_safeI) + apply (rule idle_squareI) + apply (auto dest!: Step1_2_2 [temp_use] Step1_4_2 [temp_use]) + done + +lemma S3safe: "|- $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)" + apply clarsimp + apply (rule unchanged_safeI) + apply (auto dest!: Step1_2_3 [temp_use]) + apply (auto simp: square_def UNext_def dest!: Step1_4_3a [temp_use] Step1_4_3b [temp_use]) + done + +lemma S4safe: "|- $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)" + apply clarsimp + apply (rule unchanged_safeI) + apply (auto dest!: Step1_2_4 [temp_use]) + apply (auto simp: square_def UNext_def RNext_def + dest!: Step1_4_4a [temp_use] Step1_4_4b [temp_use] Step1_4_4c [temp_use]) + done + +lemma S5safe: "|- $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)" + apply clarsimp + apply (rule unchanged_safeI) + apply (auto dest!: Step1_2_5 [temp_use]) + apply (auto simp: square_def UNext_def dest!: Step1_4_5a [temp_use] Step1_4_5b [temp_use]) + done + +lemma S6safe: "|- $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)" + apply clarsimp + apply (rule unchanged_safeI) + apply (auto dest!: Step1_2_6 [temp_use]) + apply (auto simp: square_def UNext_def RNext_def + dest!: Step1_4_6a [temp_use] Step1_4_6b [temp_use]) + done + +(* ---------------------------------------------------------------------- + Step 1.5: Temporal refinement proof, based on previous steps. +*) + +section "The liveness part" + +(* Liveness assertions for the different implementation states, based on the + fairness conditions. Prove subgoals of WF1 / SF1 rules as separate lemmas + for readability. Reuse action proofs from safety part. +*) + +(* ------------------------------ State S1 ------------------------------ *) + +lemma S1_successors: "|- $S1 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + --> (S1 rmhist p)$ | (S2 rmhist p)$" + apply (tactic "split_idle_tac [] 1") + apply (auto dest!: Step1_2_1 [temp_use]) + done + +(* Show that the implementation can satisfy the high-level fairness requirements + by entering the state S1 infinitely often. +*) + +lemma S1_RNextdisabled: "|- S1 rmhist p --> + ~Enabled (_(rtrner memCh!p, resbar rmhist!p))" + apply (tactic {* action_simp_tac (simpset () addsimps [thm "angle_def", + thm "S_def", thm "S1_def"]) [notI] [thm "enabledE", temp_elim (thm "Memoryidle")] 1 *}) + apply force + done + +lemma S1_Returndisabled: "|- S1 rmhist p --> + ~Enabled (_(rtrner memCh!p, resbar rmhist!p))" + by (tactic {* action_simp_tac (simpset () addsimps [thm "angle_def", thm "MemReturn_def", + thm "Return_def", thm "S_def", thm "S1_def"]) [notI] [thm "enabledE"] 1 *}) + +lemma RNext_fair: "|- []<>S1 rmhist p + --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + by (auto simp: WF_alt [try_rewrite] intro!: S1_RNextdisabled [temp_use] + elim!: STL4E [temp_use] DmdImplE [temp_use]) + +lemma Return_fair: "|- []<>S1 rmhist p + --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + by (auto simp: WF_alt [try_rewrite] + intro!: S1_Returndisabled [temp_use] elim!: STL4E [temp_use] DmdImplE [temp_use]) + +(* ------------------------------ State S2 ------------------------------ *) + +lemma S2_successors: "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + --> (S2 rmhist p)$ | (S3 rmhist p)$" + apply (tactic "split_idle_tac [] 1") + apply (auto dest!: Step1_2_2 [temp_use]) + done + +lemma S2MClkFwd_successors: "|- ($S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) + & _(c p) + --> (S3 rmhist p)$" + by (auto simp: angle_def dest!: Step1_2_2 [temp_use]) + +lemma S2MClkFwd_enabled: "|- $S2 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + --> $Enabled (_(c p))" + apply (auto simp: c_def intro!: MClkFwd_ch_enabled [temp_use] MClkFwd_enabled [temp_use]) + apply (cut_tac MI_base) + apply (blast dest: base_pair) + apply (simp_all add: S_def S2_def) + done + +lemma S2_live: "|- [](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)" + by (rule WF1 S2_successors S2MClkFwd_successors S2MClkFwd_enabled)+ + +(* ------------------------------ State S3 ------------------------------ *) + +lemma S3_successors: "|- $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)$" + apply (tactic "split_idle_tac [] 1") + apply (auto dest!: Step1_2_3 [temp_use]) + done + +lemma S3RPC_successors: "|- ($S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) + & _(r p) + --> (S4 rmhist p | S6 rmhist p)$" + apply (auto simp: angle_def dest!: Step1_2_3 [temp_use]) + done + +lemma S3RPC_enabled: "|- $S3 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + --> $Enabled (_(r p))" + apply (auto simp: r_def intro!: RPCFail_Next_enabled [temp_use] RPCFail_enabled [temp_use]) + apply (cut_tac MI_base) + apply (blast dest: base_pair) + apply (simp_all add: S_def S3_def) + done + +lemma S3_live: "|- [](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)" + by (rule WF1 S3_successors S3RPC_successors S3RPC_enabled)+ + +(* ------------- State S4 -------------------------------------------------- *) + +lemma S4_successors: "|- $S4 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + & (ALL l. $MemInv mm l) + --> (S4 rmhist p)$ | (S5 rmhist p)$" + apply (tactic "split_idle_tac [] 1") + apply (auto dest!: Step1_2_4 [temp_use]) + done + +(* --------- State S4a: S4 /\ (ires p = NotAResult) ------------------------ *) + +lemma S4a_successors: "|- $(S4 rmhist p & ires!p = #NotAResult) + & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l) + --> (S4 rmhist p & ires!p = #NotAResult)$ + | ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$" + apply (tactic {* split_idle_tac [thm "m_def"] 1 *}) + apply (auto dest!: Step1_2_4 [temp_use]) + done + +lemma S4aRNext_successors: "|- ($(S4 rmhist p & ires!p = #NotAResult) + & ImpNext p & [HNext rmhist p]_(c p,r p,m p,rmhist!p) & (ALL l. $MemInv mm l)) + & _(m p) + --> ((S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p)$" + by (auto simp: angle_def + dest!: Step1_2_4 [temp_use] ReadResult [temp_use] WriteResult [temp_use]) + +lemma S4aRNext_enabled: "|- $(S4 rmhist p & ires!p = #NotAResult) + & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l) + --> $Enabled (_(m p))" + apply (auto simp: m_def intro!: RNext_enabled [temp_use]) + apply (cut_tac MI_base) + apply (blast dest: base_pair) + apply (simp add: S_def S4_def) + done + +lemma S4a_live: "|- [](ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + & (ALL 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)" + by (rule WF1 S4a_successors S4aRNext_successors S4aRNext_enabled)+ + +(* ---------- State S4b: S4 /\ (ires p # NotAResult) --------------------------- *) + +lemma S4b_successors: "|- $(S4 rmhist p & ires!p ~= #NotAResult) + & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) & (ALL l. $MemInv mm l) + --> (S4 rmhist p & ires!p ~= #NotAResult)$ | (S5 rmhist p)$" + apply (tactic {* split_idle_tac [thm "m_def"] 1 *}) + apply (auto dest!: WriteResult [temp_use] Step1_2_4 [temp_use] ReadResult [temp_use]) + done + +lemma S4bReturn_successors: "|- ($(S4 rmhist p & ires!p ~= #NotAResult) + & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + & (ALL l. $MemInv mm l)) & _(m p) + --> (S5 rmhist p)$" + by (force simp: angle_def dest!: Step1_2_4 [temp_use] dest: ReturnNotReadWrite [temp_use]) + +lemma S4bReturn_enabled: "|- $(S4 rmhist p & ires!p ~= #NotAResult) + & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + & (ALL l. $MemInv mm l) + --> $Enabled (_(m p))" + apply (auto simp: m_def intro!: MemReturn_enabled [temp_use]) + apply (cut_tac MI_base) + apply (blast dest: base_pair) + apply (simp add: S_def S4_def) + done + +lemma S4b_live: "|- [](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)" + by (rule WF1 S4b_successors S4bReturn_successors S4bReturn_enabled)+ + +(* ------------------------------ State S5 ------------------------------ *) + +lemma S5_successors: "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + --> (S5 rmhist p)$ | (S6 rmhist p)$" + apply (tactic "split_idle_tac [] 1") + apply (auto dest!: Step1_2_5 [temp_use]) + done + +lemma S5RPC_successors: "|- ($S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) + & _(r p) + --> (S6 rmhist p)$" + by (auto simp: angle_def dest!: Step1_2_5 [temp_use]) + +lemma S5RPC_enabled: "|- $S5 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p) + --> $Enabled (_(r p))" + apply (auto simp: r_def intro!: RPCFail_Next_enabled [temp_use] RPCFail_enabled [temp_use]) + apply (cut_tac MI_base) + apply (blast dest: base_pair) + apply (simp_all add: S_def S5_def) + done + +lemma S5_live: "|- [](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)" + by (rule WF1 S5_successors S5RPC_successors S5RPC_enabled)+ + +(* ------------------------------ State S6 ------------------------------ *) + +lemma S6_successors: "|- $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)$" + apply (tactic "split_idle_tac [] 1") + apply (auto dest!: Step1_2_6 [temp_use]) + done + +lemma S6MClkReply_successors: + "|- ($S6 rmhist p & ImpNext p & [HNext rmhist p]_(c p,r p,m p, rmhist!p)) + & _(c p) + --> (S1 rmhist p)$" + by (auto simp: angle_def dest!: Step1_2_6 [temp_use] MClkReplyNotRetry [temp_use]) + +lemma MClkReplyS6: + "|- $ImpInv rmhist p & _(c p) --> $S6 rmhist p" + by (tactic {* action_simp_tac (simpset () addsimps [thm "angle_def", + thm "MClkReply_def", thm "Return_def", thm "ImpInv_def", thm "S_def", + thm "S1_def", thm "S2_def", thm "S3_def", thm "S4_def", thm "S5_def"]) [] [] 1 *}) + +lemma S6MClkReply_enabled: "|- S6 rmhist p --> Enabled (_(c p))" + apply (auto simp: c_def intro!: MClkReply_enabled [temp_use]) + apply (cut_tac MI_base) + apply (blast dest: base_pair) + apply (tactic {* ALLGOALS (action_simp_tac (simpset () + addsimps [thm "S_def", thm "S6_def"]) [] []) *}) + done + +lemma S6_live: "|- [](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)" + apply clarsimp + apply (subgoal_tac "sigma |= []<> (_ (c p))") + apply (erule InfiniteEnsures) + apply assumption + apply (tactic {* action_simp_tac (simpset()) [] + (map temp_elim [thm "MClkReplyS6", thm "S6MClkReply_successors"]) 1 *}) + apply (auto simp: SF_def) + apply (erule contrapos_np) + apply (auto intro!: S6MClkReply_enabled [temp_use] elim!: STL4E [temp_use] DmdImplE [temp_use]) + done + +(* --------------- aggregate leadsto properties----------------------------- *) + +lemma S5S6LeadstoS6: "sigma |= S5 rmhist p ~> S6 rmhist p + ==> sigma |= (S5 rmhist p | S6 rmhist p) ~> S6 rmhist p" + by (auto intro!: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use]) + +lemma S4bS5S6LeadstoS6: "[| 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" + by (auto intro!: LatticeDisjunctionIntro [temp_use] + S5S6LeadstoS6 [temp_use] intro: LatticeTransitivity [temp_use]) + +lemma S4S5S6LeadstoS6: "[| 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" + apply (subgoal_tac "sigma |= (S4 rmhist p & ires!p = #NotAResult) | + (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p ~> S6 rmhist p") + apply (erule_tac G = "PRED ((S4 rmhist p & ires!p = #NotAResult) | + (S4 rmhist p & ires!p ~= #NotAResult) | S5 rmhist p | S6 rmhist p)" in + LatticeTransitivity [temp_use]) + apply (force simp: Init_defs intro!: ImplLeadsto_gen [temp_use] necT [temp_use]) + apply (rule LatticeDisjunctionIntro [temp_use]) + apply (erule LatticeTransitivity [temp_use]) + apply (erule LatticeTriangle2 [temp_use]) + apply assumption + apply (auto intro!: S4bS5S6LeadstoS6 [temp_use]) + done + +lemma S3S4S5S6LeadstoS6: "[| 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" + apply (rule LatticeDisjunctionIntro [temp_use]) + apply (erule LatticeTriangle2 [temp_use]) + apply (rule S4S5S6LeadstoS6 [THEN LatticeTransitivity [temp_use]]) + apply (auto intro!: S4S5S6LeadstoS6 [temp_use] necT [temp_use] + intro: ImplLeadsto_gen [temp_use] simp: Init_defs) + done + +lemma S2S3S4S5S6LeadstoS6: "[| 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" + apply (rule LatticeDisjunctionIntro [temp_use]) + apply (rule LatticeTransitivity [temp_use]) + prefer 2 apply assumption + apply (rule S3S4S5S6LeadstoS6 [THEN LatticeTransitivity [temp_use]]) + apply (auto intro!: S3S4S5S6LeadstoS6 [temp_use] necT [temp_use] + intro: ImplLeadsto_gen [temp_use] simp: Init_defs) + done + +lemma NotS1LeadstoS6: "[| 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" + apply (rule S2S3S4S5S6LeadstoS6 [THEN LatticeTransitivity [temp_use]]) + apply assumption+ + apply (erule INV_leadsto [temp_use]) + apply (rule ImplLeadsto_gen [temp_use]) + apply (rule necT [temp_use]) + apply (auto simp: ImpInv_def Init_defs intro!: necT [temp_use]) + done + +lemma S1Infinite: "[| sigma |= ~S1 rmhist p ~> S6 rmhist p; + sigma |= []<>S6 rmhist p --> []<>S1 rmhist p |] + ==> sigma |= []<>S1 rmhist p" + apply (rule classical) + apply (tactic {* asm_lr_simp_tac (simpset() addsimps + [temp_use (thm "NotBox"), temp_rewrite (thm "NotDmd")]) 1 *}) + apply (auto elim!: leadsto_infinite [temp_use] mp dest!: DBImplBD [temp_use]) + done + +section "Refinement proof (step 1.5)" + +(* Prove invariants of the implementation: + a. memory invariant + b. "implementation invariant": always in states S1,...,S6 +*) +lemma Step1_5_1a: "|- IPImp p --> (ALL l. []$MemInv mm l)" + by (auto simp: IPImp_def box_stp_act [temp_use] intro!: MemoryInvariantAll [temp_use]) + +lemma Step1_5_1b: "|- Init(ImpInit p & HInit rmhist p) & [](ImpNext p) + & [][HNext rmhist p]_(c p, r p, m p, rmhist!p) & [](ALL l. $MemInv mm l) + --> []ImpInv rmhist p" + apply (tactic "inv_tac MI_css 1") + apply (auto simp: Init_def ImpInv_def box_stp_act [temp_use] + dest!: Step1_1 [temp_use] dest: S1_successors [temp_use] S2_successors [temp_use] + S3_successors [temp_use] S4_successors [temp_use] S5_successors [temp_use] + S6_successors [temp_use]) + done + +(*** Initialization ***) +lemma Step1_5_2a: "|- Init(ImpInit p & HInit rmhist p) --> Init(PInit (resbar rmhist) p)" + by (auto simp: Init_def intro!: Step1_1 [temp_use] Step1_3 [temp_use]) + +(*** step simulation ***) +lemma Step1_5_2b: "|- [](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)" + by (auto simp: ImpInv_def elim!: STL4E [temp_use] + dest!: S1safe [temp_use] S2safe [temp_use] S3safe [temp_use] S4safe [temp_use] + S5safe [temp_use] S6safe [temp_use]) + +(*** Liveness ***) +lemma GoodImpl: "|- IPImp p & HistP rmhist p + --> Init(ImpInit p & HInit rmhist p) + & [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) + & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) + & ImpLive p" + apply clarsimp + apply (subgoal_tac "sigma |= Init (ImpInit p & HInit rmhist p) & [] (ImpNext p) & + [][HNext rmhist p]_ (c p, r p, m p, rmhist!p) & [] (ALL l. $MemInv mm l)") + apply (auto simp: split_box_conj [try_rewrite] box_stp_act [try_rewrite] + dest!: Step1_5_1b [temp_use]) + apply (force simp: IPImp_def MClkIPSpec_def RPCIPSpec_def RPSpec_def + ImpLive_def c_def r_def m_def) + apply (force simp: IPImp_def MClkIPSpec_def RPCIPSpec_def RPSpec_def + HistP_def Init_def ImpInit_def) + apply (force simp: IPImp_def MClkIPSpec_def RPCIPSpec_def RPSpec_def + ImpNext_def c_def r_def m_def split_box_conj [temp_use]) + apply (force simp: HistP_def) + apply (force simp: allT [temp_use] dest!: Step1_5_1a [temp_use]) + done + +(* The implementation is infinitely often in state S1... *) +lemma Step1_5_3a: "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) + & [](ALL l. $MemInv mm l) + & []($ImpInv rmhist p) & ImpLive p + --> []<>S1 rmhist p" + apply (clarsimp simp: ImpLive_def) + apply (rule S1Infinite) + apply (force simp: split_box_conj [try_rewrite] box_stp_act [try_rewrite] + intro!: NotS1LeadstoS6 [temp_use] S2_live [temp_use] S3_live [temp_use] + S4a_live [temp_use] S4b_live [temp_use] S5_live [temp_use]) + apply (auto simp: split_box_conj [temp_use] intro!: S6_live [temp_use]) + done + +(* ... and therefore satisfies the fairness requirements of the specification *) +lemma Step1_5_3b: "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) + & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p + --> WF(RNext memCh mm (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + by (auto intro!: RNext_fair [temp_use] Step1_5_3a [temp_use]) + +lemma Step1_5_3c: "|- [](ImpNext p & [HNext rmhist p]_(c p, r p, m p, rmhist!p)) + & [](ALL l. $MemInv mm l) & []($ImpInv rmhist p) & ImpLive p + --> WF(MemReturn memCh (resbar rmhist) p)_(rtrner memCh!p, resbar rmhist!p)" + by (auto intro!: Return_fair [temp_use] Step1_5_3a [temp_use]) + +(* QED step of step 1 *) +lemma Step1: "|- IPImp p & HistP rmhist p --> UPSpec memCh mm (resbar rmhist) p" + by (auto simp: UPSpec_def split_box_conj [temp_use] + dest!: GoodImpl [temp_use] intro!: Step1_5_2a [temp_use] Step1_5_2b [temp_use] + Step1_5_3b [temp_use] Step1_5_3c [temp_use]) + +(* ------------------------------ Step 2 ------------------------------ *) +section "Step 2" + +lemma Step2_2a: "|- 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)" + apply clarsimp + apply (drule WriteS4 [action_use]) + apply assumption + apply (tactic "split_idle_tac [] 1") + apply (auto simp: ImpNext_def dest!: S4EnvUnch [temp_use] S4ClerkUnch [temp_use] + S4RPCUnch [temp_use]) + apply (auto simp: square_def dest: S4Write [temp_use]) + done + +lemma Step2_2: "|- (ALL p. ImpNext p) + & (ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) + & (ALL p. $ImpInv rmhist p) + & [EX q. Write rmCh mm ires q l]_(mm!l) + --> [EX q. Write memCh mm (resbar rmhist) q l]_(mm!l)" + apply (auto intro!: squareCI elim!: squareE) + apply (assumption | rule exI Step1_4_4b [action_use])+ + apply (force intro!: WriteS4 [temp_use]) + apply (auto dest!: Step2_2a [temp_use]) + done + +lemma Step2_lemma: "|- []( (ALL p. ImpNext p) + & (ALL p. [HNext rmhist p]_(c p, r p, m p, rmhist!p)) + & (ALL p. $ImpInv rmhist p) + & [EX q. Write rmCh mm ires q l]_(mm!l)) + --> [][EX q. Write memCh mm (resbar rmhist) q l]_(mm!l)" + by (force elim!: STL4E [temp_use] dest!: Step2_2 [temp_use]) + +lemma Step2: "|- #l : #MemLoc & (ALL p. IPImp p & HistP rmhist p) + --> MSpec memCh mm (resbar rmhist) l" + apply (auto simp: MSpec_def) + apply (force simp: IPImp_def MSpec_def) + apply (auto intro!: Step2_lemma [temp_use] simp: split_box_conj [temp_use] all_box [temp_use]) + prefer 4 + apply (force simp: IPImp_def MSpec_def) + apply (auto simp: split_box_conj [temp_use] elim!: allE dest!: GoodImpl [temp_use]) + done + +(* ----------------------------- Main theorem --------------------------------- *) +section "Memory implementation" + +(* The combination of a legal caller, the memory clerk, the RPC component, + and a reliable memory implement the unreliable memory. +*) + +(* Implementation of internal specification by combination of implementation + and history variable with explicit refinement mapping +*) +lemma Impl_IUSpec: "|- Implementation & Hist rmhist --> IUSpec memCh mm (resbar rmhist)" + by (auto simp: IUSpec_def Implementation_def IPImp_def MClkISpec_def + RPCISpec_def IRSpec_def Hist_def intro!: Step1 [temp_use] Step2 [temp_use]) + +(* The main theorem: introduce hiding and eliminate history variable. *) +lemma Implementation: "|- Implementation --> USpec memCh" + apply clarsimp + apply (frule History [temp_use]) + apply (auto simp: USpec_def intro: eexI [temp_use] Impl_IUSpec [temp_use] + MI_base [temp_use] elim!: eexE) + done end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemoryParameters.ML --- a/src/HOL/TLA/Memory/MemoryParameters.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,19 +0,0 @@ -(* - File: MemoryParameters.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: memory parameters (ML file) -*) - -Addsimps ([BadArgNoMemVal,MemFailNoMemVal,InitValMemVal,NotAResultNotVal, - NotAResultNotOK, NotAResultNotBA, NotAResultNotMF] - @ (map (fn x => x RS not_sym) - [NotAResultNotOK, NotAResultNotBA, NotAResultNotMF])); - -val prems = goal (the_context ()) "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P"; -by (resolve_tac prems 1); -by (cut_facts_tac (NotAResultNotVal::prems) 1); -by (Force_tac 1); -qed "MemValNotAResultE"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/MemoryParameters.thy --- a/src/HOL/TLA/Memory/MemoryParameters.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/MemoryParameters.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: MemoryParameters - Logic Image: TLA - - RPC-Memory example: Memory parameters -*) +header {* RPC-Memory example: Memory parameters *} theory MemoryParameters imports RPCMemoryParams @@ -41,6 +38,12 @@ NotAResultNotBA: "NotAResult ~= BadArg" NotAResultNotMF: "NotAResult ~= MemFailure" -ML {* use_legacy_bindings (the_context ()) *} +lemmas [simp] = + BadArgNoMemVal MemFailNoMemVal InitValMemVal NotAResultNotVal + NotAResultNotOK NotAResultNotBA NotAResultNotMF + NotAResultNotOK [symmetric] NotAResultNotBA [symmetric] NotAResultNotMF [symmetric] + +lemma MemValNotAResultE: "[| x : MemVal; (x ~= NotAResult ==> P) |] ==> P" + using NotAResultNotVal by blast end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/ProcedureInterface.ML --- a/src/HOL/TLA/Memory/ProcedureInterface.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,28 +0,0 @@ -(* - File: ProcedureInterface.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - Procedure interface (theorems and proofs) -*) - -Addsimps [slice_def]; -val mem_css = (claset(), simpset()); - -(* ---------------------------------------------------------------------------- *) - -val Procedure_defs = [caller_def, rtrner_def, Calling_def, - Call_def, Return_def, - PLegalCaller_def, LegalCaller_def, - PLegalReturner_def, LegalReturner_def]; - -(* Calls and returns change their subchannel *) -Goal "|- Call ch p v --> _((caller ch)!p)"; -by (auto_tac (mem_css addsimps2 [angle_def,Call_def,caller_def,Calling_def])); -qed "Call_changed"; - -Goal "|- Return ch p v --> _((rtrner ch)!p)"; -by (auto_tac (mem_css addsimps2 [angle_def,Return_def,rtrner_def,Calling_def])); -qed "Return_changed"; - diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/ProcedureInterface.thy --- a/src/HOL/TLA/Memory/ProcedureInterface.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/ProcedureInterface.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: ProcedureInterface - Logic Image: TLA - - Procedure interface for RPC-Memory components. -*) +header {* Procedure interface for RPC-Memory components *} theory ProcedureInterface imports TLA RPCMemoryParams @@ -84,6 +81,16 @@ [][ ? v. Return ch p v ]_((rtrner ch)!p)" LegalReturner_def: "LegalReturner ch == TEMP (! p. PLegalReturner ch p)" -ML {* use_legacy_bindings (the_context ()) *} +declare slice_def [simp] + +lemmas Procedure_defs = caller_def rtrner_def Calling_def Call_def Return_def + PLegalCaller_def LegalCaller_def PLegalReturner_def LegalReturner_def + +(* Calls and returns change their subchannel *) +lemma Call_changed: "|- Call ch p v --> _((caller ch)!p)" + by (auto simp: angle_def Call_def caller_def Calling_def) + +lemma Return_changed: "|- Return ch p v --> _((rtrner ch)!p)" + by (auto simp: angle_def Return_def rtrner_def Calling_def) end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/RPC.ML --- a/src/HOL/TLA/Memory/RPC.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,53 +0,0 @@ -(* - File: RPC.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: RPC specification (theorems and proofs) -*) - -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. -*) - -Goal "|- ~$(Calling send p) --> ~RPCNext send rcv rst p"; -by (auto_tac (mem_css addsimps2 (Return_def::RPC_action_defs))); -qed "RPCidle"; - -Goal "|- $(Calling rcv p) & $(rst!p) = #rpcB --> ~RPCNext send rcv rst p"; -by (auto_tac (mem_css addsimps2 RPC_action_defs)); -qed "RPCbusy"; - -(* RPC failure actions are visible. *) -Goal "|- RPCFail send rcv rst p --> \ -\ _(rst!p, rtrner send!p, caller rcv!p)"; -by (auto_tac (claset() addSDs [Return_changed], - simpset() addsimps [angle_def,RPCNext_def,RPCFail_def])); -qed "RPCFail_vis"; - -Goal "|- Enabled (RPCFail send rcv rst p) --> \ -\ Enabled (_(rst!p, rtrner send!p, caller rcv!p))"; -by (force_tac (mem_css addSEs2 [enabled_mono,RPCFail_vis]) 1); -qed "RPCFail_Next_enabled"; - -(* Enabledness of some actions *) -Goal "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> \ -\ |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)"; -by (action_simp_tac (simpset() addsimps [RPCFail_def,Return_def,caller_def,rtrner_def]) - [exI] [base_enabled,Pair_inject] 1); -qed "RPCFail_enabled"; - -Goal "!!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)"; -by (action_simp_tac (simpset() addsimps [RPCReply_def,Return_def,caller_def,rtrner_def]) - [exI] [base_enabled,Pair_inject] 1); -qed "RPCReply_enabled"; diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/RPC.thy --- a/src/HOL/TLA/Memory/RPC.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/RPC.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: RPC - Logic Image: TLA - - RPC-Memory example: RPC specification -*) +header {* RPC-Memory example: RPC specification *} theory RPC imports RPCParameters ProcedureInterface Memory @@ -77,6 +74,44 @@ RPCISpec_def: "RPCISpec send rcv rst == TEMP (ALL p. RPCIPSpec send rcv rst p)" -ML {* use_legacy_bindings (the_context ()) *} + +lemmas RPC_action_defs = + RPCInit_def RPCFwd_def RPCReject_def RPCFail_def RPCReply_def RPCNext_def + +lemmas RPC_temp_defs = RPCIPSpec_def RPCISpec_def + + +(* The RPC component engages in an action for process p only if there is an outstanding, + unanswered call for that process. +*) + +lemma RPCidle: "|- ~$(Calling send p) --> ~RPCNext send rcv rst p" + by (auto simp: Return_def RPC_action_defs) + +lemma RPCbusy: "|- $(Calling rcv p) & $(rst!p) = #rpcB --> ~RPCNext send rcv rst p" + by (auto simp: RPC_action_defs) + +(* RPC failure actions are visible. *) +lemma RPCFail_vis: "|- RPCFail send rcv rst p --> + _(rst!p, rtrner send!p, caller rcv!p)" + by (auto dest!: Return_changed [temp_use] simp: angle_def RPCNext_def RPCFail_def) + +lemma RPCFail_Next_enabled: "|- Enabled (RPCFail send rcv rst p) --> + Enabled (_(rst!p, rtrner send!p, caller rcv!p))" + by (force elim!: enabled_mono [temp_use] RPCFail_vis [temp_use]) + +(* Enabledness of some actions *) +lemma RPCFail_enabled: "!!p. basevars (rtrner send!p, caller rcv!p, rst!p) ==> + |- ~Calling rcv p & Calling send p --> Enabled (RPCFail send rcv rst p)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "RPCFail_def", + thm "Return_def", thm "caller_def", thm "rtrner_def"]) [exI] + [thm "base_enabled", thm "Pair_inject"] 1 *}) + +lemma RPCReply_enabled: "!!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)" + by (tactic {* action_simp_tac (simpset () addsimps [thm "RPCReply_def", + thm "Return_def", thm "caller_def", thm "rtrner_def"]) [exI] + [thm "base_enabled", thm "Pair_inject"] 1 *}) end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/RPCMemoryParams.thy --- a/src/HOL/TLA/Memory/RPCMemoryParams.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/RPCMemoryParams.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich +*) - Theory Name: RPCMemoryParams - Logic Image: TLA - - Basic declarations for the RPC-memory example. -*) +header {* Basic declarations for the RPC-memory example *} theory RPCMemoryParams imports Main diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/RPCParameters.ML --- a/src/HOL/TLA/Memory/RPCParameters.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,11 +0,0 @@ -(* - File: RPCParameters.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1997 University of Munich - - RPC-Memory example: RPC parameters (theorems and proofs) -*) - -Addsimps ([RFNoMemVal, NotAResultNotRF, OKNotRF, BANotRF] - @ (map (fn x => x RS not_sym) [NotAResultNotRF, OKNotRF, BANotRF])); diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Memory/RPCParameters.thy --- a/src/HOL/TLA/Memory/RPCParameters.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Memory/RPCParameters.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,19 +3,19 @@ ID: $Id$ Author: Stephan Merz Copyright: 1997 University of Munich - - Theory Name: RPCParameters - Logic Image: TLA +*) - RPC-Memory example: RPC parameters - For simplicity, specify the instance of RPC that is used in the - memory implementation. -*) +header {* RPC-Memory example: RPC parameters *} theory RPCParameters imports MemoryParameters begin +(* + For simplicity, specify the instance of RPC that is used in the + memory implementation. +*) + datatype rpcOp = memcall memOp | othercall Vals datatype rpcState = rpcA | rpcB @@ -46,6 +46,7 @@ case ra of (memcall m) => m | (othercall v) => arbitrary" -ML {* use_legacy_bindings (the_context ()) *} +lemmas [simp] = RFNoMemVal NotAResultNotRF OKNotRF BANotRF + NotAResultNotRF [symmetric] OKNotRF [symmetric] BANotRF [symmetric] end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Stfun.ML --- a/src/HOL/TLA/Stfun.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,59 +0,0 @@ -(* - File: Stfun.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1998 University of Munich - -Lemmas and tactics for states and state functions. -*) - -Goalw [basevars_def] "!!vs. basevars vs ==> EX u. vs u = c"; -by (res_inst_tac [("b","c"),("f","vs")] rangeE 1); -by Auto_tac; -qed "basevars"; - -Goal "!!x y. basevars (x,y) ==> basevars x"; -by (simp_tac (simpset() addsimps [basevars_def]) 1); -by (rtac equalityI 1); - by (rtac subset_UNIV 1); -by (rtac subsetI 1); -by (dres_inst_tac [("c", "(xa, arbitrary)")] basevars 1); -by Auto_tac; -qed "base_pair1"; - -Goal "!!x y. basevars (x,y) ==> basevars y"; -by (simp_tac (simpset() addsimps [basevars_def]) 1); -by (rtac equalityI 1); - by (rtac subset_UNIV 1); -by (rtac subsetI 1); -by (dres_inst_tac [("c", "(arbitrary, xa)")] basevars 1); -by Auto_tac; -qed "base_pair2"; - -Goal "!!x y. basevars (x,y) ==> basevars x & basevars y"; -by (rtac conjI 1); -by (etac base_pair1 1); -by (etac base_pair2 1); -qed "base_pair"; - -(* 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. -*) -Goalw [basevars_def] "basevars (v::unit stfun)"; -by Auto_tac; -qed "unit_base"; - -(* [| basevars v; !!x. v x = c ==> Q |] ==> Q *) -bind_thm("baseE", (standard (basevars RS exE))); - -(* ------------------------------------------------------------------------------- - The following shows that there should not be duplicates in a "stvars" tuple: - -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); - -------------------------------------------------------------------------------- *) diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/Stfun.thy --- a/src/HOL/TLA/Stfun.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/Stfun.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1998 University of Munich +*) - Theory Name: Stfun - Logic Image: HOL - -States and state functions for TLA as an "intensional" logic. -*) +header {* States and state functions for TLA as an "intensional" logic *} theory Stfun imports Intensional @@ -56,6 +53,62 @@ *) basevars_def: "stvars vs == range vs = UNIV" -ML {* use_legacy_bindings (the_context ()) *} + +lemma basevars: "!!vs. basevars vs ==> EX u. vs u = c" + apply (unfold basevars_def) + apply (rule_tac b = c and f = vs in rangeE) + apply auto + done + +lemma base_pair1: "!!x y. basevars (x,y) ==> basevars x" + apply (simp (no_asm) add: basevars_def) + apply (rule equalityI) + apply (rule subset_UNIV) + apply (rule subsetI) + apply (drule_tac c = "(xa, arbitrary) " in basevars) + apply auto + done + +lemma base_pair2: "!!x y. basevars (x,y) ==> basevars y" + apply (simp (no_asm) add: basevars_def) + apply (rule equalityI) + apply (rule subset_UNIV) + apply (rule subsetI) + apply (drule_tac c = "(arbitrary, xa) " in basevars) + apply auto + done + +lemma base_pair: "!!x y. basevars (x,y) ==> basevars x & basevars y" + apply (rule conjI) + apply (erule base_pair1) + apply (erule base_pair2) + done + +(* 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. +*) +lemma unit_base: "basevars (v::unit stfun)" + apply (unfold basevars_def) + apply auto + done + +lemma baseE: "[| basevars v; !!x. v x = c ==> Q |] ==> Q" + apply (erule basevars [THEN exE]) + apply blast + done + + +(* ------------------------------------------------------------------------------- + The following shows that there should not be duplicates in a "stvars" tuple: +*) + +lemma "!!v. basevars (v::bool stfun, v) ==> False" + apply (erule baseE) + apply (subgoal_tac "(LIFT (v,v)) x = (True, False)") + prefer 2 + apply assumption + apply simp + done end diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/TLA.ML --- a/src/HOL/TLA/TLA.ML Fri Dec 01 17:22:33 2006 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1069 +0,0 @@ -(* - File: TLA/TLA.ML - ID: $Id$ - Author: Stephan Merz - Copyright: 1998 University of Munich - -Lemmas and tactics for temporal reasoning. -*) - -(* Specialize intensional introduction/elimination rules for temporal formulas *) - -val [prem] = goal (the_context ()) "(!!sigma. sigma |= (F::temporal)) ==> |- F"; -by (REPEAT (resolve_tac [prem,intI] 1)); -qed "tempI"; - -val [prem] = goal (the_context ()) "|- (F::temporal) ==> sigma |= F"; -by (rtac (prem RS intD) 1); -qed "tempD"; - - -(* ======== Functions to "unlift" temporal theorems ====== *) - -(* 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_unlift th = - (rewrite_rule action_rews (th RS tempD)) - handle _ => action_unlift th; - -(* Turn |- F = G into meta-level rewrite rule F == G *) -val temp_rewrite = int_rewrite; - -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 = (claset(), simpset()); -val temp_cs = op addss temp_css; - -(* Modify the functions that add rules to simpsets, classical sets, - and clasimpsets in order to accept "lifted" theorems -*) - -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; - -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); - -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); - - -(* ------------------------------------------------------------------------- *) -(*** "Simple temporal logic": only [] and <> ***) -(* ------------------------------------------------------------------------- *) -section "Simple temporal logic"; - -(* []~F == []~Init F *) -bind_thm("boxNotInit", - rewrite_rule Init_simps (read_instantiate [("F", "LIFT ~F")] boxInit)); - -Goalw [dmd_def] "TEMP <>F == TEMP <> Init F"; -by (rewtac (read_instantiate [("F", "LIFT ~F")] boxInit)); -by (simp_tac (simpset() addsimps Init_simps) 1); -qed "dmdInit"; - -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); - -(* The "polymorphic" (generic) variant *) -Goal "|- []F --> Init F"; -by (rewtac (read_instantiate [("F", "F")] boxInit)); -by (rtac STL2 1); -qed "STL2_gen"; - -(* see also STL2_pr below: "|- []P --> Init P & Init (P`)" *) - - -(* Dual versions for <> *) -Goalw [dmd_def] "|- F --> <> F"; -by (auto_tac (temp_css addSDs2 [STL2])); -qed "InitDmd"; - -Goal "|- Init F --> <>F"; -by (Clarsimp_tac 1); -by (dtac (temp_use InitDmd) 1); -by (asm_full_simp_tac (simpset() addsimps [dmdInitD]) 1); -qed "InitDmd_gen"; - - -(* ------------------------ STL3 ------------------------------------------- *) -Goal "|- ([][]F) = ([]F)"; -by (force_tac (temp_css addEs2 [transT,STL2]) 1); -qed "STL3"; - -(* corresponding elimination rule introduces double boxes: - [| (sigma |= []F); (sigma |= [][]F) ==> PROP W |] ==> PROP W -*) -bind_thm("dup_boxE", make_elim((temp_unlift STL3) RS iffD2)); -bind_thm("dup_boxD", (temp_unlift STL3) RS iffD1); - -(* dual versions for <> *) -Goal "|- (<><>F) = (<>F)"; -by (auto_tac (temp_css addsimps2 [dmd_def,STL3])); -qed "DmdDmd"; -bind_thm("dup_dmdE", make_elim((temp_unlift DmdDmd) RS iffD2)); -bind_thm("dup_dmdD", (temp_unlift DmdDmd) RS iffD1); - - -(* ------------------------ STL4 ------------------------------------------- *) -val [prem] = goal (the_context ()) "|- F --> G ==> |- []F --> []G"; -by (Clarsimp_tac 1); -by (rtac (temp_use normalT) 1); -by (rtac (temp_use (prem RS necT)) 1); -by (atac 1); -qed "STL4"; - -(* Unlifted version as an elimination rule *) -val prems = goal (the_context ()) "[| sigma |= []F; |- F --> G |] ==> sigma |= []G"; -by (REPEAT (resolve_tac (prems @ [temp_use STL4]) 1)); -qed "STL4E"; - -val [prem] = goal (the_context ()) "|- Init F --> Init G ==> |- []F --> []G"; -by (rtac (rewrite_rule [boxInitD] (prem RS STL4)) 1); -qed "STL4_gen"; - -val prems = goal (the_context ()) - "[| sigma |= []F; |- Init F --> Init G |] ==> sigma |= []G"; -by (REPEAT (resolve_tac (prems @ [temp_use STL4_gen]) 1)); -qed "STL4E_gen"; - -(* see also STL4Edup below, which allows an auxiliary boxed formula: - []A /\ F => G - ----------------- - []A /\ []F => []G -*) - -(* The dual versions for <> *) -val [prem] = goalw (the_context ()) [dmd_def] - "|- F --> G ==> |- <>F --> <>G"; -by (fast_tac (temp_cs addSIs [prem] addSEs [STL4E]) 1); -qed "DmdImpl"; - -val prems = goal (the_context ()) "[| sigma |= <>F; |- F --> G |] ==> sigma |= <>G"; -by (REPEAT (resolve_tac (prems @ [temp_use DmdImpl]) 1)); -qed "DmdImplE"; - - -(* ------------------------ STL5 ------------------------------------------- *) -Goal "|- ([]F & []G) = ([](F & G))"; -by Auto_tac; -by (subgoal_tac "sigma |= [](G --> (F & G))" 1); -by (etac (temp_use normalT) 1); -by (ALLGOALS (fast_tac (temp_cs addSEs [STL4E]))); -qed "STL5"; - -(* rewrite rule to split conjunctions under boxes *) -bind_thm("split_box_conj", (temp_unlift STL5) RS sym); - -(* 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! -*) -val prems = goal (the_context ()) - "[| sigma |= []F; sigma |= []G; sigma |= [](F&G) ==> PROP R |] ==> PROP R"; -by (REPEAT (resolve_tac (prems @ [(temp_unlift STL5) RS iffD1, conjI]) 1)); -qed "box_conjE"; - -(* 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 in order to simulate "double elim-resolution". -*) - -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]); - -fun merge_temp_box_tac i = - REPEAT_DETERM (EVERY [etac box_conjE_temp i, atac i, - eres_inst_tac [("'a","behavior")] box_thin i]); - -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 box_conjE_act i, atac i, - eres_inst_tac [("'a","state * state")] box_thin i]); - - -(* rewrite rule to push universal quantification through box: - (sigma |= [](! x. F x)) = (! x. (sigma |= []F x)) -*) -bind_thm("all_box", standard((temp_unlift allT) RS sym)); - -bind_thm ("contrapos_np", thm "contrapos_np"); - -Goal "|- (<>(F | G)) = (<>F | <>G)"; -by (auto_tac (temp_css addsimps2 [dmd_def,split_box_conj])); -by (ALLGOALS (EVERY' [etac contrapos_np, - merge_box_tac, - fast_tac (temp_cs addSEs [STL4E])])); -qed "DmdOr"; - -Goal "|- (EX x. <>(F x)) = (<>(EX x. F x))"; -by (auto_tac (temp_css addsimps2 [dmd_def,Not_Rex,all_box])); -qed "exT"; - -bind_thm("ex_dmd", standard((temp_unlift exT) RS sym)); - - -Goal "!!sigma. [| sigma |= []A; sigma |= []F; |- F & []A --> G |] ==> sigma |= []G"; -by (etac dup_boxE 1); -by (merge_box_tac 1); -by (etac STL4E 1); -by (atac 1); -qed "STL4Edup"; - -Goalw [dmd_def] - "!!sigma. [| sigma |= <>F; sigma |= [](F --> G) |] ==> sigma |= <>G"; -by Auto_tac; -by (etac notE 1); -by (merge_box_tac 1); -by (fast_tac (temp_cs addSEs [STL4E]) 1); -qed "DmdImpl2"; - -val [prem1,prem2,prem3] = goal (the_context ()) - "[| sigma |= []<>F; sigma |= []G; |- F & G --> H |] ==> sigma |= []<>H"; -by (cut_facts_tac [prem1,prem2] 1); -by (eres_inst_tac [("F","G")] dup_boxE 1); -by (merge_box_tac 1); -by (fast_tac (temp_cs addSEs [STL4E,DmdImpl2] addSIs [prem3]) 1); -qed "InfImpl"; - -(* ------------------------ STL6 ------------------------------------------- *) -(* Used in the proof of STL6, but useful in itself. *) -Goalw [dmd_def] "|- []F & <>G --> <>([]F & G)"; -by (Clarsimp_tac 1); -by (etac dup_boxE 1); -by (merge_box_tac 1); -by (etac contrapos_np 1); -by (fast_tac (temp_cs addSEs [STL4E]) 1); -qed "BoxDmd"; - -(* weaker than BoxDmd, but more polymorphic (and often just right) *) -Goalw [dmd_def] "|- []F & <>G --> <>(F & G)"; -by (Clarsimp_tac 1); -by (merge_box_tac 1); -by (fast_tac (temp_cs addSEs [notE,STL4E]) 1); -qed "BoxDmd_simple"; - -Goalw [dmd_def] "|- []F & <>G --> <>(G & F)"; -by (Clarsimp_tac 1); -by (merge_box_tac 1); -by (fast_tac (temp_cs addSEs [notE,STL4E]) 1); -qed "BoxDmd2_simple"; - -val [p1,p2,p3] = goal (the_context ()) - "[| sigma |= []A; sigma |= <>F; |- []A & F --> G |] ==> sigma |= <>G"; -by (rtac ((p2 RS (p1 RS (temp_use BoxDmd))) RS DmdImplE) 1); -by (rtac p3 1); -qed "DmdImpldup"; - -Goal "|- <>[]F & <>[]G --> <>[](F & G)"; -by (auto_tac (temp_css addsimps2 [symmetric (temp_rewrite STL5)])); -by (dtac (temp_use linT) 1); -by (atac 1); -by (etac thin_rl 1); -by (rtac ((temp_unlift DmdDmd) RS iffD1) 1); -by (etac disjE 1); -by (etac DmdImplE 1); -by (rtac BoxDmd 1); -by (etac DmdImplE 1); -by Auto_tac; -by (dtac (temp_use BoxDmd) 1); -by (atac 1); -by (etac thin_rl 1); -by (fast_tac (temp_cs addSEs [DmdImplE]) 1); -qed "STL6"; - - -(* ------------------------ True / False ----------------------------------------- *) -section "Simplification of constants"; - -Goal "|- ([]#P) = #P"; -by (rtac tempI 1); -by (case_tac "P" 1); -by (auto_tac (temp_css addSIs2 [necT] addDs2 [STL2_gen] - addsimps2 Init_simps)); -qed "BoxConst"; - -Goalw [dmd_def] "|- (<>#P) = #P"; -by (case_tac "P" 1); -by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [BoxConst]))); -qed "DmdConst"; - -val temp_simps = map temp_rewrite [BoxConst, DmdConst]; - -(* Make these rewrites active by default *) -Addsimps temp_simps; -val temp_css = temp_css addsimps2 temp_simps; -val temp_cs = op addss temp_css; - - -(* ------------------------ Further rewrites ----------------------------------------- *) -section "Further rewrites"; - -Goalw [dmd_def] "|- (~[]F) = (<>~F)"; -by (Simp_tac 1); -qed "NotBox"; - -Goalw [dmd_def] "|- (~<>F) = ([]~F)"; -by (Simp_tac 1); -qed "NotDmd"; - -(* These are not by default included in temp_css, because they could be harmful, - 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]); - -Goal "|- ([]<>[]F) = (<>[]F)"; -by (auto_tac (temp_css addSDs2 [STL2])); -by (rtac ccontr 1); -by (subgoal_tac "sigma |= <>[][]F & <>[]~[]F" 1); -by (etac thin_rl 1); -by Auto_tac; -by (dtac (temp_use STL6) 1); -by (atac 1); -by (Asm_full_simp_tac 1); -by (ALLGOALS (asm_full_simp_tac (simpset() addsimps more_temp_simps))); -qed "BoxDmdBox"; - -Goalw [dmd_def] "|- (<>[]<>F) = ([]<>F)"; -by (auto_tac (temp_css addsimps2 [rewrite_rule [dmd_def] BoxDmdBox])); -qed "DmdBoxDmd"; - -val more_temp_simps = more_temp_simps @ (map temp_rewrite [BoxDmdBox, DmdBoxDmd]); - - -(* ------------------------ Miscellaneous ----------------------------------- *) - -Goal "!!sigma. [| sigma |= []F | []G |] ==> sigma |= [](F | G)"; -by (fast_tac (temp_cs addSEs [STL4E]) 1); -qed "BoxOr"; - -(* "persistently implies infinitely often" *) -Goal "|- <>[]F --> []<>F"; -by (Clarsimp_tac 1); -by (rtac ccontr 1); -by (asm_full_simp_tac (simpset() addsimps more_temp_simps) 1); -by (dtac (temp_use STL6) 1); -by (atac 1); -by (Asm_full_simp_tac 1); -qed "DBImplBD"; - -Goal "|- []<>F & <>[]G --> []<>(F & G)"; -by (Clarsimp_tac 1); -by (rtac ccontr 1); -by (rewrite_goals_tac more_temp_simps); -by (dtac (temp_use STL6) 1); -by (atac 1); -by (subgoal_tac "sigma |= <>[]~F" 1); - by (force_tac (temp_css addsimps2 [dmd_def]) 1); -by (fast_tac (temp_cs addEs [DmdImplE,STL4E]) 1); -qed "BoxDmdDmdBox"; - - -(* ------------------------------------------------------------------------- *) -(*** TLA-specific theorems: primed formulas ***) -(* ------------------------------------------------------------------------- *) -section "priming"; - -(* ------------------------ TLA2 ------------------------------------------- *) -Goal "|- []P --> Init P & Init P`"; -by (fast_tac (temp_cs addSIs [primeI, STL2_gen]) 1); -qed "STL2_pr"; - -(* Auxiliary lemma allows priming of boxed actions *) -Goal "|- []P --> []($P & P$)"; -by (Clarsimp_tac 1); -by (etac dup_boxE 1); -by (rewtac boxInit_act); -by (etac STL4E 1); -by (auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_pr])); -qed "BoxPrime"; - -val prems = goal (the_context ()) "|- $P & P$ --> A ==> |- []P --> []A"; -by (Clarsimp_tac 1); -by (dtac (temp_use BoxPrime) 1); -by (auto_tac (temp_css addsimps2 [Init_stp_act_rev] - addSIs2 prems addSEs2 [STL4E])); -qed "TLA2"; - -val prems = goal (the_context ()) - "[| sigma |= []P; |- $P & P$ --> A |] ==> sigma |= []A"; -by (REPEAT (resolve_tac (prems @ (prems RL [temp_use TLA2])) 1)); -qed "TLA2E"; - -Goalw [dmd_def] "|- (<>P`) --> (<>P)"; -by (fast_tac (temp_cs addSEs [TLA2E]) 1); -qed "DmdPrime"; - -bind_thm("PrimeDmd", (temp_use InitDmd_gen) RS (temp_use DmdPrime)); - -(* ------------------------ INV1, stable --------------------------------------- *) -section "stable, invariant"; - -val prems = goal (the_context ()) - "[| sigma |= []H; sigma |= Init P; |- H --> (Init P & ~[]F --> Init(P`) & F) |] \ -\ ==> sigma |= []F"; -by (rtac (temp_use indT) 1); -by (REPEAT (resolve_tac (prems @ (prems RL [STL4E])) 1)); -qed "ind_rule"; - -Goalw [boxInit_act] "|- ([]$P) = ([]P)"; -by (simp_tac (simpset() addsimps Init_simps) 1); -qed "box_stp_act"; -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)); - -val more_temp_simps = (temp_rewrite box_stp_act)::more_temp_simps; - -Goalw [stable_def,boxInit_stp,boxInit_act] - "|- (Init P) --> (stable P) --> []P"; -by (Clarsimp_tac 1); -by (etac ind_rule 1); -by (auto_tac (temp_css addsimps2 Init_simps addEs2 [ind_rule])); -qed "INV1"; - -Goalw [stable_def] - "!!P. |- $P & A --> P` ==> |- []A --> stable P"; -by (fast_tac (temp_cs addSEs [STL4E]) 1); -qed "StableT"; - -val prems = goal (the_context ()) - "[| sigma |= []A; |- $P & A --> P` |] ==> sigma |= stable P"; -by (REPEAT (resolve_tac (prems @ [temp_use StableT]) 1)); -qed "Stable"; - -(* Generalization of INV1 *) -Goalw [stable_def] "|- (stable P) --> [](Init P --> []P)"; -by (Clarsimp_tac 1); -by (etac dup_boxE 1); -by (force_tac (temp_css addsimps2 [stable_def] addEs2 [STL4E, INV1]) 1); -qed "StableBox"; - -Goal "|- (stable P) & <>P --> <>[]P"; -by (Clarsimp_tac 1); -by (rtac DmdImpl2 1); -by (etac (temp_use StableBox) 2); -by (asm_simp_tac (simpset() addsimps [dmdInitD]) 1); -qed "DmdStable"; - -(* ---------------- (Semi-)automatic invariant tactics ---------------------- *) - -(* inv_tac reduces goals of the form ... ==> sigma |= []P *) -fun inv_tac css = SELECT_GOAL - (EVERY [auto_tac css, - TRY (merge_box_tac 1), - 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 these simple cases the simplifier seems to be more useful than the - auto-tactic, which applies too much propositional logic and simplifies - too late. -*) - -fun auto_inv_tac ss = SELECT_GOAL - ((inv_tac (claset(),ss) 1) THEN - (TRYALL (action_simp_tac (ss addsimps [Init_stp,Init_act]) [] [squareE]))); - - -Goalw [dmd_def] "|- []($P --> P` | Q`) --> (stable P) | <>Q"; -by (clarsimp_tac (temp_css addSDs2 [BoxPrime]) 1); -by (merge_box_tac 1); -by (etac contrapos_np 1); -by (fast_tac (temp_cs addSEs [Stable]) 1); -qed "unless"; - - -(* --------------------- Recursive expansions --------------------------------------- *) -section "recursive expansions"; - -(* Recursive expansions of [] and <> for state predicates *) -Goal "|- ([]P) = (Init P & []P`)"; -by (auto_tac (temp_css addSIs2 [STL2_gen])); -by (fast_tac (temp_cs addSEs [TLA2E]) 1); -by (auto_tac (temp_css addsimps2 [stable_def] addSEs2 [INV1,STL4E])); -qed "BoxRec"; - -Goalw [dmd_def, temp_rewrite BoxRec] "|- (<>P) = (Init P | <>P`)"; -by (auto_tac (temp_css addsimps2 Init_simps)); -qed "DmdRec"; - -Goal "!!sigma. [| sigma |= <>P; sigma |= []~P` |] ==> sigma |= Init P"; -by (force_tac (temp_css addsimps2 [DmdRec,dmd_def]) 1); -qed "DmdRec2"; - -Goal "|- ([]<>P) = ([]<>P`)"; -by Auto_tac; -by (rtac classical 1); -by (rtac (temp_use DBImplBD) 1); -by (subgoal_tac "sigma |= <>[]P" 1); - by (fast_tac (temp_cs addSEs [DmdImplE,TLA2E]) 1); - by (subgoal_tac "sigma |= <>[](<>P & []~P`)" 1); - by (force_tac (temp_css addsimps2 [boxInit_stp] - addSEs2 [DmdImplE,STL4E,DmdRec2]) 1); - by (force_tac (temp_css addSIs2 [STL6] addsimps2 more_temp_simps) 1); -by (fast_tac (temp_cs addIs [DmdPrime] addSEs [STL4E]) 1); -qed "InfinitePrime"; - -val prems = goalw (the_context ()) [temp_rewrite InfinitePrime] - "[| sigma |= []N; sigma |= []<>A; |- A & N --> P` |] ==> sigma |= []<>P"; -by (rtac InfImpl 1); -by (REPEAT (resolve_tac prems 1)); -qed "InfiniteEnsures"; - -(* ------------------------ fairness ------------------------------------------- *) -section "fairness"; - -(* alternative definitions of fairness *) -Goalw [WF_def,dmd_def] - "|- WF(A)_v = ([]<>~Enabled(_v) | []<>_v)"; -by (fast_tac temp_cs 1); -qed "WF_alt"; - -Goalw [SF_def,dmd_def] - "|- SF(A)_v = (<>[]~Enabled(_v) | []<>_v)"; -by (fast_tac temp_cs 1); -qed "SF_alt"; - -(* theorems to "box" fairness conditions *) -Goal "|- WF(A)_v --> []WF(A)_v"; -by (auto_tac (temp_css addsimps2 (WF_alt::more_temp_simps) - addSIs2 [BoxOr])); -qed "BoxWFI"; - -Goal "|- ([]WF(A)_v) = WF(A)_v"; -by (fast_tac (temp_cs addSIs [BoxWFI] addSDs [STL2]) 1); -qed "WF_Box"; - -Goal "|- SF(A)_v --> []SF(A)_v"; -by (auto_tac (temp_css addsimps2 (SF_alt::more_temp_simps) - addSIs2 [BoxOr])); -qed "BoxSFI"; - -Goal "|- ([]SF(A)_v) = SF(A)_v"; -by (fast_tac (temp_cs addSIs [BoxSFI] addSDs [STL2]) 1); -qed "SF_Box"; - -val more_temp_simps = more_temp_simps @ (map temp_rewrite [WF_Box, SF_Box]); - -Goalw [SF_def,WF_def] "|- SF(A)_v --> WF(A)_v"; -by (fast_tac (temp_cs addSDs [DBImplBD]) 1); -qed "SFImplWF"; - -(* 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)); - - -(* ------------------------------ leads-to ------------------------------ *) - -section "~>"; - -Goalw [leadsto_def] "|- (Init F) & (F ~> G) --> <>G"; -by (auto_tac (temp_css addSDs2 [STL2])); -qed "leadsto_init"; - -(* |- F & (F ~> G) --> <>G *) -bind_thm("leadsto_init_temp", - rewrite_rule Init_simps (read_instantiate [("'a","behavior")] leadsto_init)); - -Goalw [leadsto_def] "|- ([]<>Init F --> []<>G) = (<>(F ~> G))"; -by Auto_tac; -by (asm_full_simp_tac (simpset() addsimps more_temp_simps) 1); -by (fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1); -by (fast_tac (temp_cs addSIs [InitDmd] addSEs [STL4E]) 1); -by (subgoal_tac "sigma |= []<><>G" 1); -by (asm_full_simp_tac (simpset() addsimps more_temp_simps) 1); -by (dtac (temp_use BoxDmdDmdBox) 1); -by (atac 1); -by (fast_tac (temp_cs addSEs [DmdImplE,STL4E]) 1); -qed "streett_leadsto"; - -Goal "|- []<>F & (F ~> G) --> []<>G"; -by (Clarsimp_tac 1); -by (etac ((temp_use InitDmd) RS - ((temp_unlift streett_leadsto) RS iffD2 RS mp)) 1); -by (asm_simp_tac (simpset() addsimps [dmdInitD]) 1); -qed "leadsto_infinite"; - -(* In particular, strong fairness is a Streett condition. The following - rules are sometimes easier to use than WF2 or SF2 below. -*) -Goalw [SF_def] "|- (Enabled(_v) ~> _v) --> SF(A)_v"; -by (clarsimp_tac (temp_css addSEs2 [leadsto_infinite]) 1); -qed "leadsto_SF"; - -Goal "|- (Enabled(_v) ~> _v) --> WF(A)_v"; -by (clarsimp_tac (temp_css addSIs2 [SFImplWF, leadsto_SF]) 1); -qed "leadsto_WF"; - -(* introduce an invariant into the proof of a leadsto assertion. - []I --> ((P ~> Q) = (P /\ I ~> Q)) -*) -Goalw [leadsto_def] "|- []I & (P & I ~> Q) --> (P ~> Q)"; -by (Clarsimp_tac 1); -by (etac STL4Edup 1); -by (atac 1); -by (auto_tac (temp_css addsimps2 Init_simps addSDs2 [STL2_gen])); -qed "INV_leadsto"; - -Goalw [leadsto_def,dmd_def] - "|- (Init F & []~G ~> G) --> (F ~> G)"; -by (force_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E]) 1); -qed "leadsto_classical"; - -Goalw [leadsto_def] "|- (F ~> #False) = ([]~F)"; -by (simp_tac (simpset() addsimps [boxNotInitD]) 1); -qed "leadsto_false"; - -Goalw [leadsto_def] "|- ((EX x. F x) ~> G) = (ALL x. (F x ~> G))"; -by (auto_tac (temp_css addsimps2 allT::Init_simps addSEs2 [STL4E])); -qed "leadsto_exists"; - -(* basic leadsto properties, cf. Unity *) - -Goalw [leadsto_def] "|- [](Init F --> Init G) --> (F ~> G)"; -by (auto_tac (temp_css addSIs2 [InitDmd_gen] addSEs2 [STL4E_gen] - addsimps2 Init_simps)); -qed "ImplLeadsto_gen"; - -bind_thm("ImplLeadsto", - rewrite_rule Init_simps - (read_instantiate [("'a","behavior"), ("'b","behavior")] ImplLeadsto_gen)); - -Goal "!!F G. |- F --> G ==> |- F ~> G"; -by (auto_tac (temp_css addsimps2 [Init_def] - addSIs2 [ImplLeadsto_gen,necT])); -qed "ImplLeadsto_simple"; - -val [prem] = goalw (the_context ()) [leadsto_def] - "|- A & $P --> Q` ==> |- []A --> (P ~> Q)"; -by (clarsimp_tac (temp_css addSEs2 [INV_leadsto]) 1); -by (etac STL4E_gen 1); -by (auto_tac (temp_css addsimps2 Init_defs addSIs2 [PrimeDmd,prem])); -qed "EnsuresLeadsto"; - -Goalw [leadsto_def] "|- []($P --> Q`) --> (P ~> Q)"; -by (Clarsimp_tac 1); -by (etac STL4E_gen 1); -by (auto_tac (temp_css addsimps2 Init_simps addSIs2 [PrimeDmd])); -qed "EnsuresLeadsto2"; - -val [p1,p2] = goalw (the_context ()) [leadsto_def] - "[| |- $P & N --> P` | Q`; \ -\ |- ($P & N) & A --> Q` \ -\ |] ==> |- []N & []([]P --> <>A) --> (P ~> Q)"; -by (Clarsimp_tac 1); -by (etac STL4Edup 1); -by (atac 1); -by (Clarsimp_tac 1); -by (subgoal_tac "sigmaa |= []($P --> P` | Q`)" 1); - by (dtac (temp_use unless) 1); - by (clarsimp_tac (temp_css addSDs2 [INV1]) 1); - by (rtac ((temp_use (p2 RS DmdImpl)) RS (temp_use DmdPrime)) 1); - by (force_tac (temp_css addSIs2 [BoxDmd_simple] - addsimps2 [split_box_conj,box_stp_act]) 1); -by (force_tac (temp_css addEs2 [STL4E] addDs2 [p1]) 1); -qed "ensures"; - -val prems = goal (the_context ()) - "[| |- $P & N --> P` | Q`; \ -\ |- ($P & N) & A --> Q` \ -\ |] ==> |- []N & []<>A --> (P ~> Q)"; -by (Clarsimp_tac 1); -by (rtac (temp_use ensures) 1); -by (TRYALL (ares_tac prems)); -by (force_tac (temp_css addSEs2 [STL4E]) 1); -qed "ensures_simple"; - -val prems = goal (the_context ()) - "[| sigma |= []<>P; sigma |= []A; |- A & $P --> Q` |] ==> sigma |= []<>Q"; -by (REPEAT (resolve_tac (prems @ - (map temp_use [leadsto_infinite, EnsuresLeadsto])) 1)); -qed "EnsuresInfinite"; - - -(*** Gronning's lattice rules (taken from TLP) ***) -section "Lattice rules"; - -Goalw [leadsto_def] "|- F ~> F"; -by (REPEAT (resolve_tac [necT,InitDmd_gen] 1)); -qed "LatticeReflexivity"; - -Goalw [leadsto_def] "|- (G ~> H) & (F ~> G) --> (F ~> H)"; -by (Clarsimp_tac 1); -by (etac dup_boxE 1); (* [][](Init G --> H) *) -by (merge_box_tac 1); -by (clarsimp_tac (temp_css addSEs2 [STL4E]) 1); -by (rtac dup_dmdD 1); -by (subgoal_tac "sigmaa |= <>Init G" 1); - by (etac DmdImpl2 1); - by (atac 1); -by (asm_simp_tac (simpset() addsimps [dmdInitD]) 1); -qed "LatticeTransitivity"; - -Goalw [leadsto_def] "|- (F | G ~> H) --> (F ~> H)"; -by (auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E])); -qed "LatticeDisjunctionElim1"; - -Goalw [leadsto_def] "|- (F | G ~> H) --> (G ~> H)"; -by (auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E])); -qed "LatticeDisjunctionElim2"; - -Goalw [leadsto_def] "|- (F ~> H) & (G ~> H) --> (F | G ~> H)"; -by (Clarsimp_tac 1); -by (merge_box_tac 1); -by (auto_tac (temp_css addsimps2 Init_simps addSEs2 [STL4E])); -qed "LatticeDisjunctionIntro"; - -Goal "|- (F | G ~> H) = ((F ~> H) & (G ~> H))"; -by (auto_tac (temp_css addIs2 [LatticeDisjunctionIntro, - LatticeDisjunctionElim1, LatticeDisjunctionElim2])); -qed "LatticeDisjunction"; - -Goal "|- (A ~> B | C) & (B ~> D) & (C ~> D) --> (A ~> D)"; -by (Clarsimp_tac 1); -by (subgoal_tac "sigma |= (B | C) ~> D" 1); -by (eres_inst_tac [("G", "LIFT (B | C)")] (temp_use LatticeTransitivity) 1); -by (ALLGOALS (fast_tac (temp_cs addSIs [LatticeDisjunctionIntro]))); -qed "LatticeDiamond"; - -Goal "|- (A ~> D | B) & (B ~> D) --> (A ~> D)"; -by (Clarsimp_tac 1); -by (subgoal_tac "sigma |= (D | B) ~> D" 1); -by (eres_inst_tac [("G", "LIFT (D | B)")] (temp_use LatticeTransitivity) 1); -by (atac 1); -by (auto_tac (temp_css addIs2 [LatticeDisjunctionIntro,LatticeReflexivity])); -qed "LatticeTriangle"; - -Goal "|- (A ~> B | D) & (B ~> D) --> (A ~> D)"; -by (Clarsimp_tac 1); -by (subgoal_tac "sigma |= B | D ~> D" 1); -by (eres_inst_tac [("G", "LIFT (B | D)")] (temp_use LatticeTransitivity) 1); -by (atac 1); -by (auto_tac (temp_css addIs2 [LatticeDisjunctionIntro,LatticeReflexivity])); -qed "LatticeTriangle2"; - -(*** Lamport's fairness rules ***) -section "Fairness rules"; - -val prems = goal (the_context ()) - "[| |- $P & N --> P` | Q`; \ -\ |- ($P & N) & _v --> Q`; \ -\ |- $P & N --> $(Enabled(_v)) |] \ -\ ==> |- []N & WF(A)_v --> (P ~> Q)"; -by (clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1); -by (rtac (temp_use ensures) 1); -by (TRYALL (ares_tac prems)); -by (etac STL4Edup 1); -by (atac 1); -by (clarsimp_tac (temp_css addsimps2 [WF_def]) 1); -by (rtac (temp_use STL2) 1); -by (clarsimp_tac (temp_css addSEs2 [mp] addSIs2 [InitDmd]) 1); -by (resolve_tac ((map temp_use (prems RL [STL4])) RL [box_stp_actD]) 1); -by (asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_actI]) 1); -qed "WF1"; - -(* Sometimes easier to use; designed for action B rather than state predicate Q *) -val [prem1,prem2,prem3] = goalw (the_context ()) [leadsto_def] - "[| |- N & $P --> $Enabled (_v); \ -\ |- N & _v --> B; \ -\ |- [](N & [~A]_v) --> stable P |] \ -\ ==> |- []N & WF(A)_v --> (P ~> B)"; -by (clarsimp_tac (temp_css addSDs2 [BoxWFI]) 1); -by (etac STL4Edup 1); -by (atac 1); -by (Clarsimp_tac 1); -by (rtac (temp_use (prem2 RS DmdImpl)) 1); -by (rtac (temp_use BoxDmd_simple) 1); -by (atac 1); -by (rtac classical 1); -by (rtac (temp_use STL2) 1); -by (clarsimp_tac (temp_css addsimps2 [WF_def] addSEs2 [mp] addSIs2 [InitDmd]) 1); -by (rtac ((temp_use (prem1 RS STL4)) RS box_stp_actD) 1); -by (asm_simp_tac (simpset() addsimps [split_box_conj,box_stp_act]) 1); -by (etac (temp_use INV1) 1); -by (rtac (temp_use prem3) 1); -by (asm_full_simp_tac (simpset() addsimps [split_box_conj,temp_use NotDmd,not_angle]) 1); -qed "WF_leadsto"; - -val prems = goal (the_context ()) - "[| |- $P & N --> P` | Q`; \ -\ |- ($P & N) & _v --> Q`; \ -\ |- []P & []N & []F --> <>Enabled(_v) |] \ -\ ==> |- []N & SF(A)_v & []F --> (P ~> Q)"; -by (clarsimp_tac (temp_css addSDs2 [BoxSFI]) 1); -by (rtac (temp_use ensures) 1); -by (TRYALL (ares_tac prems)); -by (eres_inst_tac [("F","F")] dup_boxE 1); -by (merge_temp_box_tac 1); -by (etac STL4Edup 1); -by (atac 1); -by (clarsimp_tac (temp_css addsimps2 [SF_def]) 1); -by (rtac (temp_use STL2) 1); -by (etac mp 1); -by (resolve_tac (map temp_use (prems RL [STL4])) 1); -by (asm_simp_tac (simpset() addsimps [split_box_conj, STL3]) 1); -qed "SF1"; - -val [prem1,prem2,prem3,prem4] = goal (the_context ()) - "[| |- 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"; -by (clarsimp_tac (temp_css addSDs2 [BoxWFI, (temp_use BoxDmdBox) RS iffD2] - addsimps2 [read_instantiate [("A","M")] WF_def]) 1); -by (eres_inst_tac [("F","F")] dup_boxE 1); -by (merge_temp_box_tac 1); -by (etac STL4Edup 1); -by (atac 1); -by (clarsimp_tac (temp_css addSIs2 - [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1); -by (rtac classical 1); -by (subgoal_tac "sigmaa |= <>(($P & P` & N) & _f)" 1); - by (force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1); -by (rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1); -by (asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1); -by (merge_act_box_tac 1); -by (forward_tac [temp_use prem4] 1); -by (TRYALL atac); -by (dtac (temp_use STL6) 1); -by (atac 1); -by (eres_inst_tac [("V","sigmaa |= <>[]P")] thin_rl 1); -by (eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1); -by (dtac (temp_use BoxWFI) 1); -by (eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1); -by (merge_temp_box_tac 1); -by (etac DmdImpldup 1); -by (atac 1); -by (auto_tac (temp_css addsimps2 [split_box_conj,STL3,WF_Box,box_stp_act])); - by (force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1); -by (rtac (temp_use STL2) 1); -by (force_tac (temp_css addsimps2 [WF_def,split_box_conj] addSEs2 [mp] - addSIs2 [InitDmd, prem3 RS STL4]) 1); -qed "WF2"; - -val [prem1,prem2,prem3,prem4] = goal (the_context ()) - "[| |- 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"; -by (clarsimp_tac (temp_css addSDs2 [BoxSFI] - addsimps2 [read_instantiate [("A","M")] SF_def]) 1); -by (eres_inst_tac [("F","F")] dup_boxE 1); -by (eres_inst_tac [("F","TEMP <>Enabled(_g)")] dup_boxE 1); -by (merge_temp_box_tac 1); -by (etac STL4Edup 1); -by (atac 1); -by (clarsimp_tac (temp_css addSIs2 - [(temp_use BoxDmd_simple) RS (temp_use (prem1 RS DmdImpl))]) 1); -by (rtac classical 1); -by (subgoal_tac "sigmaa |= <>(($P & P` & N) & _f)" 1); - by (force_tac (temp_css addsimps2 [angle_def] addSIs2 [prem2] addSEs2 [DmdImplE]) 1); -by (rtac (temp_use (rewrite_rule [temp_rewrite DmdDmd] (BoxDmd_simple RS DmdImpl))) 1); -by (asm_full_simp_tac (simpset() addsimps [temp_use NotDmd, not_angle]) 1); -by (merge_act_box_tac 1); -by (forward_tac [temp_use prem4] 1); -by (TRYALL atac); -by (eres_inst_tac [("V","sigmaa |= []F")] thin_rl 1); -by (dtac (temp_use BoxSFI) 1); -by (eres_inst_tac [("F","TEMP <>Enabled(_g)")] dup_boxE 1); -by (eres_inst_tac [("F", "ACT N & [~B]_f")] dup_boxE 1); -by (merge_temp_box_tac 1); -by (etac DmdImpldup 1); -by (atac 1); -by (auto_tac (temp_css addsimps2 [split_box_conj,STL3,SF_Box,box_stp_act])); - by (force_tac (temp_css addSEs2 [read_instantiate [("P","P")] TLA2E]) 1); -by (rtac (temp_use STL2) 1); -by (force_tac (temp_css addsimps2 [SF_def,split_box_conj] addSEs2 [mp,InfImpl] - addSIs2 [prem3]) 1); -qed "SF2"; - -(* ------------------------------------------------------------------------- *) -(*** Liveness proofs by well-founded orderings ***) -(* ------------------------------------------------------------------------- *) -section "Well-founded orderings"; - -val p1::prems = goal (the_context ()) - "[| wf r; \ -\ !!x. sigma |= F x ~> (G | (EX y. #((y,x):r) & F y)) \ -\ |] ==> sigma |= F x ~> G"; -by (rtac (p1 RS wf_induct) 1); -by (rtac (temp_use LatticeTriangle) 1); -by (resolve_tac prems 1); -by (auto_tac (temp_css addsimps2 [leadsto_exists])); -by (case_tac "(y,x):r" 1); - by (Force_tac 1); -by (force_tac (temp_css addsimps2 leadsto_def::Init_simps addSIs2 [necT]) 1); -qed "wf_leadsto"; - -(* If r is well-founded, state function v cannot decrease forever *) -Goal "!!r. wf r ==> |- [][ (v`, $v) : #r ]_v --> <>[][#False]_v"; -by (Clarsimp_tac 1); -by (rtac ccontr 1); -by (subgoal_tac "sigma |= (EX x. v=#x) ~> #False" 1); - by (dtac ((temp_use leadsto_false) RS iffD1 RS (temp_use STL2_gen)) 1); - by (force_tac (temp_css addsimps2 Init_defs) 1); -by (clarsimp_tac (temp_css addsimps2 [leadsto_exists,not_square]@more_temp_simps) 1); -by (etac wf_leadsto 1); -by (rtac (temp_use ensures_simple) 1); -by (TRYALL atac); -by (auto_tac (temp_css addsimps2 [square_def,angle_def])); -qed "wf_not_box_decrease"; - -(* "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, then there - have to be infinitely many non-stuttering steps where v doesn't decrease. -*) -val [prem] = goal (the_context ()) - "wf r ==> |- []<>((v`, $v) : #r) --> []<><(v`, $v) ~: #r>_v"; -by (Clarsimp_tac 1); -by (rtac ccontr 1); -by (asm_full_simp_tac (simpset() addsimps not_angle::more_temp_simps) 1); -by (dtac (prem RS (temp_use wf_not_dmd_box_decrease)) 1); -by (dtac (temp_use BoxDmdDmdBox) 1); -by (atac 1); -by (subgoal_tac "sigma |= []<>((#False)::action)" 1); - by (Force_tac 1); -by (etac STL4E 1); -by (rtac DmdImpl 1); -by (force_tac (temp_css addIs2 [prem RS wf_irrefl]) 1); -qed "wf_box_dmd_decrease"; - -(* In particular, for natural numbers, if n decreases infinitely often - then it has to increase infinitely often. -*) -Goal "!!n::nat stfun. |- []<>(n` < $n) --> []<>($n < n`)"; -by (Clarsimp_tac 1); -by (subgoal_tac "sigma |= []<><~( (n`,$n) : #less_than )>_n" 1); - by (etac thin_rl 1); - by (etac STL4E 1); - by (rtac DmdImpl 1); - by (clarsimp_tac (temp_css addsimps2 [angle_def]) 1); -by (rtac (temp_use wf_box_dmd_decrease) 1); -by (auto_tac (temp_css addSEs2 [STL4E,DmdImplE])); -qed "nat_box_dmd_decrease"; - - -(* ------------------------------------------------------------------------- *) -(*** Flexible quantification over state variables ***) -(* ------------------------------------------------------------------------- *) -section "Flexible quantification"; - -val [prem1,prem2] = goal (the_context ()) - "[| basevars vs; (!!x. basevars (x,vs) ==> sigma |= F x) |]\ -\ ==> sigma |= (AALL x. F x)"; -by (auto_tac (temp_css addsimps2 [aall_def] addSEs2 [eexE] - addSIs2 [prem1] addSDs2 [prem2])); -qed "aallI"; - -Goalw [aall_def] "|- (AALL x. F x) --> F x"; -by (Clarsimp_tac 1); -by (etac contrapos_np 1); -by (force_tac (temp_css addSIs2 [eexI]) 1); -qed "aallE"; - -(* monotonicity of quantification *) -val [min,maj] = goal (the_context ()) - "[| sigma |= EEX x. F x; !!x. sigma |= F x --> G x |] ==> sigma |= EEX x. G x"; -by (rtac (unit_base RS (min RS eexE)) 1); -by (rtac (temp_use eexI) 1); -by (etac ((rewrite_rule intensional_rews maj) RS mp) 1); -qed "eex_mono"; - -val [min,maj] = goal (the_context ()) - "[| sigma |= AALL x. F(x); !!x. sigma |= F(x) --> G(x) |] ==> sigma |= AALL x. G(x)"; -by (rtac (unit_base RS aallI) 1); -by (rtac ((rewrite_rule intensional_rews maj) RS mp) 1); -by (rtac (min RS (temp_use aallE)) 1); -qed "aall_mono"; - -(* Derived history introduction rule *) -val [p1,p2,p3,p4,p5] = goal (the_context ()) - "[| 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)"; -by (rtac ((temp_use history) RS eexE) 1); - by (rtac p3 1); -by (rtac (temp_use eexI) 1); -by (Clarsimp_tac 1); -by (rtac conjI 1); -by (cut_facts_tac [p2] 2); -by (merge_box_tac 2); -by (force_tac (temp_css addSEs2 [STL4E,p5]) 2); -by (cut_facts_tac [p1] 1); -by (force_tac (temp_css addsimps2 Init_defs addSEs2 [p4]) 1); -qed "historyI"; - -(* ---------------------------------------------------------------------- - example of a history variable: existence of a clock - -Goal "|- EEX h. Init(h = #True) & [](h` = (~$h))"; -by (rtac tempI 1); -by (rtac historyI 1); -by (REPEAT (force_tac (temp_css addsimps2 Init_defs addIs2 [unit_base, necT]) 1)); -(** solved **) - ----------------------------------------------------------------------- *) diff -r 17098171d46a -r 6f79647cf536 src/HOL/TLA/TLA.thy --- a/src/HOL/TLA/TLA.thy Fri Dec 01 17:22:33 2006 +0100 +++ b/src/HOL/TLA/TLA.thy Sat Dec 02 02:52:02 2006 +0100 @@ -3,12 +3,9 @@ ID: $Id$ Author: Stephan Merz Copyright: 1998 University of Munich +*) - Theory Name: TLA - Logic Image: HOL - -The temporal level of TLA. -*) +header {* The temporal level of TLA *} theory TLA imports Init @@ -99,6 +96,1108 @@ |] ==> G sigma" history: "|- EEX h. Init(h = ha) & [](!x. $h = #x --> h` = hb x)" -ML {* use_legacy_bindings (the_context ()) *} + +(* Specialize intensional introduction/elimination rules for temporal formulas *) + +lemma tempI: "(!!sigma. sigma |= (F::temporal)) ==> |- F" + apply (rule intI) + apply (erule meta_spec) + done + +lemma tempD: "|- (F::temporal) ==> sigma |= F" + by (erule intD) + + +(* ======== Functions to "unlift" temporal theorems ====== *) + +ML {* +(* The following functions are specialized versions of the corresponding + functions defined in theory Intensional in that they introduce a + "world" parameter of type "behavior". +*) +local + val action_rews = thms "action_rews"; + val tempD = thm "tempD"; +in + +fun temp_unlift th = + (rewrite_rule action_rews (th RS tempD)) handle THM _ => action_unlift th; + +(* Turn |- F = G into meta-level rewrite rule F == G *) +val temp_rewrite = int_rewrite + +fun temp_use th = + case (concl_of th) of + Const _ $ (Const ("Intensional.Valid", _) $ _) => + ((flatten (temp_unlift th)) handle THM _ => th) + | _ => th; + +fun try_rewrite th = temp_rewrite th handle THM _ => temp_use th; + +end +*} + +setup {* + Attrib.add_attributes [ + ("temp_unlift", Attrib.no_args (Thm.rule_attribute (K temp_unlift)), ""), + ("temp_rewrite", Attrib.no_args (Thm.rule_attribute (K temp_rewrite)), ""), + ("temp_use", Attrib.no_args (Thm.rule_attribute (K temp_use)), ""), + ("try_rewrite", Attrib.no_args (Thm.rule_attribute (K try_rewrite)), "")] +*} + +(* Update classical reasoner---will be updated once more below! *) + +declare tempI [intro!] +declare tempD [dest] +ML {* +val temp_css = (claset(), simpset()) +val temp_cs = op addss temp_css +*} + +(* Modify the functions that add rules to simpsets, classical sets, + and clasimpsets in order to accept "lifted" theorems +*) + +(* ------------------------------------------------------------------------- *) +(*** "Simple temporal logic": only [] and <> ***) +(* ------------------------------------------------------------------------- *) +section "Simple temporal logic" + +(* []~F == []~Init F *) +lemmas boxNotInit = boxInit [of "LIFT ~F", unfolded Init_simps, standard] + +lemma dmdInit: "TEMP <>F == TEMP <> Init F" + apply (unfold dmd_def) + apply (unfold boxInit [of "LIFT ~F"]) + apply (simp (no_asm) add: Init_simps) + done + +lemmas dmdNotInit = dmdInit [of "LIFT ~F", unfolded Init_simps, standard] + +(* boxInit and dmdInit cannot be used as rewrites, because they loop. + Non-looping instances for state predicates and actions are occasionally useful. +*) +lemmas boxInit_stp = boxInit [where 'a = state, standard] +lemmas boxInit_act = boxInit [where 'a = "state * state", standard] +lemmas dmdInit_stp = dmdInit [where 'a = state, standard] +lemmas dmdInit_act = dmdInit [where 'a = "state * state", standard] + +(* The symmetric equations can be used to get rid of Init *) +lemmas boxInitD = boxInit [symmetric] +lemmas dmdInitD = dmdInit [symmetric] +lemmas boxNotInitD = boxNotInit [symmetric] +lemmas dmdNotInitD = dmdNotInit [symmetric] + +lemmas Init_simps = Init_simps boxInitD dmdInitD boxNotInitD dmdNotInitD + +(* ------------------------ STL2 ------------------------------------------- *) +lemmas STL2 = reflT + +(* The "polymorphic" (generic) variant *) +lemma STL2_gen: "|- []F --> Init F" + apply (unfold boxInit [of F]) + apply (rule STL2) + done + +(* see also STL2_pr below: "|- []P --> Init P & Init (P`)" *) + + +(* Dual versions for <> *) +lemma InitDmd: "|- F --> <> F" + apply (unfold dmd_def) + apply (auto dest!: STL2 [temp_use]) + done + +lemma InitDmd_gen: "|- Init F --> <>F" + apply clarsimp + apply (drule InitDmd [temp_use]) + apply (simp add: dmdInitD) + done + + +(* ------------------------ STL3 ------------------------------------------- *) +lemma STL3: "|- ([][]F) = ([]F)" + by (auto elim: transT [temp_use] STL2 [temp_use]) + +(* corresponding elimination rule introduces double boxes: + [| (sigma |= []F); (sigma |= [][]F) ==> PROP W |] ==> PROP W +*) +lemmas dup_boxE = STL3 [temp_unlift, THEN iffD2, elim_format] +lemmas dup_boxD = STL3 [temp_unlift, THEN iffD1, standard] + +(* dual versions for <> *) +lemma DmdDmd: "|- (<><>F) = (<>F)" + by (auto simp add: dmd_def [try_rewrite] STL3 [try_rewrite]) + +lemmas dup_dmdE = DmdDmd [temp_unlift, THEN iffD2, elim_format] +lemmas dup_dmdD = DmdDmd [temp_unlift, THEN iffD1, standard] + + +(* ------------------------ STL4 ------------------------------------------- *) +lemma STL4: + assumes "|- F --> G" + shows "|- []F --> []G" + apply clarsimp + apply (rule normalT [temp_use]) + apply (rule assms [THEN necT, temp_use]) + apply assumption + done + +(* Unlifted version as an elimination rule *) +lemma STL4E: "[| sigma |= []F; |- F --> G |] ==> sigma |= []G" + by (erule (1) STL4 [temp_use]) + +lemma STL4_gen: "|- Init F --> Init G ==> |- []F --> []G" + apply (drule STL4) + apply (simp add: boxInitD) + done + +lemma STL4E_gen: "[| sigma |= []F; |- Init F --> Init G |] ==> sigma |= []G" + by (erule (1) STL4_gen [temp_use]) + +(* see also STL4Edup below, which allows an auxiliary boxed formula: + []A /\ F => G + ----------------- + []A /\ []F => []G +*) + +(* The dual versions for <> *) +lemma DmdImpl: + assumes prem: "|- F --> G" + shows "|- <>F --> <>G" + apply (unfold dmd_def) + apply (fastsimp intro!: prem [temp_use] elim!: STL4E [temp_use]) + done + +lemma DmdImplE: "[| sigma |= <>F; |- F --> G |] ==> sigma |= <>G" + by (erule (1) DmdImpl [temp_use]) + +(* ------------------------ STL5 ------------------------------------------- *) +lemma STL5: "|- ([]F & []G) = ([](F & G))" + apply auto + apply (subgoal_tac "sigma |= [] (G --> (F & G))") + apply (erule normalT [temp_use]) + apply (fastsimp elim!: STL4E [temp_use])+ + done + +(* rewrite rule to split conjunctions under boxes *) +lemmas split_box_conj = STL5 [temp_unlift, symmetric, standard] + + +(* 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! +*) +lemma box_conjE: + assumes "sigma |= []F" + and "sigma |= []G" + and "sigma |= [](F&G) ==> PROP R" + shows "PROP R" + by (rule assms STL5 [temp_unlift, THEN iffD1] conjI)+ + +(* Instances of box_conjE for state predicates, actions, and temporals + in case the general rule is "too polymorphic". +*) +lemmas box_conjE_temp = box_conjE [where 'a = behavior, standard] +lemmas box_conjE_stp = box_conjE [where 'a = state, standard] +lemmas box_conjE_act = box_conjE [where 'a = "state * state", standard] + +(* Define a tactic that tries to merge all boxes in an antecedent. The definition is + a bit kludgy in order to simulate "double elim-resolution". +*) + +lemma box_thin: "[| sigma |= []F; PROP W |] ==> PROP W" . + +ML {* +local + val box_conjE = thm "box_conjE"; + val box_thin = thm "box_thin"; + val box_conjE_temp = thm "box_conjE_temp"; + val box_conjE_stp = thm "box_conjE_stp"; + val box_conjE_act = thm "box_conjE_act"; +in + +fun merge_box_tac i = + REPEAT_DETERM (EVERY [etac box_conjE i, atac i, etac box_thin i]) + +fun merge_temp_box_tac i = + REPEAT_DETERM (EVERY [etac box_conjE_temp i, atac i, + eres_inst_tac [("'a","behavior")] box_thin i]) + +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 box_conjE_act i, atac i, + eres_inst_tac [("'a","state * state")] box_thin i]) end +*} + +(* rewrite rule to push universal quantification through box: + (sigma |= [](! x. F x)) = (! x. (sigma |= []F x)) +*) +lemmas all_box = allT [temp_unlift, symmetric, standard] + +lemma DmdOr: "|- (<>(F | G)) = (<>F | <>G)" + apply (auto simp add: dmd_def split_box_conj [try_rewrite]) + apply (erule contrapos_np, tactic "merge_box_tac 1", + fastsimp elim!: STL4E [temp_use])+ + done + +lemma exT: "|- (EX x. <>(F x)) = (<>(EX x. F x))" + by (auto simp: dmd_def Not_Rex [try_rewrite] all_box [try_rewrite]) + +lemmas ex_dmd = exT [temp_unlift, symmetric, standard] + +lemma STL4Edup: "!!sigma. [| sigma |= []A; sigma |= []F; |- F & []A --> G |] ==> sigma |= []G" + apply (erule dup_boxE) + apply (tactic "merge_box_tac 1") + apply (erule STL4E) + apply assumption + done + +lemma DmdImpl2: + "!!sigma. [| sigma |= <>F; sigma |= [](F --> G) |] ==> sigma |= <>G" + apply (unfold dmd_def) + apply auto + apply (erule notE) + apply (tactic "merge_box_tac 1") + apply (fastsimp elim!: STL4E [temp_use]) + done + +lemma InfImpl: + assumes 1: "sigma |= []<>F" + and 2: "sigma |= []G" + and 3: "|- F & G --> H" + shows "sigma |= []<>H" + apply (insert 1 2) + apply (erule_tac F = G in dup_boxE) + apply (tactic "merge_box_tac 1") + apply (fastsimp elim!: STL4E [temp_use] DmdImpl2 [temp_use] intro!: 3 [temp_use]) + done + +(* ------------------------ STL6 ------------------------------------------- *) +(* Used in the proof of STL6, but useful in itself. *) +lemma BoxDmd: "|- []F & <>G --> <>([]F & G)" + apply (unfold dmd_def) + apply clarsimp + apply (erule dup_boxE) + apply (tactic "merge_box_tac 1") + apply (erule contrapos_np) + apply (fastsimp elim!: STL4E [temp_use]) + done + +(* weaker than BoxDmd, but more polymorphic (and often just right) *) +lemma BoxDmd_simple: "|- []F & <>G --> <>(F & G)" + apply (unfold dmd_def) + apply clarsimp + apply (tactic "merge_box_tac 1") + apply (fastsimp elim!: notE STL4E [temp_use]) + done + +lemma BoxDmd2_simple: "|- []F & <>G --> <>(G & F)" + apply (unfold dmd_def) + apply clarsimp + apply (tactic "merge_box_tac 1") + apply (fastsimp elim!: notE STL4E [temp_use]) + done + +lemma DmdImpldup: + assumes 1: "sigma |= []A" + and 2: "sigma |= <>F" + and 3: "|- []A & F --> G" + shows "sigma |= <>G" + apply (rule 2 [THEN 1 [THEN BoxDmd [temp_use]], THEN DmdImplE]) + apply (rule 3) + done + +lemma STL6: "|- <>[]F & <>[]G --> <>[](F & G)" + apply (auto simp: STL5 [temp_rewrite, symmetric]) + apply (drule linT [temp_use]) + apply assumption + apply (erule thin_rl) + apply (rule DmdDmd [temp_unlift, THEN iffD1]) + apply (erule disjE) + apply (erule DmdImplE) + apply (rule BoxDmd) + apply (erule DmdImplE) + apply auto + apply (drule BoxDmd [temp_use]) + apply assumption + apply (erule thin_rl) + apply (fastsimp elim!: DmdImplE [temp_use]) + done + + +(* ------------------------ True / False ----------------------------------------- *) +section "Simplification of constants" + +lemma BoxConst: "|- ([]#P) = #P" + apply (rule tempI) + apply (cases P) + apply (auto intro!: necT [temp_use] dest: STL2_gen [temp_use] simp: Init_simps) + done + +lemma DmdConst: "|- (<>#P) = #P" + apply (unfold dmd_def) + apply (cases P) + apply (simp_all add: BoxConst [try_rewrite]) + done + +lemmas temp_simps [temp_rewrite, simp] = BoxConst DmdConst + +(* Make these rewrites active by default *) +ML {* +val temp_css = temp_css addsimps2 (thms "temp_simps") +val temp_cs = op addss temp_css +*} + + +(* ------------------------ Further rewrites ----------------------------------------- *) +section "Further rewrites" + +lemma NotBox: "|- (~[]F) = (<>~F)" + by (simp add: dmd_def) + +lemma NotDmd: "|- (~<>F) = ([]~F)" + by (simp add: dmd_def) + +(* These are not declared by default, because they could be harmful, + e.g. []F & ~[]F becomes []F & <>~F !! *) +lemmas more_temp_simps = + STL3 [temp_rewrite] DmdDmd [temp_rewrite] NotBox [temp_rewrite] NotDmd [temp_rewrite] + NotBox [temp_unlift, THEN eq_reflection] + NotDmd [temp_unlift, THEN eq_reflection] + +lemma BoxDmdBox: "|- ([]<>[]F) = (<>[]F)" + apply (auto dest!: STL2 [temp_use]) + apply (rule ccontr) + apply (subgoal_tac "sigma |= <>[][]F & <>[]~[]F") + apply (erule thin_rl) + apply auto + apply (drule STL6 [temp_use]) + apply assumption + apply simp + apply (simp_all add: more_temp_simps) + done + +lemma DmdBoxDmd: "|- (<>[]<>F) = ([]<>F)" + apply (unfold dmd_def) + apply (auto simp: BoxDmdBox [unfolded dmd_def, try_rewrite]) + done + +lemmas more_temp_simps = more_temp_simps BoxDmdBox [temp_rewrite] DmdBoxDmd [temp_rewrite] + + +(* ------------------------ Miscellaneous ----------------------------------- *) + +lemma BoxOr: "!!sigma. [| sigma |= []F | []G |] ==> sigma |= [](F | G)" + by (fastsimp elim!: STL4E [temp_use]) + +(* "persistently implies infinitely often" *) +lemma DBImplBD: "|- <>[]F --> []<>F" + apply clarsimp + apply (rule ccontr) + apply (simp add: more_temp_simps) + apply (drule STL6 [temp_use]) + apply assumption + apply simp + done + +lemma BoxDmdDmdBox: "|- []<>F & <>[]G --> []<>(F & G)" + apply clarsimp + apply (rule ccontr) + apply (unfold more_temp_simps) + apply (drule STL6 [temp_use]) + apply assumption + apply (subgoal_tac "sigma |= <>[]~F") + apply (force simp: dmd_def) + apply (fastsimp elim: DmdImplE [temp_use] STL4E [temp_use]) + done + + +(* ------------------------------------------------------------------------- *) +(*** TLA-specific theorems: primed formulas ***) +(* ------------------------------------------------------------------------- *) +section "priming" + +(* ------------------------ TLA2 ------------------------------------------- *) +lemma STL2_pr: "|- []P --> Init P & Init P`" + by (fastsimp intro!: STL2_gen [temp_use] primeI [temp_use]) + +(* Auxiliary lemma allows priming of boxed actions *) +lemma BoxPrime: "|- []P --> []($P & P$)" + apply clarsimp + apply (erule dup_boxE) + apply (unfold boxInit_act) + apply (erule STL4E) + apply (auto simp: Init_simps dest!: STL2_pr [temp_use]) + done + +lemma TLA2: + assumes "|- $P & P$ --> A" + shows "|- []P --> []A" + apply clarsimp + apply (drule BoxPrime [temp_use]) + apply (auto simp: Init_stp_act_rev [try_rewrite] intro!: prems [temp_use] + elim!: STL4E [temp_use]) + done + +lemma TLA2E: "[| sigma |= []P; |- $P & P$ --> A |] ==> sigma |= []A" + by (erule (1) TLA2 [temp_use]) + +lemma DmdPrime: "|- (<>P`) --> (<>P)" + apply (unfold dmd_def) + apply (fastsimp elim!: TLA2E [temp_use]) + done + +lemmas PrimeDmd = InitDmd_gen [temp_use, THEN DmdPrime [temp_use], standard] + +(* ------------------------ INV1, stable --------------------------------------- *) +section "stable, invariant" + +lemma ind_rule: + "[| sigma |= []H; sigma |= Init P; |- H --> (Init P & ~[]F --> Init(P`) & F) |] + ==> sigma |= []F" + apply (rule indT [temp_use]) + apply (erule (2) STL4E) + done + +lemma box_stp_act: "|- ([]$P) = ([]P)" + by (simp add: boxInit_act Init_simps) + +lemmas box_stp_actI = box_stp_act [temp_use, THEN iffD2, standard] +lemmas box_stp_actD = box_stp_act [temp_use, THEN iffD1, standard] + +lemmas more_temp_simps = box_stp_act [temp_rewrite] more_temp_simps + +lemma INV1: + "|- (Init P) --> (stable P) --> []P" + apply (unfold stable_def boxInit_stp boxInit_act) + apply clarsimp + apply (erule ind_rule) + apply (auto simp: Init_simps elim: ind_rule) + done + +lemma StableT: + "!!P. |- $P & A --> P` ==> |- []A --> stable P" + apply (unfold stable_def) + apply (fastsimp elim!: STL4E [temp_use]) + done + +lemma Stable: "[| sigma |= []A; |- $P & A --> P` |] ==> sigma |= stable P" + by (erule (1) StableT [temp_use]) + +(* Generalization of INV1 *) +lemma StableBox: "|- (stable P) --> [](Init P --> []P)" + apply (unfold stable_def) + apply clarsimp + apply (erule dup_boxE) + apply (force simp: stable_def elim: STL4E [temp_use] INV1 [temp_use]) + done + +lemma DmdStable: "|- (stable P) & <>P --> <>[]P" + apply clarsimp + apply (rule DmdImpl2) + prefer 2 + apply (erule StableBox [temp_use]) + apply (simp add: dmdInitD) + done + +(* ---------------- (Semi-)automatic invariant tactics ---------------------- *) + +ML {* +local + val INV1 = thm "INV1"; + val Stable = thm "Stable"; + val Init_stp = thm "Init_stp"; + val Init_act = thm "Init_act"; + val squareE = thm "squareE"; +in + +(* inv_tac reduces goals of the form ... ==> sigma |= []P *) +fun inv_tac css = SELECT_GOAL + (EVERY [auto_tac css, + TRY (merge_box_tac 1), + 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 these simple cases the simplifier seems to be more useful than the + auto-tactic, which applies too much propositional logic and simplifies + too late. +*) +fun auto_inv_tac ss = SELECT_GOAL + ((inv_tac (claset(),ss) 1) THEN + (TRYALL (action_simp_tac (ss addsimps [Init_stp, Init_act]) [] [squareE]))); +end +*} + +lemma unless: "|- []($P --> P` | Q`) --> (stable P) | <>Q" + apply (unfold dmd_def) + apply (clarsimp dest!: BoxPrime [temp_use]) + apply (tactic "merge_box_tac 1") + apply (erule contrapos_np) + apply (fastsimp elim!: Stable [temp_use]) + done + + +(* --------------------- Recursive expansions --------------------------------------- *) +section "recursive expansions" + +(* Recursive expansions of [] and <> for state predicates *) +lemma BoxRec: "|- ([]P) = (Init P & []P`)" + apply (auto intro!: STL2_gen [temp_use]) + apply (fastsimp elim!: TLA2E [temp_use]) + apply (auto simp: stable_def elim!: INV1 [temp_use] STL4E [temp_use]) + done + +lemma DmdRec: "|- (<>P) = (Init P | <>P`)" + apply (unfold dmd_def BoxRec [temp_rewrite]) + apply (auto simp: Init_simps) + done + +lemma DmdRec2: "!!sigma. [| sigma |= <>P; sigma |= []~P` |] ==> sigma |= Init P" + apply (force simp: DmdRec [temp_rewrite] dmd_def) + done + +lemma InfinitePrime: "|- ([]<>P) = ([]<>P`)" + apply auto + apply (rule classical) + apply (rule DBImplBD [temp_use]) + apply (subgoal_tac "sigma |= <>[]P") + apply (fastsimp elim!: DmdImplE [temp_use] TLA2E [temp_use]) + apply (subgoal_tac "sigma |= <>[] (<>P & []~P`)") + apply (force simp: boxInit_stp [temp_use] + elim!: DmdImplE [temp_use] STL4E [temp_use] DmdRec2 [temp_use]) + apply (force intro!: STL6 [temp_use] simp: more_temp_simps) + apply (fastsimp intro: DmdPrime [temp_use] elim!: STL4E [temp_use]) + done + +lemma InfiniteEnsures: + "[| sigma |= []N; sigma |= []<>A; |- A & N --> P` |] ==> sigma |= []<>P" + apply (unfold InfinitePrime [temp_rewrite]) + apply (rule InfImpl) + apply assumption+ + done + +(* ------------------------ fairness ------------------------------------------- *) +section "fairness" + +(* alternative definitions of fairness *) +lemma WF_alt: "|- WF(A)_v = ([]<>~Enabled(_v) | []<>_v)" + apply (unfold WF_def dmd_def) + apply fastsimp + done + +lemma SF_alt: "|- SF(A)_v = (<>[]~Enabled(_v) | []<>_v)" + apply (unfold SF_def dmd_def) + apply fastsimp + done + +(* theorems to "box" fairness conditions *) +lemma BoxWFI: "|- WF(A)_v --> []WF(A)_v" + by (auto simp: WF_alt [try_rewrite] more_temp_simps intro!: BoxOr [temp_use]) + +lemma WF_Box: "|- ([]WF(A)_v) = WF(A)_v" + by (fastsimp intro!: BoxWFI [temp_use] dest!: STL2 [temp_use]) + +lemma BoxSFI: "|- SF(A)_v --> []SF(A)_v" + by (auto simp: SF_alt [try_rewrite] more_temp_simps intro!: BoxOr [temp_use]) + +lemma SF_Box: "|- ([]SF(A)_v) = SF(A)_v" + by (fastsimp intro!: BoxSFI [temp_use] dest!: STL2 [temp_use]) + +lemmas more_temp_simps = more_temp_simps WF_Box [temp_rewrite] SF_Box [temp_rewrite] + +lemma SFImplWF: "|- SF(A)_v --> WF(A)_v" + apply (unfold SF_def WF_def) + apply (fastsimp dest!: DBImplBD [temp_use]) + done + +(* A tactic that "boxes" all fairness conditions. Apply more_temp_simps to "unbox". *) +ML {* +local + val BoxWFI = thm "BoxWFI"; + val BoxSFI = thm "BoxSFI"; +in +val box_fair_tac = SELECT_GOAL (REPEAT (dresolve_tac [BoxWFI,BoxSFI] 1)) +end +*} + + +(* ------------------------------ leads-to ------------------------------ *) + +section "~>" + +lemma leadsto_init: "|- (Init F) & (F ~> G) --> <>G" + apply (unfold leadsto_def) + apply (auto dest!: STL2 [temp_use]) + done + +(* |- F & (F ~> G) --> <>G *) +lemmas leadsto_init_temp = leadsto_init [where 'a = behavior, unfolded Init_simps, standard] + +lemma streett_leadsto: "|- ([]<>Init F --> []<>G) = (<>(F ~> G))" + apply (unfold leadsto_def) + apply auto + apply (simp add: more_temp_simps) + apply (fastsimp elim!: DmdImplE [temp_use] STL4E [temp_use]) + apply (fastsimp intro!: InitDmd [temp_use] elim!: STL4E [temp_use]) + apply (subgoal_tac "sigma |= []<><>G") + apply (simp add: more_temp_simps) + apply (drule BoxDmdDmdBox [temp_use]) + apply assumption + apply (fastsimp elim!: DmdImplE [temp_use] STL4E [temp_use]) + done + +lemma leadsto_infinite: "|- []<>F & (F ~> G) --> []<>G" + apply clarsimp + apply (erule InitDmd [temp_use, THEN streett_leadsto [temp_unlift, THEN iffD2, THEN mp]]) + apply (simp add: dmdInitD) + done + +(* In particular, strong fairness is a Streett condition. The following + rules are sometimes easier to use than WF2 or SF2 below. +*) +lemma leadsto_SF: "|- (Enabled(_v) ~> _v) --> SF(A)_v" + apply (unfold SF_def) + apply (clarsimp elim!: leadsto_infinite [temp_use]) + done + +lemma leadsto_WF: "|- (Enabled(_v) ~> _v) --> WF(A)_v" + by (clarsimp intro!: SFImplWF [temp_use] leadsto_SF [temp_use]) + +(* introduce an invariant into the proof of a leadsto assertion. + []I --> ((P ~> Q) = (P /\ I ~> Q)) +*) +lemma INV_leadsto: "|- []I & (P & I ~> Q) --> (P ~> Q)" + apply (unfold leadsto_def) + apply clarsimp + apply (erule STL4Edup) + apply assumption + apply (auto simp: Init_simps dest!: STL2_gen [temp_use]) + done + +lemma leadsto_classical: "|- (Init F & []~G ~> G) --> (F ~> G)" + apply (unfold leadsto_def dmd_def) + apply (force simp: Init_simps elim!: STL4E [temp_use]) + done + +lemma leadsto_false: "|- (F ~> #False) = ([]~F)" + apply (unfold leadsto_def) + apply (simp add: boxNotInitD) + done + +lemma leadsto_exists: "|- ((EX x. F x) ~> G) = (ALL x. (F x ~> G))" + apply (unfold leadsto_def) + apply (auto simp: allT [try_rewrite] Init_simps elim!: STL4E [temp_use]) + done + +(* basic leadsto properties, cf. Unity *) + +lemma ImplLeadsto_gen: "|- [](Init F --> Init G) --> (F ~> G)" + apply (unfold leadsto_def) + apply (auto intro!: InitDmd_gen [temp_use] + elim!: STL4E_gen [temp_use] simp: Init_simps) + done + +lemmas ImplLeadsto = ImplLeadsto_gen [where 'a = behavior and 'b = behavior, + unfolded Init_simps, standard] + +lemma ImplLeadsto_simple: "!!F G. |- F --> G ==> |- F ~> G" + by (auto simp: Init_def intro!: ImplLeadsto_gen [temp_use] necT [temp_use]) + +lemma EnsuresLeadsto: + assumes "|- A & $P --> Q`" + shows "|- []A --> (P ~> Q)" + apply (unfold leadsto_def) + apply (clarsimp elim!: INV_leadsto [temp_use]) + apply (erule STL4E_gen) + apply (auto simp: Init_defs intro!: PrimeDmd [temp_use] assms [temp_use]) + done + +lemma EnsuresLeadsto2: "|- []($P --> Q`) --> (P ~> Q)" + apply (unfold leadsto_def) + apply clarsimp + apply (erule STL4E_gen) + apply (auto simp: Init_simps intro!: PrimeDmd [temp_use]) + done + +lemma ensures: + assumes 1: "|- $P & N --> P` | Q`" + and 2: "|- ($P & N) & A --> Q`" + shows "|- []N & []([]P --> <>A) --> (P ~> Q)" + apply (unfold leadsto_def) + apply clarsimp + apply (erule STL4Edup) + apply assumption + apply clarsimp + apply (subgoal_tac "sigmaa |= [] ($P --> P` | Q`) ") + apply (drule unless [temp_use]) + apply (clarsimp dest!: INV1 [temp_use]) + apply (rule 2 [THEN DmdImpl, temp_use, THEN DmdPrime [temp_use]]) + apply (force intro!: BoxDmd_simple [temp_use] + simp: split_box_conj [try_rewrite] box_stp_act [try_rewrite]) + apply (force elim: STL4E [temp_use] dest: 1 [temp_use]) + done + +lemma ensures_simple: + "[| |- $P & N --> P` | Q`; + |- ($P & N) & A --> Q` + |] ==> |- []N & []<>A --> (P ~> Q)" + apply clarsimp + apply (erule (2) ensures [temp_use]) + apply (force elim!: STL4E [temp_use]) + done + +lemma EnsuresInfinite: + "[| sigma |= []<>P; sigma |= []A; |- A & $P --> Q` |] ==> sigma |= []<>Q" + apply (erule leadsto_infinite [temp_use]) + apply (erule EnsuresLeadsto [temp_use]) + apply assumption + done + + +(*** Gronning's lattice rules (taken from TLP) ***) +section "Lattice rules" + +lemma LatticeReflexivity: "|- F ~> F" + apply (unfold leadsto_def) + apply (rule necT InitDmd_gen)+ + done + +lemma LatticeTransitivity: "|- (G ~> H) & (F ~> G) --> (F ~> H)" + apply (unfold leadsto_def) + apply clarsimp + apply (erule dup_boxE) (* [][] (Init G --> H) *) + apply (tactic "merge_box_tac 1") + apply (clarsimp elim!: STL4E [temp_use]) + apply (rule dup_dmdD) + apply (subgoal_tac "sigmaa |= <>Init G") + apply (erule DmdImpl2) + apply assumption + apply (simp add: dmdInitD) + done + +lemma LatticeDisjunctionElim1: "|- (F | G ~> H) --> (F ~> H)" + apply (unfold leadsto_def) + apply (auto simp: Init_simps elim!: STL4E [temp_use]) + done + +lemma LatticeDisjunctionElim2: "|- (F | G ~> H) --> (G ~> H)" + apply (unfold leadsto_def) + apply (auto simp: Init_simps elim!: STL4E [temp_use]) + done + +lemma LatticeDisjunctionIntro: "|- (F ~> H) & (G ~> H) --> (F | G ~> H)" + apply (unfold leadsto_def) + apply clarsimp + apply (tactic "merge_box_tac 1") + apply (auto simp: Init_simps elim!: STL4E [temp_use]) + done + +lemma LatticeDisjunction: "|- (F | G ~> H) = ((F ~> H) & (G ~> H))" + by (auto intro: LatticeDisjunctionIntro [temp_use] + LatticeDisjunctionElim1 [temp_use] + LatticeDisjunctionElim2 [temp_use]) + +lemma LatticeDiamond: "|- (A ~> B | C) & (B ~> D) & (C ~> D) --> (A ~> D)" + apply clarsimp + apply (subgoal_tac "sigma |= (B | C) ~> D") + apply (erule_tac G = "LIFT (B | C)" in LatticeTransitivity [temp_use]) + apply (fastsimp intro!: LatticeDisjunctionIntro [temp_use])+ + done + +lemma LatticeTriangle: "|- (A ~> D | B) & (B ~> D) --> (A ~> D)" + apply clarsimp + apply (subgoal_tac "sigma |= (D | B) ~> D") + apply (erule_tac G = "LIFT (D | B)" in LatticeTransitivity [temp_use]) + apply assumption + apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use]) + done + +lemma LatticeTriangle2: "|- (A ~> B | D) & (B ~> D) --> (A ~> D)" + apply clarsimp + apply (subgoal_tac "sigma |= B | D ~> D") + apply (erule_tac G = "LIFT (B | D)" in LatticeTransitivity [temp_use]) + apply assumption + apply (auto intro: LatticeDisjunctionIntro [temp_use] LatticeReflexivity [temp_use]) + done + +(*** Lamport's fairness rules ***) +section "Fairness rules" + +lemma WF1: + "[| |- $P & N --> P` | Q`; + |- ($P & N) & _v --> Q`; + |- $P & N --> $(Enabled(_v)) |] + ==> |- []N & WF(A)_v --> (P ~> Q)" + apply (clarsimp dest!: BoxWFI [temp_use]) + apply (erule (2) ensures [temp_use]) + apply (erule (1) STL4Edup) + apply (clarsimp simp: WF_def) + apply (rule STL2 [temp_use]) + apply (clarsimp elim!: mp intro!: InitDmd [temp_use]) + apply (erule STL4 [temp_use, THEN box_stp_actD [temp_use]]) + apply (simp add: split_box_conj box_stp_actI) + done + +(* Sometimes easier to use; designed for action B rather than state predicate Q *) +lemma WF_leadsto: + assumes 1: "|- N & $P --> $Enabled (_v)" + and 2: "|- N & _v --> B" + and 3: "|- [](N & [~A]_v) --> stable P" + shows "|- []N & WF(A)_v --> (P ~> B)" + apply (unfold leadsto_def) + apply (clarsimp dest!: BoxWFI [temp_use]) + apply (erule (1) STL4Edup) + apply clarsimp + apply (rule 2 [THEN DmdImpl, temp_use]) + apply (rule BoxDmd_simple [temp_use]) + apply assumption + apply (rule classical) + apply (rule STL2 [temp_use]) + apply (clarsimp simp: WF_def elim!: mp intro!: InitDmd [temp_use]) + apply (rule 1 [THEN STL4, temp_use, THEN box_stp_actD]) + apply (simp (no_asm_simp) add: split_box_conj [try_rewrite] box_stp_act [try_rewrite]) + apply (erule INV1 [temp_use]) + apply (rule 3 [temp_use]) + apply (simp add: split_box_conj [try_rewrite] NotDmd [temp_use] not_angle [try_rewrite]) + done + +lemma SF1: + "[| |- $P & N --> P` | Q`; + |- ($P & N) & _v --> Q`; + |- []P & []N & []F --> <>Enabled(_v) |] + ==> |- []N & SF(A)_v & []F --> (P ~> Q)" + apply (clarsimp dest!: BoxSFI [temp_use]) + apply (erule (2) ensures [temp_use]) + apply (erule_tac F = F in dup_boxE) + apply (tactic "merge_temp_box_tac 1") + apply (erule STL4Edup) + apply assumption + apply (clarsimp simp: SF_def) + apply (rule STL2 [temp_use]) + apply (erule mp) + apply (erule STL4 [temp_use]) + apply (simp add: split_box_conj [try_rewrite] STL3 [try_rewrite]) + done + +lemma WF2: + assumes 1: "|- N & _f --> _g" + and 2: "|- $P & P` & _f --> B" + and 3: "|- P & Enabled(_g) --> Enabled(_f)" + and 4: "|- [](N & [~B]_f) & WF(A)_f & []F & <>[]Enabled(_g) --> <>[]P" + shows "|- []N & WF(A)_f & []F --> WF(M)_g" + apply (clarsimp dest!: BoxWFI [temp_use] BoxDmdBox [temp_use, THEN iffD2] + simp: WF_def [where A = M]) + apply (erule_tac F = F in dup_boxE) + apply (tactic "merge_temp_box_tac 1") + apply (erule STL4Edup) + apply assumption + apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]]) + apply (rule classical) + apply (subgoal_tac "sigmaa |= <> (($P & P` & N) & _f)") + apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use]) + apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use]) + apply (simp add: NotDmd [temp_use] not_angle [try_rewrite]) + apply (tactic "merge_act_box_tac 1") + apply (frule 4 [temp_use]) + apply assumption+ + apply (drule STL6 [temp_use]) + apply assumption + apply (erule_tac V = "sigmaa |= <>[]P" in thin_rl) + apply (erule_tac V = "sigmaa |= []F" in thin_rl) + apply (drule BoxWFI [temp_use]) + apply (erule_tac F = "ACT N & [~B]_f" in dup_boxE) + apply (tactic "merge_temp_box_tac 1") + apply (erule DmdImpldup) + apply assumption + apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite] + WF_Box [try_rewrite] box_stp_act [try_rewrite]) + apply (force elim!: TLA2E [where P = P, temp_use]) + apply (rule STL2 [temp_use]) + apply (force simp: WF_def split_box_conj [try_rewrite] + elim!: mp intro!: InitDmd [temp_use] 3 [THEN STL4, temp_use]) + done + +lemma SF2: + assumes 1: "|- N & _f --> _g" + and 2: "|- $P & P` & _f --> B" + and 3: "|- P & Enabled(_g) --> Enabled(_f)" + and 4: "|- [](N & [~B]_f) & SF(A)_f & []F & []<>Enabled(_g) --> <>[]P" + shows "|- []N & SF(A)_f & []F --> SF(M)_g" + apply (clarsimp dest!: BoxSFI [temp_use] simp: 2 [try_rewrite] SF_def [where A = M]) + apply (erule_tac F = F in dup_boxE) + apply (erule_tac F = "TEMP <>Enabled (_g) " in dup_boxE) + apply (tactic "merge_temp_box_tac 1") + apply (erule STL4Edup) + apply assumption + apply (clarsimp intro!: BoxDmd_simple [temp_use, THEN 1 [THEN DmdImpl, temp_use]]) + apply (rule classical) + apply (subgoal_tac "sigmaa |= <> (($P & P` & N) & _f)") + apply (force simp: angle_def intro!: 2 [temp_use] elim!: DmdImplE [temp_use]) + apply (rule BoxDmd_simple [THEN DmdImpl, unfolded DmdDmd [temp_rewrite], temp_use]) + apply (simp add: NotDmd [temp_use] not_angle [try_rewrite]) + apply (tactic "merge_act_box_tac 1") + apply (frule 4 [temp_use]) + apply assumption+ + apply (erule_tac V = "sigmaa |= []F" in thin_rl) + apply (drule BoxSFI [temp_use]) + apply (erule_tac F = "TEMP <>Enabled (_g)" in dup_boxE) + apply (erule_tac F = "ACT N & [~B]_f" in dup_boxE) + apply (tactic "merge_temp_box_tac 1") + apply (erule DmdImpldup) + apply assumption + apply (auto simp: split_box_conj [try_rewrite] STL3 [try_rewrite] + SF_Box [try_rewrite] box_stp_act [try_rewrite]) + apply (force elim!: TLA2E [where P = P, temp_use]) + apply (rule STL2 [temp_use]) + apply (force simp: SF_def split_box_conj [try_rewrite] + elim!: mp InfImpl [temp_use] intro!: 3 [temp_use]) + done + +(* ------------------------------------------------------------------------- *) +(*** Liveness proofs by well-founded orderings ***) +(* ------------------------------------------------------------------------- *) +section "Well-founded orderings" + +lemma wf_leadsto: + assumes 1: "wf r" + and 2: "!!x. sigma |= F x ~> (G | (EX y. #((y,x):r) & F y)) " + shows "sigma |= F x ~> G" + apply (rule 1 [THEN wf_induct]) + apply (rule LatticeTriangle [temp_use]) + apply (rule 2) + apply (auto simp: leadsto_exists [try_rewrite]) + apply (case_tac "(y,x) :r") + apply force + apply (force simp: leadsto_def Init_simps intro!: necT [temp_use]) + done + +(* If r is well-founded, state function v cannot decrease forever *) +lemma wf_not_box_decrease: "!!r. wf r ==> |- [][ (v`, $v) : #r ]_v --> <>[][#False]_v" + apply clarsimp + apply (rule ccontr) + apply (subgoal_tac "sigma |= (EX x. v=#x) ~> #False") + apply (drule leadsto_false [temp_use, THEN iffD1, THEN STL2_gen [temp_use]]) + apply (force simp: Init_defs) + apply (clarsimp simp: leadsto_exists [try_rewrite] not_square [try_rewrite] more_temp_simps) + apply (erule wf_leadsto) + apply (rule ensures_simple [temp_use]) + apply (tactic "TRYALL atac") + apply (auto simp: square_def angle_def) + done + +(* "wf r ==> |- <>[][ (v`, $v) : #r ]_v --> <>[][#False]_v" *) +lemmas wf_not_dmd_box_decrease = + wf_not_box_decrease [THEN DmdImpl, unfolded more_temp_simps, standard] + +(* If there are infinitely many steps where v decreases, then there + have to be infinitely many non-stuttering steps where v doesn't decrease. +*) +lemma wf_box_dmd_decrease: + assumes 1: "wf r" + shows "|- []<>((v`, $v) : #r) --> []<><(v`, $v) ~: #r>_v" + apply clarsimp + apply (rule ccontr) + apply (simp add: not_angle [try_rewrite] more_temp_simps) + apply (drule 1 [THEN wf_not_dmd_box_decrease [temp_use]]) + apply (drule BoxDmdDmdBox [temp_use]) + apply assumption + apply (subgoal_tac "sigma |= []<> ((#False) ::action)") + apply force + apply (erule STL4E) + apply (rule DmdImpl) + apply (force intro: 1 [THEN wf_irrefl, temp_use]) + done + +(* In particular, for natural numbers, if n decreases infinitely often + then it has to increase infinitely often. +*) +lemma nat_box_dmd_decrease: "!!n::nat stfun. |- []<>(n` < $n) --> []<>($n < n`)" + apply clarsimp + apply (subgoal_tac "sigma |= []<><~ ((n`,$n) : #less_than) >_n") + apply (erule thin_rl) + apply (erule STL4E) + apply (rule DmdImpl) + apply (clarsimp simp: angle_def [try_rewrite]) + apply (rule wf_box_dmd_decrease [temp_use]) + apply (auto elim!: STL4E [temp_use] DmdImplE [temp_use]) + done + + +(* ------------------------------------------------------------------------- *) +(*** Flexible quantification over state variables ***) +(* ------------------------------------------------------------------------- *) +section "Flexible quantification" + +lemma aallI: + assumes 1: "basevars vs" + and 2: "(!!x. basevars (x,vs) ==> sigma |= F x)" + shows "sigma |= (AALL x. F x)" + by (auto simp: aall_def elim!: eexE [temp_use] intro!: 1 dest!: 2 [temp_use]) + +lemma aallE: "|- (AALL x. F x) --> F x" + apply (unfold aall_def) + apply clarsimp + apply (erule contrapos_np) + apply (force intro!: eexI [temp_use]) + done + +(* monotonicity of quantification *) +lemma eex_mono: + assumes 1: "sigma |= EEX x. F x" + and 2: "!!x. sigma |= F x --> G x" + shows "sigma |= EEX x. G x" + apply (rule unit_base [THEN 1 [THEN eexE]]) + apply (rule eexI [temp_use]) + apply (erule 2 [unfolded intensional_rews, THEN mp]) + done + +lemma aall_mono: + assumes 1: "sigma |= AALL x. F(x)" + and 2: "!!x. sigma |= F(x) --> G(x)" + shows "sigma |= AALL x. G(x)" + apply (rule unit_base [THEN aallI]) + apply (rule 2 [unfolded intensional_rews, THEN mp]) + apply (rule 1 [THEN aallE [temp_use]]) + done + +(* Derived history introduction rule *) +lemma historyI: + assumes 1: "sigma |= Init I" + and 2: "sigma |= []N" + and 3: "basevars vs" + and 4: "!!h. basevars(h,vs) ==> |- I & h = ha --> HI h" + and 5: "!!h s t. [| basevars(h,vs); N (s,t); h t = hb (h s) (s,t) |] ==> HN h (s,t)" + shows "sigma |= EEX h. Init (HI h) & [](HN h)" + apply (rule history [temp_use, THEN eexE]) + apply (rule 3) + apply (rule eexI [temp_use]) + apply clarsimp + apply (rule conjI) + prefer 2 + apply (insert 2) + apply (tactic "merge_box_tac 1") + apply (force elim!: STL4E [temp_use] 5 [temp_use]) + apply (insert 1) + apply (force simp: Init_defs elim!: 4 [temp_use]) + done + +(* ---------------------------------------------------------------------- + example of a history variable: existence of a clock +*) + +lemma "|- EEX h. Init(h = #True) & [](h` = (~$h))" + apply (rule tempI) + apply (rule historyI) + apply (force simp: Init_defs intro!: unit_base [temp_use] necT [temp_use])+ + done + +end +