lh-l4v/lib/sep_algebra/Sep_Rotate.thy

39 lines
895 B
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_Rotate
imports Sep_Select
begin
(* We can build rotators on the basis of selectors *)
2019-05-24 02:20:31 +00:00
ML \<open>
2014-07-14 19:32:44 +00:00
(* generic rotator *)
fun range lo hi =
let
fun r lo = if lo > hi then [] else lo::r (lo+1)
in r lo end
2017-07-12 05:13:51 +00:00
fun rotator lens tactic ctxt i st =
let
val len = case Seq.pull ((lens THEN' resolve0_tac [@{thm iffI}]) i st) of
NONE => 0
| SOME (thm, _) => conj_length ctxt (Thm.cprem_of thm i)
val nums = range 1 len
val selector = sep_select_tactic lens
val tac' = map (fn x => selector [x] ctxt THEN' tactic) nums
in
(selector [1] ctxt THEN' FIRST' tac') i st
end
fun rotator' ctxt lens tactic = rotator lens tactic ctxt
fun sep_apply_tactic ctxt lens_tac thms = lens_tac THEN' eresolve_tac ctxt thms
2019-05-24 02:20:31 +00:00
\<close>
2014-07-14 19:32:44 +00:00
end