author | nipkow |
Thu, 26 Jun 2025 17:30:33 +0200 | |
changeset 82772 | 59b937edcff8 |
parent 80693 | e451d6230535 |
permissions | -rw-r--r-- |
23489 | 1 |
(* Title: HOL/Tools/Qelim/qelim.ML |
23904 | 2 |
Author: Amine Chaieb, TU Muenchen |
23466 | 3 |
|
23904 | 4 |
Generic quantifier elimination conversions for HOL formulae. |
23466 | 5 |
*) |
6 |
||
7 |
signature QELIM = |
|
8 |
sig |
|
80692 | 9 |
val gen_qelim_conv: Proof.context -> conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a -> |
10 |
('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv |
|
11 |
val standard_qelim_conv: Proof.context -> |
|
12 |
(cterm list -> conv) -> (cterm list -> conv) -> |
|
13 |
(cterm list -> conv) -> conv |
|
23466 | 14 |
end; |
15 |
||
23489 | 16 |
structure Qelim: QELIM = |
23466 | 17 |
struct |
18 |
||
61075 | 19 |
fun gen_qelim_conv ctxt precv postcv simpex_conv ins env atcv ncv qcv = |
80692 | 20 |
let |
21 |
fun conv env p = |
|
80693 | 22 |
(case Thm.term_of p of |
23 |
\<^Const_>\<open>conj for _ _\<close> => Conv.binop_conv (conv env) p |
|
24 |
| \<^Const_>\<open>disj for _ _\<close> => Conv.binop_conv (conv env) p |
|
25 |
| \<^Const_>\<open>implies for _ _\<close> => Conv.binop_conv (conv env) p |
|
26 |
| \<^Const_>\<open>HOL.eq _ for _ _\<close> => Conv.binop_conv (conv env) p |
|
27 |
| \<^Const_>\<open>Not for _\<close> => Conv.arg_conv (conv env) p |
|
28 |
| \<^Const_>\<open>Ex _ for \<open>Abs (s, _, _)\<close>\<close> => |
|
80692 | 29 |
let |
30 |
val (e,p0) = Thm.dest_comb p |
|
31 |
val (x,p') = Thm.dest_abs_global p0 |
|
32 |
val env' = ins x env |
|
33 |
val th = |
|
34 |
Thm.abstract_rule s x ((conv env' then_conv ncv env') p') |
|
35 |
|> Drule.arg_cong_rule e |
|
36 |
val th' = simpex_conv (Thm.rhs_of th) |
|
37 |
val (_, r) = Thm.dest_equals (Thm.cprop_of th') |
|
38 |
in |
|
39 |
if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th)) |
|
40 |
else Thm.transitive (Thm.transitive th th') (conv env r) |
|
41 |
end |
|
80693 | 42 |
| \<^Const_>\<open>Ex _ for _\<close> => (Thm.eta_long_conversion then_conv conv env) p |
43 |
| \<^Const_>\<open>All _ for _\<close> => |
|
80692 | 44 |
let |
45 |
val allT = Thm.ctyp_of_cterm (Thm.dest_fun p) |
|
46 |
val T = Thm.dest_ctyp0 (Thm.dest_ctyp0 allT) |
|
80693 | 47 |
val P = Thm.dest_arg p |
48 |
val th = \<^instantiate>\<open>'a = T and P in lemma "\<forall>x::'a. P x \<equiv> \<nexists>x. \<not> P x" by simp\<close> |
|
80692 | 49 |
in Thm.transitive th (conv env (Thm.rhs_of th)) end |
80693 | 50 |
| _ => atcv env p) |
80692 | 51 |
in precv then_conv (conv env) then_conv postcv end |
52 |
||
23489 | 53 |
|
23904 | 54 |
(* Instantiation of some parameter for most common cases *) |
23489 | 55 |
|
23904 | 56 |
local |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
57 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
58 |
val ss = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
59 |
simpset_of |
69593 | 60 |
(put_simpset HOL_basic_ss \<^context> |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
61 |
addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib}); |
80692 | 62 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
63 |
fun pcv ctxt = Simplifier.rewrite (put_simpset ss ctxt) |
23466 | 64 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
65 |
in |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
66 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
67 |
fun standard_qelim_conv ctxt atcv ncv qcv p = |
74269 | 68 |
let |
69 |
val pcv = pcv ctxt |
|
74274 | 70 |
val env = Cterms.list_set_rev (Cterms.build (Drule.add_frees_cterm p)) |
74269 | 71 |
in gen_qelim_conv ctxt pcv pcv pcv cons env atcv ncv qcv p end |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
45654
diff
changeset
|
72 |
|
23904 | 73 |
end; |
23466 | 74 |
|
75 |
end; |