(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory Sep_ImpI imports Sep_Provers Sep_Cancel_Set Sep_Tactic_Helpers begin lemma sep_wand_lens: "(\s. T s = Q s) \ ((P \* T) \* R) s \ ((P \* Q) \* R) s" apply (sep_erule_full refl_imp) apply (clarsimp simp: sep_impl_def) done lemma sep_wand_lens': "(\s. T s = Q s) \ ((T \* P) \* R) s \ ((Q \* P) \* R) s" apply (sep_erule_full refl_imp, erule sep_curry[rotated]) apply (clarsimp) apply (erule sep_mp) done (* Removing wands from the conclusions *) ML \ fun sep_wand_lens ctxt = resolve_tac ctxt[@{thm sep_wand_lens}] fun sep_wand_lens' ctxt = resolve_tac ctxt [@{thm sep_wand_lens'}] fun sep_wand_rule_tac tac ctxt = let val r = rotator' ctxt in tac |> r (sep_wand_lens' ctxt) |> r (sep_wand_lens ctxt) |> r (sep_select ctxt) end fun sep_wand_rule_tac' thms ctxt = let val r = rotator' ctxt in eresolve_tac ctxt thms |> r (sep_wand_lens ctxt) |> r (sep_select ctxt) |> r (sep_asm_select ctxt) end fun sep_wand_rule_method thms ctxt = SIMPLE_METHOD' (sep_wand_rule_tac thms ctxt) fun sep_wand_rule_method' thms ctxt = SIMPLE_METHOD' (sep_wand_rule_tac' thms ctxt) \ lemma sep_wand_match: "(\s. Q s \ Q' s) \ (R \* R') s ==> (Q \* R \* Q' \* R') s" apply (erule sep_curry[rotated]) apply (sep_select_asm 1 3) apply (sep_drule (direct) sep_mp_frame) apply (sep_erule_full refl_imp, clarsimp) done lemma sep_wand_trivial: "(\s. Q s \ Q' s) \ R' s ==> (Q \* Q' \* R') s" apply (erule sep_curry[rotated]) apply (sep_erule_full refl_imp) apply (clarsimp) done lemma sep_wand_collapse: "(P \* Q \* R) s \ (P \* Q \* R) s " apply (erule sep_curry[rotated])+ apply (clarsimp simp: sep_conj_assoc) apply (erule sep_mp) done lemma sep_wand_match_less_safe: assumes drule: " \s. (Q' \* R) s \ ((P \* R') \* Q' \* R'' ) s " shows "(Q' \* R) s \ (\s. Q' s \ Q s) \ ((P \* Q \* R') \* R'') s" apply (drule drule) apply (sep_erule_full refl_imp) apply (erule sep_conj_sep_impl) apply (clarsimp simp: sep_conj_assoc) apply (sep_select_asm 1 3) apply (sep_drule (direct) sep_mp_frame, sep_erule_full refl_imp) apply (clarsimp) done ML \ fun sep_match_trivial_tac ctxt = let fun flip f a b = f b a val sep_cancel = flip (sep_apply_tactic ctxt) (SepCancel_Rules.get ctxt |> rev) fun f x = x |> rotate_prems ~1 |> (fn x => [x]) |> eresolve0_tac |> sep_cancel val sep_thms = map f [@{thm sep_wand_trivial}, @{thm sep_wand_match}] in sep_wand_rule_tac (resolve0_tac [@{thm sep_rule}] THEN' FIRST' sep_thms) ctxt end fun sep_safe_method ctxt = SIMPLE_METHOD' (sep_match_trivial_tac ctxt) \ method_setup sep_safe = \ Scan.succeed (sep_safe_method) \ end