lh-l4v/lib/sep_algebra/Sep_MP.thy

54 lines
2.2 KiB
Plaintext
Raw Normal View History

2014-07-14 19:32:44 +00:00
(*
2020-03-09 06:18:30 +00:00
* Copyright 2020, Data61, CSIRO (ABN 41 687 119 230)
2014-07-14 19:32:44 +00:00
*
2020-03-09 06:18:30 +00:00
* SPDX-License-Identifier: BSD-2-Clause
2014-07-14 19:32:44 +00:00
*)
theory Sep_MP
imports Sep_Tactic_Helpers Sep_Provers Sep_Cancel_Set
begin
lemma sep_mp_gen: "((Q \<longrightarrow>* R) \<and>* Q') s \<Longrightarrow> (\<And>s. Q' s \<Longrightarrow> Q s) \<Longrightarrow> R s"
by (clarsimp simp: sep_conj_def sep_impl_def)
lemma sep_mp_frame_gen: "\<lbrakk>((Q \<longrightarrow>* R) \<and>* Q' \<and>* R') s; (\<And>s. Q' s \<Longrightarrow> Q s)\<rbrakk> \<Longrightarrow> (R \<and>* R') s"
by (metis sep_conj_left_commute sep_globalise sep_mp_frame)
2017-07-12 05:13:51 +00:00
lemma sep_impl_simpl:
2014-07-14 19:32:44 +00:00
"(P \<and>* Q \<longrightarrow>* R) s \<Longrightarrow> (P \<longrightarrow>* Q \<longrightarrow>* R) s"
apply (erule sep_conj_sep_impl)
apply (erule sep_conj_sep_impl)
apply (clarsimp simp: sep_conj_assoc)
apply (erule sep_mp)
done
lemma sep_wand_frame_lens: "((P \<longrightarrow>* Q) \<and>* R) s \<Longrightarrow> (\<And>s. T s = R s) ==> ((P \<longrightarrow>* Q) \<and>* T) s"
by (metis sep_conj_commute sep_conj_impl1)
2019-06-05 10:18:48 +00:00
ML \<open>
2017-07-12 05:13:51 +00:00
fun sep_wand_frame_drule ctxt =
2015-05-14 09:40:55 +00:00
let val lens = dresolve_tac ctxt [@{thm sep_wand_frame_lens}]
val lens' = dresolve_tac ctxt [@{thm sep_asm_eq}]
2014-07-14 19:32:44 +00:00
val r = rotator' ctxt
val sep_cancel_thms = rev (SepCancel_Rules.get ctxt)
2015-05-14 09:40:55 +00:00
in sep_apply_tactic ctxt (dresolve_tac ctxt [@{thm sep_mp_frame_gen}]) sep_cancel_thms |> r lens |> r lens'
2017-07-12 05:13:51 +00:00
end;
2014-07-14 19:32:44 +00:00
fun sep_mp_solver ctxt =
let val sep_mp = sep_apply_tactic ctxt (dresolve0_tac [@{thm sep_mp_gen}]) ((rev o SepCancel_Rules.get) ctxt)
2014-07-14 19:32:44 +00:00
val taclist = [sep_drule_comb_tac false [@{thm sep_empty_imp}] ctxt,
sep_drule_tac sep_mp ctxt,
2015-05-14 09:40:55 +00:00
sep_drule_tac (sep_drule_tactic ctxt [@{thm sep_impl_simpl}]) ctxt,
2014-07-14 19:32:44 +00:00
sep_wand_frame_drule ctxt ]
val check = DETERM o (sep_drule_tac (sep_select_tactic (dresolve0_tac [@{thm sep_wand_frame_lens}]) [1] ctxt) ctxt)
2017-07-12 05:13:51 +00:00
2014-07-14 19:32:44 +00:00
in CHANGED_PROP o (check THEN_ALL_NEW REPEAT_ALL_NEW ( FIRST' taclist) )
end;
2017-07-12 05:13:51 +00:00
val sep_mp_method = SIMPLE_METHOD' o sep_mp_solver
2019-06-05 10:18:48 +00:00
\<close>
2017-07-12 05:13:51 +00:00
2019-06-05 10:18:48 +00:00
method_setup sep_mp = \<open>Scan.succeed sep_mp_method\<close>
2015-05-14 09:40:55 +00:00
2014-07-14 19:32:44 +00:00
end