| author | haftmann | 
| Tue, 23 Apr 2013 11:14:50 +0200 | |
| changeset 51734 | d504e349e951 | 
| parent 51717 | 9e7d1c139569 | 
| child 54742 | 7a86358a3c0b | 
| permissions | -rw-r--r-- | 
| 37744 | 1 | (* Title: HOL/Decision_Procs/langford.ML | 
| 29269 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29265diff
changeset | 2 | Author: Amine Chaieb, TU Muenchen | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29265diff
changeset | 3 | *) | 
| 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 wenzelm parents: 
29265diff
changeset | 4 | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 5 | signature LANGFORD_QE = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 6 | sig | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 7 | val dlo_tac : Proof.context -> int -> tactic | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 8 | val dlo_conv : Proof.context -> cterm -> thm | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 9 | end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 10 | |
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 11 | structure LangfordQE :LANGFORD_QE = | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 12 | struct | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 13 | |
| 30452 | 14 | val dest_set = | 
| 15 | let | |
| 16 | fun h acc ct = | |
| 17 | case term_of ct of | |
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
30452diff
changeset | 18 |      Const (@{const_name Orderings.bot}, _) => acc
 | 
| 30452 | 19 |    | Const (@{const_name insert}, _) $ _ $ t => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct)
 | 
| 20 | in h [] end; | |
| 21 | ||
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 22 | fun prove_finite cT u = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 23 | let val [th0,th1] = map (instantiate' [SOME cT] []) @{thms "finite.intros"}
 | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 24 | fun ins x th = | 
| 36945 | 25 | Thm.implies_elim (instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 26 | (Thm.cprop_of th), SOME x] th1) th | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 27 | in fold ins u th0 end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 28 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 29 | fun simp_rule ctxt = | 
| 45654 | 30 | Conv.fconv_rule | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 31 | (Conv.arg_conv | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 32 |       (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms ball_simps simp_thms})));
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 33 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 34 | fun basic_dloqe ctxt stupid dlo_qeth dlo_qeth_nolb dlo_qeth_noub gather ep = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 35 | case term_of ep of | 
| 38558 | 36 |   Const(@{const_name Ex},_)$_ => 
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 37 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 38 | val p = Thm.dest_arg ep | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 39 | val ths = | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 40 | simplify (put_simpset HOL_basic_ss ctxt addsimps gather) (instantiate' [] [SOME p] stupid) | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 41 | val (L,U) = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 42 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 43 | val (x,q) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of ths)) | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 44 | in (Thm.dest_arg1 q |> Thm.dest_arg1, Thm.dest_arg q |> Thm.dest_arg1) | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 45 | end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 46 | fun proveneF S = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 47 | let val (a,A) = Thm.dest_comb S |>> Thm.dest_arg | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 48 | val cT = ctyp_of_term a | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 49 | val ne = instantiate' [SOME cT] [SOME a, SOME A] | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 50 |                     @{thm insert_not_empty}
 | 
| 30452 | 51 | val f = prove_finite cT (dest_set S) | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 52 | in (ne, f) end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 53 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 54 | val qe = case (term_of L, term_of U) of | 
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
30452diff
changeset | 55 |       (Const (@{const_name Orderings.bot}, _),_) =>  
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 56 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 57 | val (neU,fU) = proveneF U | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 58 | in simp_rule ctxt (Thm.transitive ths (dlo_qeth_nolb OF [neU, fU])) end | 
| 32264 
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
 haftmann parents: 
30452diff
changeset | 59 |     | (_,Const (@{const_name Orderings.bot}, _)) =>  
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 60 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 61 | val (neL,fL) = proveneF L | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 62 | in simp_rule ctxt (Thm.transitive ths (dlo_qeth_noub OF [neL, fL])) end | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 63 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 64 | | (_,_) => | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 65 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 66 | val (neL,fL) = proveneF L | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 67 | val (neU,fU) = proveneF U | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 68 | in simp_rule ctxt (Thm.transitive ths (dlo_qeth OF [neL, neU, fL, fU])) | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 69 | end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 70 | in qe end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 71 | | _ => error "dlo_qe : Not an existential formula"; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 72 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 73 | val all_conjuncts = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 74 | let fun h acc ct = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 75 | case term_of ct of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 76 |    @{term HOL.conj}$_$_ => h (h acc (Thm.dest_arg ct)) (Thm.dest_arg1 ct)
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 77 | | _ => ct::acc | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 78 | in h [] end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 79 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 80 | fun conjuncts ct = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 81 | case term_of ct of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 82 |   @{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct))
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 83 | | _ => [ct]; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 84 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 85 | fun fold1 f = foldr1 (uncurry f); | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 86 | |
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 87 | val list_conj = fold1 (fn c => fn c' => Thm.apply (Thm.apply @{cterm HOL.conj} c) c') ;
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 88 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 89 | fun mk_conj_tab th = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 90 | let fun h acc th = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 91 | case prop_of th of | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 92 |    @{term "Trueprop"}$(@{term HOL.conj}$p$q) => 
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 93 | h (h acc (th RS conjunct2)) (th RS conjunct1) | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 94 |   | @{term "Trueprop"}$p => (p,th)::acc
 | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 95 | in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 96 | |
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 97 | fun is_conj (@{term HOL.conj}$_$_) = true
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 98 | | is_conj _ = false; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 99 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 100 | fun prove_conj tab cjs = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 101 | case cjs of | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 102 | [c] => if is_conj (term_of c) then prove_conj tab (conjuncts c) else tab c | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 103 | | c::cs => conjI OF [prove_conj tab [c], prove_conj tab cs]; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 104 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 105 | fun conj_aci_rule eq = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 106 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 107 | val (l,r) = Thm.dest_equals eq | 
| 36945 | 108 | fun tabl c = the (Termtab.lookup (mk_conj_tab (Thm.assume l)) (term_of c)) | 
| 109 | fun tabr c = the (Termtab.lookup (mk_conj_tab (Thm.assume r)) (term_of c)) | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 110 | val ll = Thm.dest_arg l | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 111 | val rr = Thm.dest_arg r | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 112 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 113 | val thl = prove_conj tabl (conjuncts rr) | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 114 | |> Drule.implies_intr_hyps | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 115 | val thr = prove_conj tabr (conjuncts ll) | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 116 | |> Drule.implies_intr_hyps | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 117 |   val eqI = instantiate' [] [SOME ll, SOME rr] @{thm iffI}
 | 
| 36945 | 118 | in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end; | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 119 | |
| 44121 | 120 | fun contains x ct = member (op aconv) (Misc_Legacy.term_frees (term_of ct)) (term_of x); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 121 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 122 | fun is_eqx x eq = case term_of eq of | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 123 |    Const(@{const_name HOL.eq},_)$l$r => l aconv term_of x orelse r aconv term_of x
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 124 | | _ => false ; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 125 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 126 | local | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 127 | fun proc ctxt ct = | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 128 | case term_of ct of | 
| 38558 | 129 |   Const(@{const_name Ex},_)$Abs (xn,_,_) => 
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 130 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 131 | val e = Thm.dest_fun ct | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 132 | val (x,p) = Thm.dest_abs (SOME xn) (Thm.dest_arg ct) | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 133 |     val Pp = Thm.apply @{cterm "Trueprop"} p 
 | 
| 30148 | 134 | val (eqs,neqs) = List.partition (is_eqx x) (all_conjuncts p) | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 135 | in case eqs of | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 136 | [] => | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 137 | let | 
| 30148 | 138 | val (dx,ndx) = List.partition (contains x) neqs | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 139 | in case ndx of [] => NONE | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 140 | | _ => | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 141 |             conj_aci_rule (Thm.mk_binop @{cterm "op == :: prop => _"} Pp 
 | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 142 |                  (Thm.apply @{cterm Trueprop} (list_conj (ndx @dx))))
 | 
| 36945 | 143 | |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 144 | |> Conv.fconv_rule (Conv.arg_conv | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 145 |                (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps})))
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 146 | |> SOME | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 147 | end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 148 |     | _ => conj_aci_rule (Thm.mk_binop @{cterm "op == :: prop => _"} Pp 
 | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 149 |                  (Thm.apply @{cterm Trueprop} (list_conj (eqs@neqs))))
 | 
| 36945 | 150 | |> Thm.abstract_rule xn x |> Drule.arg_cong_rule e | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 151 | |> Conv.fconv_rule (Conv.arg_conv | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 152 |                (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps})))
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 153 | |> SOME | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 154 | end | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 155 | | _ => NONE; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 156 | in val reduce_ex_simproc = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 157 | Simplifier.make_simproc | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 158 |   {lhss = [@{cpat "EX x. ?P x"}] , name = "reduce_ex_simproc",
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 159 | proc = K proc, identifier = []} | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 160 | end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 161 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 162 | fun raw_dlo_conv ctxt dlo_ss | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 163 |           ({qe_bnds, qe_nolb, qe_noub, gst, gs, atoms}:Langford_Data.entry) = 
 | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 164 | let | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 165 |   val ctxt' = put_simpset dlo_ss ctxt addsimps @{thms "dnf_simps"} addsimprocs [reduce_ex_simproc]
 | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 166 | val dnfex_conv = Simplifier.rewrite ctxt' | 
| 45654 | 167 | val pcv = | 
| 168 | Simplifier.rewrite | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 169 | (put_simpset dlo_ss ctxt | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 170 |         addsimps @{thms simp_thms ex_simps all_simps all_not_ex not_all ex_disj_distrib})
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 171 | in fn p => | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 172 | Qelim.gen_qelim_conv pcv pcv dnfex_conv cons | 
| 36945 | 173 | (Thm.add_cterm_frees p []) (K Thm.reflexive) (K Thm.reflexive) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 174 | (K (basic_dloqe ctxt gst qe_bnds qe_nolb qe_noub gs)) p | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 175 | end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 176 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 177 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 178 | val grab_atom_bop = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 179 | let | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 180 | fun h bounds tm = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 181 | (case term_of tm of | 
| 38864 
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
 haftmann parents: 
38795diff
changeset | 182 |      Const (@{const_name HOL.eq}, T) $ _ $ _ =>
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 183 | if domain_type T = HOLogic.boolT then find_args bounds tm | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 184 | else Thm.dest_fun2 tm | 
| 38558 | 185 |    | Const (@{const_name Not}, _) $ _ => h bounds (Thm.dest_arg tm)
 | 
| 186 |    | Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm)
 | |
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 187 |    | Const ("all", _) $ _ => find_body bounds (Thm.dest_arg tm)
 | 
| 38558 | 188 |    | Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm)
 | 
| 38795 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 189 |    | Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm
 | 
| 
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
 haftmann parents: 
38786diff
changeset | 190 |    | Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm
 | 
| 38786 
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
 haftmann parents: 
38558diff
changeset | 191 |    | Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 192 |    | Const ("==>", _) $ _ $ _ => find_args bounds tm
 | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 193 |    | Const ("==", _) $ _ $ _ => find_args bounds tm
 | 
| 38558 | 194 |    | Const (@{const_name Trueprop}, _) $ _ => h bounds (Thm.dest_arg tm)
 | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 195 | | _ => Thm.dest_fun2 tm) | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 196 | and find_args bounds tm = | 
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 197 | (h bounds (Thm.dest_arg tm) handle CTERM _ => h bounds (Thm.dest_arg1 tm)) | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 198 | and find_body bounds b = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 199 | let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 200 | in h (bounds + 1) b' end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 201 | in h end; | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 202 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 203 | fun dlo_instance ctxt tm = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 204 | (fst (Langford_Data.get ctxt), | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 205 | Langford_Data.match ctxt (grab_atom_bop 0 tm)); | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 206 | |
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 207 | fun dlo_conv ctxt tm = | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 208 | (case dlo_instance ctxt tm of | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 209 |     (_, NONE) => raise CTERM ("dlo_conv (langford): no corresponding instance in context!", [tm])
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 210 | | (ss, SOME instance) => raw_dlo_conv ctxt ss instance tm); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 211 | |
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 212 | fun generalize_tac f = CSUBGOAL (fn (p, i) => PRIMITIVE (fn st => | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 213 | let | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 214 |    fun all T = Drule.cterm_rule (instantiate' [SOME T] []) @{cpat "all"}
 | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
45654diff
changeset | 215 | fun gen x t = Thm.apply (all (ctyp_of_term x)) (Thm.lambda x t) | 
| 35408 | 216 | val ts = sort (fn (a,b) => Term_Ord.fast_term_ord (term_of a, term_of b)) (f p) | 
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 217 | val p' = fold_rev gen ts p | 
| 36945 | 218 | in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end)); | 
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 219 | |
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 220 | |
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 221 | fun cfrees ats ct = | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 222 | let | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 223 | val ins = insert (op aconvc) | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 224 | fun h acc t = | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 225 | case (term_of t) of | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 226 | b$_$_ => if member (op aconvc) ats (Thm.dest_fun2 t) | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 227 | then ins (Thm.dest_arg t) (ins (Thm.dest_arg1 t) acc) | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 228 | else h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 229 | | _$_ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 230 | | Abs(_,_,_) => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc)) | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 231 | | Free _ => if member (op aconvc) ats t then acc else ins t acc | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 232 | | Var _ => if member (op aconvc) ats t then acc else ins t acc | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 233 | | _ => acc | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 234 | in h [] ct end | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 235 | |
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 236 | fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => | 
| 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 237 | (case dlo_instance ctxt p of | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 238 | (ss, NONE) => simp_tac (put_simpset ss ctxt) i | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 239 | | (ss, SOME instance) => | 
| 35625 | 240 | Object_Logic.full_atomize_tac i THEN | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 241 | simp_tac (put_simpset ss ctxt) i | 
| 24083 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 242 | THEN (CONVERSION Thm.eta_long_conversion) i | 
| 
4ea3656380b1
find_body goes under meta-quantifier ; tactic generalizes free variables;
 chaieb parents: 
23906diff
changeset | 243 | THEN (TRY o generalize_tac (cfrees (#atoms instance))) i | 
| 35625 | 244 | THEN Object_Logic.full_atomize_tac i | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 245 | THEN CONVERSION (Object_Logic.judgment_conv (raw_dlo_conv ctxt ss instance)) i | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
46497diff
changeset | 246 | THEN (simp_tac (put_simpset ss ctxt) i))); | 
| 23906 
e61361aa23b2
Quantifier elimination for Dense linear orders after Langford
 chaieb parents: diff
changeset | 247 | end; |