(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Corres_Adjust_Preconds imports "Corres_UL" begin text \ Gadget for adjusting preconditions in a corres rule or similar. Probably only useful for predicates with two or more related preconditions, such as corres. Used to do some_corres_rule[adj_corres some_intro_rule], given e.g. some_intro_rule: @{prop "(s, t) \ sr \ P s \ Q t"} Will apply this rule to solve @{prop "Q t"} components in either precondition or any sub-conjunct, and will then try to put the assumptions @{prop "P s"}, @{prop "(s, t) \ sr"} into the right places. The premises of the rule can be in any given order. Concrete example at the bottom. \ named_theorems corres_adjust_precond_structures locale corres_adjust_preconds begin text \Worker predicates. Broadly speaking, a goal of the form "preconds ?A ?B ?C ?D ==> P" expects to establish P by instantiating ?A, or failing that ?B, etc. A goal of the form finalise_preconds A exists to make sure that schematic conjuncts of A are resolved to True.\ definition preconds :: "bool \ bool \ bool \ bool \ bool" where "preconds A B C D = (A \ B \ C \ D)" definition finalise_preconds :: "bool \ bool" where "finalise_preconds A = True" text \Use a precond directly to establish goal.\ lemma consume_preconds: "preconds A True True True \ A" "preconds True B True True \ B" "preconds True True C True \ C" "preconds True True True D \ D" by (simp_all add: preconds_def) lemmas consume_preconds_True = consume_preconds(1)[where A=True] text \For use as a drule, split a set of schematic preconds to give two sets that can be instantiated separately.\ lemma split_preconds_left: "preconds (A \ A') (B \ B') (C \ C') (D \ D') \ preconds A B C D" "preconds (A \ A') (B \ B') (C \ C') True \ preconds A B C True" "preconds (A \ A') (B \ B') True True \ preconds A B True True" "preconds (A \ A') True True True \ preconds A True True True" by (simp_all add: preconds_def) lemma split_preconds_right: "preconds (A \ A') (B \ B') (C \ C') (D \ D') \ preconds A' B' C' D'" "preconds (A \ A') (B \ B') (C \ C') True \ preconds A' B' C' True" "preconds (A \ A') (B \ B') True True \ preconds A' B' True True" "preconds (A \ A') True True True \ preconds A' True True True" by (simp_all add: preconds_def) text \For use as an erule. Initiate the precond process, creating a finalise goal to be solved later.\ lemma preconds_goal_initiate: "preconds A B C D \ (preconds A B C D \ Q) \ finalise_preconds (A \ B \ C \ D) \ Q" by simp text \Finalise preconds, trying to replace conjuncts with True if they are not yet instantiated.\ lemma finalise_preconds: "finalise_preconds True" "finalise_preconds A \ finalise_preconds B \ finalise_preconds (A \ B)" "finalise_preconds X" by (simp_all add: finalise_preconds_def) text \Shape of the whole process for application to regular corres goals.\ lemma corres_adjust_pre: "corres_underlying R nf nf' rs P Q f f' \ (\s s'. (s, s') \ R \ preconds (P1 s) (Q1 s') True True \ P s) \ (\s s'. (s, s') \ R \ preconds (Q2 s') (P2 s) True True \ Q s') \ corres_underlying R nf nf' rs (\s. P1 s \ P2 s) (\s'. Q1 s' \ Q2 s') f f'" apply (erule stronger_corres_guard_imp) apply (simp add: preconds_def)+ done ML \ structure Corres_Adjust_Preconds = struct val def_intros = @{thms conjI pred_conjI} (* apply an intro rule, splitting preconds assumptions to provide unique assumptions for each goal. *) fun intro_split ctxt intros i = ((resolve_tac ctxt intros THEN_ALL_NEW (TRY o assume_tac ctxt)) THEN_ALL_NEW (fn j => (EVERY (replicate (j - i) (dresolve_tac ctxt @{thms split_preconds_left} j))) THEN dresolve_tac ctxt @{thms split_preconds_right} j)) i fun handle_preconds ctxt intros = TRY o (eresolve_tac ctxt [@{thm preconds_goal_initiate}] THEN' REPEAT_ALL_NEW (eresolve_tac ctxt @{thms consume_preconds_True} ORELSE' intro_split ctxt (intros @ def_intros) ORELSE' eresolve_tac ctxt @{thms consume_preconds}) THEN' REPEAT_ALL_NEW (resolve_tac ctxt @{thms finalise_preconds}) ) fun mk_adj_preconds ctxt intros rule = let val xs = [rule] RL (Named_Theorems.get ctxt @{named_theorems corres_adjust_precond_structures}) val x = case xs of [] => raise THM ("no unifier with corres_adjust_precond_structures", 1, [rule]) | xs => hd xs in x |> ALLGOALS (handle_preconds ctxt intros) |> Seq.hd |> Simplifier.simplify (clear_simpset ctxt addsimps @{thms conj_assoc simp_thms(21-22)}) end val setup = Attrib.setup @{binding "adj_corres"} ((Attrib.thms -- Args.context) >> (fn (intros, ctxt) => Thm.rule_attribute [] (K (mk_adj_preconds ctxt intros)))) "use argument theorems to adjust a corres theorem." end \ end declare corres_adjust_preconds.corres_adjust_pre[corres_adjust_precond_structures] setup Corres_Adjust_Preconds.setup experiment begin definition test_sr :: "(nat \ nat) set" where "test_sr = {(x, y). y = 2 * x}" lemma test_corres: "corres_underlying test_sr nf nf' dc (\x. x < 40) (\y. y < 30 \ y = 6) (modify (\x. x + 2)) (modify (\y. 10))" by (simp add: corres_underlying_def simpler_modify_def test_sr_def) lemma test_adj_precond: "(x, y) \ test_sr \ x = 3 \ y = 6" by (simp add: test_sr_def) ML \ Corres_Adjust_Preconds.mk_adj_preconds @{context} [@{thm test_adj_precond}] @{thm test_corres} \ lemma foo_adj_corres: "corres_underlying test_sr nf nf' dc (\s. s < 40 \ s = 3) (\s'. s' < 30) (modify (\x. x + 2)) (modify (\y. 10))" by (rule test_corres[adj_corres test_adj_precond]) text \A more general demo of what it does.\ lemma assumes my_corres: "corres_underlying my_sr nf nf' rvrel P Q a c" assumes my_adj: "\s s'. (s,s') \ my_sr \ P s \ Q s'" shows "corres_underlying my_sr nf nf' rvrel (\s. P s \ P s) (\s'. True) a c" apply (rule my_corres[adj_corres my_adj]) done end end