| author | traytel | 
| Fri, 15 Mar 2013 10:08:23 +0100 | |
| changeset 51430 | e96447ea13c9 | 
| parent 45654 | cf10bde35973 | 
| child 51717 | 9e7d1c139569 | 
| 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 | |
| 41453 | 9 | val gen_qelim_conv: conv -> conv -> conv -> (cterm -> 'a -> 'a) -> 'a -> | 
| 10 |   ('a -> conv) -> ('a -> conv) -> ('a -> conv) -> conv
 | |
| 11 | val standard_qelim_conv: (cterm list -> conv) -> (cterm list -> conv) -> | |
| 12 | (cterm list -> conv) -> conv | |
| 23466 | 13 | end; | 
| 14 | ||
| 23489 | 15 | structure Qelim: QELIM = | 
| 23466 | 16 | struct | 
| 17 | ||
| 23489 | 18 | val all_not_ex = mk_meta_eq @{thm "all_not_ex"};
 | 
| 19 | ||
| 23524 | 20 | fun gen_qelim_conv precv postcv simpex_conv ins env atcv ncv qcv = | 
| 23489 | 21 | let | 
| 23466 | 22 | fun conv env p = | 
| 23489 | 23 | case (term_of p) of | 
| 41453 | 24 | Const(s,T)$_$_ => | 
| 23904 | 25 | if domain_type T = HOLogic.boolT | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 26 |           andalso member (op =) [@{const_name HOL.conj}, @{const_name HOL.disj},
 | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 27 |             @{const_name HOL.implies}, @{const_name HOL.eq}] s
 | 
| 41453 | 28 | then Conv.binop_conv (conv env) p | 
| 23904 | 29 | else atcv env p | 
| 41453 | 30 |   | Const(@{const_name Not},_)$_ => Conv.arg_conv (conv env) p
 | 
| 32603 | 31 |   | Const(@{const_name Ex},_)$Abs(s,_,_) =>
 | 
| 23489 | 32 | let | 
| 23466 | 33 | val (e,p0) = Thm.dest_comb p | 
| 34 | val (x,p') = Thm.dest_abs (SOME s) p0 | |
| 35 | val env' = ins x env | |
| 36 | val th = Thm.abstract_rule s x ((conv env' then_conv ncv env') p') | |
| 37 | |> Drule.arg_cong_rule e | |
| 38 | val th' = simpex_conv (Thm.rhs_of th) | |
| 39 | val (l,r) = Thm.dest_equals (cprop_of th') | |
| 23593 | 40 | in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th)) | 
| 23466 | 41 | else Thm.transitive (Thm.transitive th th') (conv env r) end | 
| 32603 | 42 |   | Const(@{const_name Ex},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p
 | 
| 43 |   | Const(@{const_name All},_)$_ =>
 | |
| 23489 | 44 | let | 
| 23466 | 45 | val p = Thm.dest_arg p | 
| 46 |      val ([(_,T)],[]) = Thm.match (@{cpat "All"}, Thm.dest_fun p)
 | |
| 23489 | 47 | val th = instantiate' [SOME T] [SOME p] all_not_ex | 
| 36945 | 48 | in Thm.transitive th (conv env (Thm.rhs_of th)) | 
| 23466 | 49 | end | 
| 50 | | _ => atcv env p | |
| 51 | in precv then_conv (conv env) then_conv postcv end | |
| 23489 | 52 | |
| 23904 | 53 | (* Instantiation of some parameter for most common cases *) | 
| 23489 | 54 | |
| 23904 | 55 | local | 
| 35410 | 56 | val pcv = | 
| 57 | Simplifier.rewrite | |
| 45654 | 58 |     (HOL_basic_ss addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib});
 | 
| 23466 | 59 | |
| 23904 | 60 | in fun standard_qelim_conv atcv ncv qcv p = | 
| 61 | gen_qelim_conv pcv pcv pcv cons (Thm.add_cterm_frees p []) atcv ncv qcv p | |
| 62 | end; | |
| 23466 | 63 | |
| 64 | end; |