| 26582 |      1 | (*  Title:      Tools/atomize_elim.ML
 | 
|  |      2 |     ID:         $Id$
 | 
|  |      3 |     Author:     Alexander Krauss, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Turn elimination rules into atomic expressions in the object logic.
 | 
|  |      6 | *)
 | 
|  |      7 | 
 | 
|  |      8 | signature ATOMIZE_ELIM =
 | 
|  |      9 | sig
 | 
|  |     10 |   val declare_atomize_elim : attribute  
 | 
|  |     11 | 
 | 
| 27571 |     12 |   val atomize_elim_conv : Proof.context -> conv
 | 
|  |     13 |   val full_atomize_elim_conv : Proof.context -> conv
 | 
| 26582 |     14 | 
 | 
|  |     15 |   val atomize_elim_tac : Proof.context -> int -> tactic
 | 
|  |     16 | 
 | 
|  |     17 |   val setup : theory -> theory
 | 
|  |     18 | end
 | 
|  |     19 | 
 | 
|  |     20 | structure AtomizeElim : ATOMIZE_ELIM =
 | 
|  |     21 | struct
 | 
|  |     22 | 
 | 
|  |     23 | (* theory data *)
 | 
| 27571 |     24 | structure AtomizeElimData = NamedThmsFun(
 | 
|  |     25 |   val name = "atomize_elim";
 | 
|  |     26 |   val description = "atomize_elim rewrite rule";
 | 
| 26582 |     27 | );
 | 
|  |     28 | 
 | 
| 27571 |     29 | val declare_atomize_elim = AtomizeElimData.add
 | 
| 26582 |     30 | 
 | 
|  |     31 | (* More conversions: *)
 | 
|  |     32 | local open Conv in
 | 
|  |     33 | 
 | 
|  |     34 | (* Compute inverse permutation *)
 | 
|  |     35 | fun invert_perm pi =
 | 
|  |     36 |       (pi @ ((0 upto (fold (curry Int.max) pi 0)) \\ pi))
 | 
|  |     37 |            |> map_index I
 | 
|  |     38 |            |> sort (int_ord o pairself snd)
 | 
|  |     39 |            |> map fst
 | 
|  |     40 | 
 | 
|  |     41 | (* rearrange_prems as a conversion *)
 | 
|  |     42 | fun rearrange_prems_conv pi ct =
 | 
|  |     43 |     let
 | 
|  |     44 |       val pi' = 0 :: map (curry op + 1) pi
 | 
| 27571 |     45 |       val fwd  = Thm.trivial ct
 | 
| 26582 |     46 |                   |> Drule.rearrange_prems pi'
 | 
| 27571 |     47 |       val back = Thm.trivial (snd (Thm.dest_implies (cprop_of fwd)))
 | 
| 26582 |     48 |                   |> Drule.rearrange_prems (invert_perm pi')
 | 
| 27571 |     49 |     in Thm.equal_intr fwd back end
 | 
| 26582 |     50 | 
 | 
|  |     51 | 
 | 
|  |     52 | (* convert !!thesis. (!!x y z. ... => thesis) ==> ... 
 | 
|  |     53 |                      ==> (!!a b c. ... => thesis)
 | 
|  |     54 |                      ==> thesis
 | 
|  |     55 | 
 | 
|  |     56 |    to
 | 
|  |     57 | 
 | 
|  |     58 |    (EX x y z. ...) | ... | (EX a b c. ...)
 | 
|  |     59 | *)
 | 
| 27571 |     60 | fun atomize_elim_conv ctxt ct =
 | 
| 26582 |     61 |     (forall_conv (K (prems_conv ~1 ObjectLogic.atomize_prems)) ctxt
 | 
| 27571 |     62 |     then_conv MetaSimplifier.rewrite true (AtomizeElimData.get ctxt)
 | 
| 26582 |     63 |     then_conv (fn ct' => if can ObjectLogic.dest_judgment ct'
 | 
|  |     64 |                          then all_conv ct'
 | 
|  |     65 |                          else raise CTERM ("atomize_elim", [ct', ct])))
 | 
|  |     66 |     ct
 | 
|  |     67 | 
 | 
|  |     68 | 
 | 
|  |     69 | (* Move the thesis-quantifier inside over other quantifiers and unrelated prems *)
 | 
|  |     70 | fun thesis_miniscope_conv inner_cv ctxt =
 | 
|  |     71 |     let
 | 
|  |     72 |       (* check if the outermost quantifier binds the conclusion *)
 | 
| 27571 |     73 |       fun is_forall_thesis t = 
 | 
|  |     74 |           case Logic.strip_assums_concl t of
 | 
|  |     75 |             (_ $ Bound i) => i = length (Logic.strip_params t) - 1
 | 
|  |     76 |           | _ => false
 | 
| 26582 |     77 | 
 | 
|  |     78 |       (* move unrelated assumptions out of the quantification *)
 | 
|  |     79 |       fun movea_conv ctxt ct =
 | 
|  |     80 |           let
 | 
|  |     81 |             val _ $ Abs (_, _, b) = term_of ct
 | 
|  |     82 |             val idxs = fold_index (fn (i, t) => not (loose_bvar1 (t, 0)) ? cons i)
 | 
|  |     83 |                        (Logic.strip_imp_prems b) []
 | 
|  |     84 |                        |> rev
 | 
|  |     85 | 
 | 
|  |     86 |             fun skip_over_assm cv = rewr_conv (Thm.symmetric Drule.norm_hhf_eq)
 | 
|  |     87 |                                     then_conv (implies_concl_conv cv)
 | 
|  |     88 |           in
 | 
|  |     89 |             (forall_conv (K (rearrange_prems_conv idxs)) ctxt
 | 
|  |     90 |              then_conv (funpow (length idxs) skip_over_assm (inner_cv ctxt)))
 | 
|  |     91 |             ct
 | 
|  |     92 |           end
 | 
|  |     93 | 
 | 
|  |     94 |       (* move current quantifier to the right *)
 | 
|  |     95 |       fun moveq_conv ctxt = 
 | 
| 27571 |     96 |           (rewr_conv @{thm swap_params} then_conv (forall_conv (moveq_conv o snd) ctxt))
 | 
| 26582 |     97 |           else_conv (movea_conv ctxt)
 | 
|  |     98 | 
 | 
|  |     99 |       fun ms_conv ctxt ct = 
 | 
|  |    100 |           if is_forall_thesis (term_of ct)
 | 
|  |    101 |           then moveq_conv ctxt ct
 | 
|  |    102 |           else (implies_concl_conv (ms_conv ctxt)
 | 
|  |    103 |             else_conv
 | 
| 27571 |    104 |             (forall_conv (ms_conv o snd) ctxt))
 | 
| 26582 |    105 |             ct
 | 
|  |    106 |     in
 | 
|  |    107 |       ms_conv ctxt 
 | 
|  |    108 |     end
 | 
|  |    109 | 
 | 
| 27571 |    110 | val full_atomize_elim_conv = thesis_miniscope_conv atomize_elim_conv
 | 
| 26582 |    111 | 
 | 
|  |    112 | end
 | 
|  |    113 | 
 | 
|  |    114 | 
 | 
| 27571 |    115 | fun atomize_elim_tac ctxt = SUBGOAL (fn (subg, i) =>
 | 
| 26582 |    116 |     let
 | 
| 27571 |    117 |       val thy = ProofContext.theory_of ctxt
 | 
|  |    118 |       val _ $ thesis = Logic.strip_assums_concl subg
 | 
| 26582 |    119 |                        
 | 
|  |    120 |       (* Introduce a quantifier first if the thesis is not bound *)
 | 
|  |    121 |       val quantify_thesis = 
 | 
|  |    122 |           if is_Bound thesis then all_tac
 | 
|  |    123 |           else let
 | 
| 27571 |    124 |               val cthesis = cterm_of thy thesis
 | 
| 26582 |    125 |               val rule = instantiate' [SOME (ctyp_of_term cthesis)] [NONE, SOME cthesis] 
 | 
|  |    126 |                          @{thm meta_spec}
 | 
|  |    127 |             in
 | 
| 27571 |    128 |               compose_tac (false, rule, 1) i
 | 
| 26582 |    129 |             end
 | 
|  |    130 |     in
 | 
|  |    131 |       quantify_thesis
 | 
| 27571 |    132 |       THEN (CONVERSION (full_atomize_elim_conv ctxt) i)
 | 
| 26582 |    133 |     end)
 | 
|  |    134 | 
 | 
|  |    135 | val setup = Method.add_methods
 | 
|  |    136 |   [("atomize_elim", Method.ctxt_args (Method.SIMPLE_METHOD' o atomize_elim_tac),
 | 
|  |    137 |     "convert obtains statement to atomic object logic goal")]
 | 
| 27571 |    138 |   #> AtomizeElimData.setup
 | 
| 26582 |    139 | 
 | 
|  |    140 | end
 |