| author | desharna | 
| Mon, 20 Mar 2023 15:02:17 +0100 | |
| changeset 77697 | f35cbb4da88a | 
| parent 74525 | c960bfcb91db | 
| child 80692 | f65b65814b81 | 
| 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 | |
| 61075 | 9 | val gen_qelim_conv: Proof.context -> conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a -> | 
| 41453 | 10 |   ('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 11 | val standard_qelim_conv: Proof.context -> | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 12 | (cterm list -> conv) -> (cterm list -> conv) -> | 
| 41453 | 13 | (cterm list -> conv) -> conv | 
| 23466 | 14 | end; | 
| 15 | ||
| 23489 | 16 | structure Qelim: QELIM = | 
| 23466 | 17 | struct | 
| 18 | ||
| 23489 | 19 | val all_not_ex = mk_meta_eq @{thm "all_not_ex"};
 | 
| 20 | ||
| 61075 | 21 | fun gen_qelim_conv ctxt precv postcv simpex_conv ins env atcv ncv qcv = | 
| 23489 | 22 | let | 
| 23466 | 23 | fun conv env p = | 
| 59582 | 24 | case Thm.term_of p of | 
| 41453 | 25 | Const(s,T)$_$_ => | 
| 23904 | 26 | if domain_type T = HOLogic.boolT | 
| 69593 | 27 | andalso member (op =) [\<^const_name>\<open>HOL.conj\<close>, \<^const_name>\<open>HOL.disj\<close>, | 
| 28 | \<^const_name>\<open>HOL.implies\<close>, \<^const_name>\<open>HOL.eq\<close>] s | |
| 41453 | 29 | then Conv.binop_conv (conv env) p | 
| 23904 | 30 | else atcv env p | 
| 69593 | 31 | | Const(\<^const_name>\<open>Not\<close>,_)$_ => Conv.arg_conv (conv env) p | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 32 | | Const(\<^const_name>\<open>Ex\<close>,_)$Abs (s, _, _) => | 
| 23489 | 33 | let | 
| 23466 | 34 | val (e,p0) = Thm.dest_comb p | 
| 74525 
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
 wenzelm parents: 
74518diff
changeset | 35 | val (x,p') = Thm.dest_abs_global p0 | 
| 23466 | 36 | val env' = ins x env | 
| 37 | val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p') | |
| 38 | |> Drule.arg_cong_rule e | |
| 39 | val th' = simpex_conv (Thm.rhs_of th) | |
| 59582 | 40 | val (_, r) = Thm.dest_equals (Thm.cprop_of th') | 
| 23593 | 41 | in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th)) | 
| 23466 | 42 | else Thm.transitive (Thm.transitive th th') (conv env r) end | 
| 69593 | 43 | | Const(\<^const_name>\<open>Ex\<close>,_)$ _ => (Thm.eta_long_conversion then_conv conv env) p | 
| 70158 | 44 | | Const(\<^const_name>\<open>All\<close>, _)$_ => | 
| 23489 | 45 | let | 
| 70158 | 46 | val allT = Thm.ctyp_of_cterm (Thm.dest_fun p) | 
| 70159 | 47 | val T = Thm.dest_ctyp0 (Thm.dest_ctyp0 allT) | 
| 23466 | 48 | val p = Thm.dest_arg p | 
| 60801 | 49 | val th = Thm.instantiate' [SOME T] [SOME p] all_not_ex | 
| 36945 | 50 | in Thm.transitive th (conv env (Thm.rhs_of th)) | 
| 23466 | 51 | end | 
| 52 | | _ => atcv env p | |
| 53 | in precv then_conv (conv env) then_conv postcv end | |
| 23489 | 54 | |
| 23904 | 55 | (* Instantiation of some parameter for most common cases *) | 
| 23489 | 56 | |
| 23904 | 57 | local | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 58 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 59 | val ss = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 60 | simpset_of | 
| 69593 | 61 | (put_simpset HOL_basic_ss \<^context> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 62 |     addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib});
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
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: 
45654diff
changeset | 65 | in | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
changeset | 66 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
45654diff
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: 
45654diff
changeset | 72 | |
| 23904 | 73 | end; | 
| 23466 | 74 | |
| 75 | end; |