25 |
25 |
26 (*-----------------------------------------------------------------*) |
26 (*-----------------------------------------------------------------*) |
27 (*cooper_pp: provefunction for the one-exstance quantifier elimination*) |
27 (*cooper_pp: provefunction for the one-exstance quantifier elimination*) |
28 (* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*) |
28 (* Here still only one problem : The proof for the arithmetical transformations done on the dvd atomic formulae*) |
29 (*-----------------------------------------------------------------*) |
29 (*-----------------------------------------------------------------*) |
|
30 |
|
31 val presburger_ss = simpset_of (theory "Presburger"); |
30 |
32 |
31 fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = |
33 fun cooper_pp sg (fm as e$Abs(xn,xT,p)) = |
32 let val (xn1,p1) = variant_abs (xn,xT,p) |
34 let val (xn1,p1) = variant_abs (xn,xT,p) |
33 in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end; |
35 in (CooperProof.cooper_prv sg (Free (xn1, xT)) p1) end; |
34 |
36 |
63 | add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs) |
65 | add_term_typed_consts (Abs (_, _, t), cs) = add_term_typed_consts (t, cs) |
64 | add_term_typed_consts (_, cs) = cs; |
66 | add_term_typed_consts (_, cs) = cs; |
65 |
67 |
66 fun term_typed_consts t = add_term_typed_consts(t,[]); |
68 fun term_typed_consts t = add_term_typed_consts(t,[]); |
67 |
69 |
68 (* put a term into eta long beta normal form *) |
|
69 fun eta_long Ts (Abs (s, T, t)) = Abs (s, T, eta_long (T :: Ts) t) |
|
70 | eta_long Ts t = (case strip_comb t of |
|
71 (Abs _, _) => eta_long Ts (Envir.beta_norm t) |
|
72 | (u, ts) => |
|
73 let |
|
74 val Us = binder_types (fastype_of1 (Ts, t)); |
|
75 val i = length Us |
|
76 in list_abs (map (pair "x") Us, |
|
77 list_comb (incr_boundvars i u, map (eta_long (rev Us @ Ts)) |
|
78 (map (incr_boundvars i) ts @ map Bound (i - 1 downto 0)))) |
|
79 end); |
|
80 |
|
81 (* Some Types*) |
70 (* Some Types*) |
82 val bT = HOLogic.boolT; |
71 val bT = HOLogic.boolT; |
83 val iT = HOLogic.intT; |
72 val iT = HOLogic.intT; |
84 val binT = HOLogic.binT; |
73 val binT = HOLogic.binT; |
85 val nT = HOLogic.natT; |
74 val nT = HOLogic.natT; |
107 ("op +", iT --> iT --> iT), |
96 ("op +", iT --> iT --> iT), |
108 ("op -", iT --> iT --> iT), |
97 ("op -", iT --> iT --> iT), |
109 ("op *", iT --> iT --> iT), |
98 ("op *", iT --> iT --> iT), |
110 ("HOL.abs", iT --> iT), |
99 ("HOL.abs", iT --> iT), |
111 ("uminus", iT --> iT), |
100 ("uminus", iT --> iT), |
|
101 ("HOL.max", iT --> iT --> iT), |
|
102 ("HOL.min", iT --> iT --> iT), |
112 |
103 |
113 ("op <=", nT --> nT --> bT), |
104 ("op <=", nT --> nT --> bT), |
114 ("op =", nT --> nT --> bT), |
105 ("op =", nT --> nT --> bT), |
115 ("op <", nT --> nT --> bT), |
106 ("op <", nT --> nT --> bT), |
116 ("Divides.op dvd", nT --> nT --> bT), |
107 ("Divides.op dvd", nT --> nT --> bT), |
118 ("Divides.op mod", nT --> nT --> nT), |
109 ("Divides.op mod", nT --> nT --> nT), |
119 ("op +", nT --> nT --> nT), |
110 ("op +", nT --> nT --> nT), |
120 ("op -", nT --> nT --> nT), |
111 ("op -", nT --> nT --> nT), |
121 ("op *", nT --> nT --> nT), |
112 ("op *", nT --> nT --> nT), |
122 ("Suc", nT --> nT), |
113 ("Suc", nT --> nT), |
|
114 ("HOL.max", nT --> nT --> nT), |
|
115 ("HOL.min", nT --> nT --> nT), |
123 |
116 |
124 ("Numeral.bin.Bit", binT --> bT --> binT), |
117 ("Numeral.bin.Bit", binT --> bT --> binT), |
125 ("Numeral.bin.Pls", binT), |
118 ("Numeral.bin.Pls", binT), |
126 ("Numeral.bin.Min", binT), |
119 ("Numeral.bin.Min", binT), |
127 ("Numeral.number_of", binT --> iT), |
120 ("Numeral.number_of", binT --> iT), |
266 val ct = cterm_of sg (HOLogic.mk_Trueprop t) |
259 val ct = cterm_of sg (HOLogic.mk_Trueprop t) |
267 (* Theorem for the nat --> int transformation *) |
260 (* Theorem for the nat --> int transformation *) |
268 val pre_thm = Seq.hd (EVERY |
261 val pre_thm = Seq.hd (EVERY |
269 [simp_tac simpset0 1, |
262 [simp_tac simpset0 1, |
270 TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), |
263 TRY (simp_tac simpset1 1), TRY (simp_tac simpset2 1), |
271 TRY (simp_tac simpset3 1)] |
264 TRY (simp_tac simpset3 1), TRY (simp_tac presburger_ss 1)] |
272 (trivial ct)) |
265 (trivial ct)) |
273 fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) |
266 fun assm_tac i = REPEAT_DETERM_N nh (assume_tac i) |
274 (* The result of the quantifier elimination *) |
267 (* The result of the quantifier elimination *) |
275 val (th, tac) = case (prop_of pre_thm) of |
268 val (th, tac) = case (prop_of pre_thm) of |
276 Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => |
269 Const ("==>", _) $ (Const ("Trueprop", _) $ t1) $ _ => |