lh-l4v/lib/sep_algebra/Sep_Select.thy

66 lines
1.9 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_Select
imports Separation_Algebra
begin
ML_file "sep_tactics.ML"
2019-06-05 10:18:48 +00:00
ML\<open>
structure SepSelect_Rules = Named_Thms (
val name = @{binding "sep_select"}
val description = "sep_select rules"
)
2019-06-05 10:18:48 +00:00
\<close>
2014-07-14 19:32:44 +00:00
setup SepSelect_Rules.setup
2019-06-05 10:18:48 +00:00
ML \<open>
2014-07-14 19:32:44 +00:00
structure SepSelectAsm_Rules = Named_Thms (
val name = @{binding "sep_select_asm"}
val description = "sep_select_asm rules"
)
2019-06-05 10:18:48 +00:00
\<close>
2014-07-14 19:32:44 +00:00
setup SepSelectAsm_Rules.setup
2019-06-05 10:18:48 +00:00
ML \<open>
2014-07-14 19:32:44 +00:00
fun sep_selects_tactic ns ctxt =
2016-03-28 08:13:28 +00:00
sep_select_tactic (resolve_tac ctxt (SepSelect_Rules.get ctxt)) ns ctxt
2017-07-12 05:13:51 +00:00
2014-07-14 19:32:44 +00:00
fun sep_select_asms_tactic ns ctxt =
2017-07-12 05:13:51 +00:00
sep_select_tactic (dresolve_tac ctxt (SepSelectAsm_Rules.get ctxt)) ns ctxt
2019-06-05 10:18:48 +00:00
\<close>
2014-07-14 19:32:44 +00:00
2019-06-05 10:18:48 +00:00
method_setup sep_select_asm = \<open>
2017-07-12 05:13:51 +00:00
Scan.lift (Scan.repeat Parse.int) >>
2016-03-28 08:13:28 +00:00
(fn ns => fn ctxt => SIMPLE_METHOD' (sep_select_asms_tactic ns ctxt))
2019-06-05 10:18:48 +00:00
\<close> "Reorder assumptions"
2014-07-14 19:32:44 +00:00
2019-06-05 10:18:48 +00:00
method_setup sep_select = \<open>
2017-07-12 05:13:51 +00:00
Scan.lift (Scan.repeat Parse.int) >>
2016-03-28 08:13:28 +00:00
(fn ns => fn ctxt => SIMPLE_METHOD' (sep_selects_tactic ns ctxt))
2019-06-05 10:18:48 +00:00
\<close> "Reorder conclusions"
2014-07-14 19:32:44 +00:00
2016-03-28 08:13:28 +00:00
lemma sep_eq [sep_select]: "(\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> T s \<Longrightarrow> (P \<and>* R) s" by clarsimp
lemma sep_asm_eq [sep_select_asm]: "(P \<and>* R) s \<Longrightarrow> (\<And>s. T s = (P \<and>* R) s) \<Longrightarrow> T s" by clarsimp
2014-07-14 19:32:44 +00:00
2019-06-05 10:18:48 +00:00
ML \<open>
(* export method form of these two for use outisde this entry *)
2017-07-12 05:13:51 +00:00
fun sep_select_method lens ns ctxt =
2017-07-12 05:13:51 +00:00
SIMPLE_METHOD' (sep_select_tactic lens ns ctxt)
fun sep_select_generic_method asm thms ns ctxt =
sep_select_method (if asm then dresolve_tac ctxt thms else resolve_tac ctxt thms) ns ctxt
2019-06-05 10:18:48 +00:00
\<close>
2019-06-05 10:18:48 +00:00
method_setup sep_select_gen = \<open>
Attrib.thms --| Scan.lift Args.colon -- Scan.lift (Scan.repeat Parse.int) -- Scan.lift (Args.mode "asm") >>
(fn ((lens, ns), asm) => sep_select_generic_method asm lens ns)
2019-06-05 10:18:48 +00:00
\<close>
2014-07-14 19:32:44 +00:00
end