| author | wenzelm | 
| Tue, 19 Sep 2006 23:01:52 +0200 | |
| changeset 20618 | 3f763be47c2f | 
| parent 20194 | c9dbce9a23a1 | 
| child 23035 | 643859163183 | 
| permissions | -rw-r--r-- | 
| 13876 | 1  | 
(* Title: HOL/Integ/qelim.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Amine Chaieb and Tobias Nipkow, TU Muenchen  | 
|
4  | 
||
5  | 
File containing the implementation of the proof protocoling  | 
|
6  | 
and proof generation for multiple quantified formulae.  | 
|
7  | 
*)  | 
|
8  | 
||
9  | 
signature QELIM =  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
10  | 
sig  | 
| 19250 | 11  | 
val tproof_of_mlift_qelim: theory -> (term -> bool) ->  | 
| 13876 | 12  | 
(string list -> term -> thm) -> (term -> thm) ->  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
13  | 
(term -> thm) -> term -> thm  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
14  | 
|
| 13876 | 15  | 
end;  | 
16  | 
||
17  | 
structure Qelim : QELIM =  | 
|
18  | 
struct  | 
|
19  | 
open CooperDec;  | 
|
20  | 
open CooperProof;  | 
|
21  | 
||
| 19250 | 22  | 
val cboolT = ctyp_of HOL.thy HOLogic.boolT;  | 
| 13876 | 23  | 
|
24  | 
(* List of the theorems to be used in the following*)  | 
|
25  | 
||
26  | 
val qe_ex_conj = thm "qe_ex_conj";  | 
|
27  | 
val qe_ex_nconj = thm "qe_ex_nconj";  | 
|
28  | 
val qe_ALL = thm "qe_ALL";  | 
|
29  | 
||
30  | 
||
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
31  | 
(*============= Compact version =====*)  | 
| 13876 | 32  | 
|
33  | 
||
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
34  | 
fun decomp_qe is_at afnp nfnp qfnp sg P =  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
35  | 
if is_at P then ([], fn [] => afnp [] P) else  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
36  | 
case P of  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
37  | 
   (Const("op &",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_conjI)
 | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
38  | 
   |(Const("op |",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_disjI)
 | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
39  | 
   |(Const("op -->",_)$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_impI)
 | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
40  | 
   |(Const("op =", Type ("fun",[Type ("bool", []),_]))$A$B) => ([A,B], fn [th1,th2] => [th1,th2] MRS qe_eqI)
 | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
41  | 
   |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
 | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
42  | 
   |(Const("Ex",_)$Abs(xn,xT,p)) => 
 | 
| 20194 | 43  | 
let val (xn1,p1) = Syntax.variant_abs(xn,xT,p)  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
44  | 
in ([p1],  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
45  | 
fn [th1_1] =>  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
46  | 
let val th2 = [th1_1,nfnp (snd (qe_get_terms th1_1))] MRS trans  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
47  | 
val eth1 = (forall_intr (cterm_of sg (Free(xn1,xT))) th2) COMP qe_exI  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
48  | 
val th3 = qfnp (snd(qe_get_terms eth1))  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
49  | 
in [eth1,th3] MRS trans  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
50  | 
end )  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
51  | 
end  | 
| 13876 | 52  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
53  | 
   |(Const("All",_)$Abs(xn,xT,p)) => ([(HOLogic.exists_const xT)$Abs(xn,xT,HOLogic.Not $ p)], fn [th] => th RS qe_ALL)
 | 
| 15531 | 54  | 
| _ => ([],fn [] => instantiate' [SOME (ctyp_of sg (type_of P))] [SOME (cterm_of sg P)] refl);  | 
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
55  | 
|
| 13876 | 56  | 
|
| 
14758
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
57  | 
fun tproof_of_mlift_qelim sg isat afnp nfnp qfnp p =  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
58  | 
let val th1 = thm_of sg (decomp_qe isat afnp nfnp qfnp sg) p  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
59  | 
val th2 = nfnp (snd (qe_get_terms th1))  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
60  | 
in [th1,th2] MRS trans  | 
| 
 
af3b71a46a1c
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller.
 
chaieb 
parents: 
13876 
diff
changeset
 | 
61  | 
end;  | 
| 13876 | 62  | 
end;  |