src/HOL/Integ/qelim.ML
author wenzelm
Wed, 09 Jun 2004 18:52:42 +0200
changeset 14898 a25550451b51
parent 14758 af3b71a46a1c
child 14981 e73f8140af78
permissions -rw-r--r--
Url.File;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/Integ/qelim.ML
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     3
    Author:     Amine Chaieb and Tobias Nipkow, TU Muenchen
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     5
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     6
File containing the implementation of the proof protocoling
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     7
and proof generation for multiple quantified formulae.
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     8
*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     9
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    10
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
    11
 sig
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
    12
 val tproof_of_mlift_qelim: Sign.sg -> (term -> bool) ->
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    13
   (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
    14
   (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
    15
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    16
end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    17
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    18
structure Qelim : QELIM =
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    19
struct
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    20
open CooperDec;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    21
open CooperProof;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    22
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    23
val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    24
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    25
(* List of the theorems to be used in the following*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    26
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    27
val qe_ex_conj = thm "qe_ex_conj";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    28
val qe_ex_nconj = thm "qe_ex_nconj";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    29
val qe_ALL = thm "qe_ALL";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    30
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    31
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
    32
(*============= Compact version =====*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    33
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    34
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
    35
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
    36
   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
    37
   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
    38
   (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
    39
   |(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
    40
   |(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
    41
   |(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
    42
   |(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
    43
   |(Const("Ex",_)$Abs(xn,xT,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
    44
      let val (xn1,p1) = variant_abs(xn,xT,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
    45
      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
    46
        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
    47
          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
    48
              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
    49
              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
    50
          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
    51
          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
    52
      end
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    53
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
    54
   |(Const("All",_)$Abs(xn,xT,p)) => ([(HOLogic.exists_const xT)$Abs(xn,xT,HOLogic.Not $ p)], fn [th] => th RS qe_ALL)
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
   | _ => ([],fn [] => instantiate' [Some (ctyp_of sg (type_of P))] [Some (cterm_of sg P)] refl);
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
    56
 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    57
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
    58
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
    59
   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
    60
       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
    61
    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
    62
    end;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    63
end;