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