src/HOL/Tools/Presburger/cooper_proof.ML
author wenzelm
Thu, 07 Apr 2005 09:25:33 +0200
changeset 15661 9ef583b08647
parent 15531 08c8dad8e399
child 16389 48832eba5b1e
permissions -rw-r--r--
reverted renaming of Some/None in comments and strings;
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/cooper_proof.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
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     5
File containing the implementation of the proof
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     6
generation for Cooper Algorithm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     7
*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
     8
14920
a7525235e20f An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
chaieb
parents: 14877
diff changeset
     9
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    10
signature COOPER_PROOF =
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    11
sig
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    12
  val qe_Not : thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    13
  val qe_conjI : thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    14
  val qe_disjI : thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    15
  val qe_impI : thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    16
  val qe_eqI : thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    17
  val qe_exI : thm
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
    18
  val list_to_set : typ -> term list -> term
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    19
  val qe_get_terms : thm -> term * term
15122
4b52eeb62807 proof_of_evalc corrected;
chaieb
parents: 15107
diff changeset
    20
  val cooper_prv  : Sign.sg -> term -> term -> thm
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    21
  val proof_of_evalc : Sign.sg -> term -> thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    22
  val proof_of_cnnf : Sign.sg -> term -> (term -> thm) -> thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    23
  val proof_of_linform : Sign.sg -> string list -> term -> thm
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
    24
  val proof_of_adjustcoeffeq : Sign.sg -> term -> int -> 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: 14479
diff changeset
    25
  val prove_elementar : Sign.sg -> string -> 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: 14479
diff changeset
    26
  val thm_of : Sign.sg -> (term -> (term list * (thm list -> thm))) -> term -> thm
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    27
end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    28
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    29
structure CooperProof : COOPER_PROOF =
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    30
struct
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    31
open CooperDec;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    32
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: 14479
diff changeset
    33
(*
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: 14479
diff changeset
    34
val presburger_ss = simpset_of (theory "Presburger")
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: 14479
diff changeset
    35
  addsimps [zdiff_def] delsimps [symmetric zdiff_def];
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: 14479
diff changeset
    36
*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    37
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    38
val presburger_ss = simpset_of (theory "Presburger")
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: 14479
diff changeset
    39
  addsimps[diff_int_def] delsimps [thm"diff_int_def_symmetric"];
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    40
val cboolT = ctyp_of (sign_of HOL.thy) HOLogic.boolT;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    41
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    42
(*Theorems that will be used later for the proofgeneration*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    43
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    44
val zdvd_iff_zmod_eq_0 = thm "zdvd_iff_zmod_eq_0";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    45
val unity_coeff_ex = thm "unity_coeff_ex";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    46
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: 14479
diff changeset
    47
(* Thorems for proving the adjustment of the coeffitients*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    48
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    49
val ac_lt_eq =  thm "ac_lt_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    50
val ac_eq_eq = thm "ac_eq_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    51
val ac_dvd_eq = thm "ac_dvd_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    52
val ac_pi_eq = thm "ac_pi_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    53
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    54
(* The logical compination of the sythetised properties*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    55
val qe_Not = thm "qe_Not";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    56
val qe_conjI = thm "qe_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    57
val qe_disjI = thm "qe_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    58
val qe_impI = thm "qe_impI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    59
val qe_eqI = thm "qe_eqI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    60
val qe_exI = thm "qe_exI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    61
val qe_ALLI = thm "qe_ALLI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    62
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: 14479
diff changeset
    63
(*Modulo D property for Pminusinf an Plusinf *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    64
val fm_modd_minf = thm "fm_modd_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    65
val not_dvd_modd_minf = thm "not_dvd_modd_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    66
val dvd_modd_minf = thm "dvd_modd_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    67
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    68
val fm_modd_pinf = thm "fm_modd_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    69
val not_dvd_modd_pinf = thm "not_dvd_modd_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    70
val dvd_modd_pinf = thm "dvd_modd_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    71
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: 14479
diff changeset
    72
(* the minusinfinity proprty*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    73
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    74
val fm_eq_minf = thm "fm_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    75
val neq_eq_minf = thm "neq_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    76
val eq_eq_minf = thm "eq_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    77
val le_eq_minf = thm "le_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    78
val len_eq_minf = thm "len_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    79
val not_dvd_eq_minf = thm "not_dvd_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    80
val dvd_eq_minf = thm "dvd_eq_minf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    81
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: 14479
diff changeset
    82
(* the Plusinfinity proprty*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    83
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    84
val fm_eq_pinf = thm "fm_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    85
val neq_eq_pinf = thm "neq_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    86
val eq_eq_pinf = thm "eq_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    87
val le_eq_pinf = thm "le_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    88
val len_eq_pinf = thm "len_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    89
val not_dvd_eq_pinf = thm "not_dvd_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    90
val dvd_eq_pinf = thm "dvd_eq_pinf";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    91
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    92
(*Logical construction of the Property*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    93
val eq_minf_conjI = thm "eq_minf_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    94
val eq_minf_disjI = thm "eq_minf_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    95
val modd_minf_disjI = thm "modd_minf_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    96
val modd_minf_conjI = thm "modd_minf_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    97
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    98
val eq_pinf_conjI = thm "eq_pinf_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
    99
val eq_pinf_disjI = thm "eq_pinf_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   100
val modd_pinf_disjI = thm "modd_pinf_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   101
val modd_pinf_conjI = thm "modd_pinf_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   102
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   103
(*Cooper Backwards...*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   104
(*Bset*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   105
val not_bst_p_fm = thm "not_bst_p_fm";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   106
val not_bst_p_ne = thm "not_bst_p_ne";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   107
val not_bst_p_eq = thm "not_bst_p_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   108
val not_bst_p_gt = thm "not_bst_p_gt";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   109
val not_bst_p_lt = thm "not_bst_p_lt";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   110
val not_bst_p_ndvd = thm "not_bst_p_ndvd";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   111
val not_bst_p_dvd = thm "not_bst_p_dvd";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   112
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   113
(*Aset*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   114
val not_ast_p_fm = thm "not_ast_p_fm";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   115
val not_ast_p_ne = thm "not_ast_p_ne";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   116
val not_ast_p_eq = thm "not_ast_p_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   117
val not_ast_p_gt = thm "not_ast_p_gt";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   118
val not_ast_p_lt = thm "not_ast_p_lt";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   119
val not_ast_p_ndvd = thm "not_ast_p_ndvd";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   120
val not_ast_p_dvd = thm "not_ast_p_dvd";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   121
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   122
(*Logical construction of the prop*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   123
(*Bset*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   124
val not_bst_p_conjI = thm "not_bst_p_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   125
val not_bst_p_disjI = thm "not_bst_p_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   126
val not_bst_p_Q_elim = thm "not_bst_p_Q_elim";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   127
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   128
(*Aset*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   129
val not_ast_p_conjI = thm "not_ast_p_conjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   130
val not_ast_p_disjI = thm "not_ast_p_disjI";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   131
val not_ast_p_Q_elim = thm "not_ast_p_Q_elim";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   132
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   133
(*Cooper*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   134
val cppi_eq = thm "cppi_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   135
val cpmi_eq = thm "cpmi_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   136
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   137
(*Others*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   138
val simp_from_to = thm "simp_from_to";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   139
val P_eqtrue = thm "P_eqtrue";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   140
val P_eqfalse = thm "P_eqfalse";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   141
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   142
(*For Proving NNF*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   143
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   144
val nnf_nn = thm "nnf_nn";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   145
val nnf_im = thm "nnf_im";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   146
val nnf_eq = thm "nnf_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   147
val nnf_sdj = thm "nnf_sdj";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   148
val nnf_ncj = thm "nnf_ncj";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   149
val nnf_nim = thm "nnf_nim";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   150
val nnf_neq = thm "nnf_neq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   151
val nnf_ndj = thm "nnf_ndj";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   152
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   153
(*For Proving term linearizition*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   154
val linearize_dvd = thm "linearize_dvd";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   155
val lf_lt = thm "lf_lt";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   156
val lf_eq = thm "lf_eq";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   157
val lf_dvd = thm "lf_dvd";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   158
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   159
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   160
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   161
(*This function norm_zero_one  replaces the occurences of Numeral1 and Numeral0*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   162
(*Respectively by their abstract representation Const("1",..) and COnst("0",..)*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   163
(*this is necessary because the theorems use this representation.*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   164
(* This function should be elminated in next versions...*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   165
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   166
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   167
fun norm_zero_one fm = case fm of
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   168
  (Const ("op *",_) $ c $ t) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   169
    if c = one then (norm_zero_one t)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   170
    else if (dest_numeral c = ~1) 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   171
         then (Const("uminus",HOLogic.intT --> HOLogic.intT) $ (norm_zero_one t))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   172
         else (HOLogic.mk_binop "op *" (norm_zero_one c,norm_zero_one t))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   173
  |(node $ rest) => ((norm_zero_one node)$(norm_zero_one rest))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   174
  |(Abs(x,T,p)) => (Abs(x,T,(norm_zero_one p)))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   175
  |_ => fm;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   176
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   177
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   178
(*function list to Set, constructs a set containing all elements of a given list.*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   179
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   180
fun list_to_set T1 l = let val T = (HOLogic.mk_setT T1) in 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   181
	case l of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   182
		[] => Const ("{}",T)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   183
		|(h::t) => Const("insert", T1 --> (T --> T)) $ h $(list_to_set T1 t)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   184
		end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   185
		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   186
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   187
(* Returns both sides of an equvalence in the theorem*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   188
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   189
fun qe_get_terms th = let val (_$(Const("op =",Type ("fun",[Type ("bool", []),_])) $ A $ B )) = prop_of th in (A,B) end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   190
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   191
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   192
(* Modified version of the simple version with minimal amount of checking and postprocessing*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   193
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   194
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   195
fun simple_prove_goal_cterm2 G tacs =
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   196
  let
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   197
    fun check NONE = error "prove_goal: tactic failed"
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   198
      | check (SOME (thm, _)) = (case nprems_of thm of
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   199
            0 => thm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   200
          | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!"))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   201
  in check (Seq.pull (EVERY tacs (trivial G))) end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   202
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   203
(*-------------------------------------------------------------*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   204
(*-------------------------------------------------------------*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   205
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   206
fun cert_Trueprop sg t = cterm_of sg (HOLogic.mk_Trueprop t);
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   207
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   208
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   209
(*This function proove elementar will be used to generate proofs at runtime*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   210
(*It is is based on the isabelle function proove_goalw_cterm and is thought to *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   211
(*prove properties such as a dvd b (essentially) that are only to make at
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   212
runtime.*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   213
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   214
fun prove_elementar sg s fm2 = case s of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   215
  (*"ss" like simplification with simpset*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   216
  "ss" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   217
    let
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: 14479
diff changeset
   218
      val ss = presburger_ss addsimps
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: 14479
diff changeset
   219
        [zdvd_iff_zmod_eq_0,unity_coeff_ex]
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   220
      val ct =  cert_Trueprop sg fm2
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   221
    in 
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   222
      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)] 
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   223
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   224
    end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   225
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   226
  (*"bl" like blast tactic*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   227
  (* Is only used in the harrisons like proof procedure *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   228
  | "bl" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   229
     let val ct = cert_Trueprop sg fm2
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   230
     in
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   231
       simple_prove_goal_cterm2 ct [blast_tac HOL_cs 1]
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   232
     end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   233
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   234
  (*"ed" like Existence disjunctions ...*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   235
  (* Is only used in the harrisons like proof procedure *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   236
  | "ed" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   237
    let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   238
      val ex_disj_tacs =
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   239
        let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   240
          val tac1 = EVERY[REPEAT(resolve_tac [disjI1,disjI2] 1), etac exI 1]
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   241
          val tac2 = EVERY[etac exE 1, rtac exI 1,
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   242
            REPEAT(resolve_tac [disjI1,disjI2] 1), assumption 1]
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   243
	in [rtac iffI 1,
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   244
          etac exE 1, REPEAT(EVERY[etac disjE 1, tac1]), tac1,
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   245
          REPEAT(EVERY[etac disjE 1, tac2]), tac2]
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   246
        end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   247
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   248
      val ct = cert_Trueprop sg fm2
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   249
    in 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   250
      simple_prove_goal_cterm2 ct ex_disj_tacs
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   251
    end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   252
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   253
  | "fa" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   254
    let val ct = cert_Trueprop sg fm2
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   255
    in simple_prove_goal_cterm2 ct [simple_arith_tac 1]  
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   256
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   257
    end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   258
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   259
  | "sa" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   260
    let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   261
      val ss = presburger_ss addsimps zadd_ac
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   262
      val ct = cert_Trueprop sg fm2
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   263
    in 
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   264
      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]  
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   265
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   266
    end
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: 14479
diff changeset
   267
  (* like Existance Conjunction *)
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: 14479
diff changeset
   268
  | "ec" =>
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: 14479
diff changeset
   269
    let
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: 14479
diff changeset
   270
      val ss = presburger_ss addsimps zadd_ac
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: 14479
diff changeset
   271
      val ct = cert_Trueprop sg fm2
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: 14479
diff changeset
   272
    in 
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: 14479
diff changeset
   273
      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (blast_tac HOL_cs 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: 14479
diff changeset
   274
    end
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   275
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   276
  | "ac" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   277
    let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   278
      val ss = HOL_basic_ss addsimps zadd_ac
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   279
      val ct = cert_Trueprop sg fm2
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   280
    in 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   281
      simple_prove_goal_cterm2 ct [simp_tac ss 1]
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   282
    end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   283
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   284
  | "lf" =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   285
    let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   286
      val ss = presburger_ss addsimps zadd_ac
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   287
      val ct = cert_Trueprop sg fm2
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   288
    in 
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   289
      simple_prove_goal_cterm2 ct [simp_tac ss 1, TRY (simple_arith_tac 1)]  
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   290
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   291
    end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   292
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: 14479
diff changeset
   293
(*=============================================================*)
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: 14479
diff changeset
   294
(*-------------------------------------------------------------*)
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: 14479
diff changeset
   295
(*              The new compact model                          *)
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: 14479
diff changeset
   296
(*-------------------------------------------------------------*)
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: 14479
diff changeset
   297
(*=============================================================*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   298
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: 14479
diff changeset
   299
fun thm_of sg decomp t = 
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: 14479
diff changeset
   300
    let val (ts,recomb) = decomp t 
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: 14479
diff changeset
   301
    in recomb (map (thm_of sg decomp) ts) 
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: 14479
diff changeset
   302
    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: 14479
diff changeset
   303
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: 14479
diff changeset
   304
(*==================================================*)
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: 14479
diff changeset
   305
(*     Compact Version for adjustcoeffeq            *)
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: 14479
diff changeset
   306
(*==================================================*)
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: 14479
diff changeset
   307
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: 14479
diff changeset
   308
fun decomp_adjustcoeffeq sg x l fm = case fm 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: 14479
diff changeset
   309
    (Const("Not",_)$(Const("op <",_) $(Const("0",_)) $(rt as (Const ("op +", _)$(Const ("op *",_) $    c $ y ) $z )))) => 
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: 14479
diff changeset
   310
     let  
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: 14479
diff changeset
   311
        val m = l div (dest_numeral c) 
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: 14479
diff changeset
   312
        val n = if (x = y) then abs (m) else 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: 14479
diff changeset
   313
        val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div n)*l) ), x)) 
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: 14479
diff changeset
   314
        val rs = if (x = y) 
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: 14479
diff changeset
   315
                 then (HOLogic.mk_binrel "op <" (zero,linear_sub [] (mk_numeral n) (HOLogic.mk_binop "op +" ( xtm ,( linear_cmul n z) )))) 
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: 14479
diff changeset
   316
                 else HOLogic.mk_binrel "op <" (zero,linear_sub [] one rt )
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: 14479
diff changeset
   317
        val ck = cterm_of sg (mk_numeral n)
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: 14479
diff changeset
   318
        val cc = cterm_of sg c
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: 14479
diff changeset
   319
        val ct = cterm_of sg z
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: 14479
diff changeset
   320
        val cx = cterm_of sg y
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: 14479
diff changeset
   321
        val pre = prove_elementar sg "lf" 
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: 14479
diff changeset
   322
            (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral n)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   323
        val th1 = (pre RS (instantiate' [] [SOME ck,SOME cc, SOME cx, SOME ct] (ac_pi_eq)))
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: 14479
diff changeset
   324
        in ([], fn [] => [th1,(prove_elementar sg "sa" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] 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: 14479
diff changeset
   325
        end
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   326
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: 14479
diff changeset
   327
  |(Const(p,_) $a $( Const ("op +", _)$(Const ("op *",_) $ 
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: 14479
diff changeset
   328
      c $ y ) $t )) => 
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: 14479
diff changeset
   329
   if (is_arith_rel fm) andalso (x = y) 
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: 14479
diff changeset
   330
   then  
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: 14479
diff changeset
   331
        let val m = l div (dest_numeral c) 
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: 14479
diff changeset
   332
           val k = (if p = "op <" then abs(m) else m)  
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: 14479
diff changeset
   333
           val xtm = (HOLogic.mk_binop "op *" ((mk_numeral ((m div k)*l) ), x))
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: 14479
diff changeset
   334
           val rs = (HOLogic.mk_binrel p ((linear_cmul k a),(HOLogic.mk_binop "op +" ( xtm ,( linear_cmul k t) )))) 
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: 14479
diff changeset
   335
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: 14479
diff changeset
   336
           val ck = cterm_of sg (mk_numeral k)
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: 14479
diff changeset
   337
           val cc = cterm_of sg c
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: 14479
diff changeset
   338
           val ct = cterm_of sg t
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: 14479
diff changeset
   339
           val cx = cterm_of sg x
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: 14479
diff changeset
   340
           val ca = cterm_of sg a
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: 14479
diff changeset
   341
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: 14479
diff changeset
   342
	   in 
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: 14479
diff changeset
   343
	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: 14479
diff changeset
   344
	  "op <" => 
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: 14479
diff changeset
   345
	let val pre = prove_elementar sg "lf" 
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: 14479
diff changeset
   346
	    (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),(mk_numeral k)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   347
            val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_lt_eq)))
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: 14479
diff changeset
   348
	in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] 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: 14479
diff changeset
   349
         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: 14479
diff changeset
   350
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: 14479
diff changeset
   351
           |"op =" =>
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: 14479
diff changeset
   352
	     let val pre = prove_elementar sg "lf" 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   353
	    (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   354
	         val th1 = (pre RS(instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct] (ac_eq_eq)))
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: 14479
diff changeset
   355
	     in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] 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: 14479
diff changeset
   356
             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: 14479
diff changeset
   357
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: 14479
diff changeset
   358
             |"Divides.op dvd" =>
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: 14479
diff changeset
   359
	       let val pre = prove_elementar sg "lf" 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   360
	   (HOLogic.Not $ (HOLogic.mk_binrel "op =" (Const("0",HOLogic.intT),(mk_numeral k))))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   361
                   val th1 = (pre RS (instantiate' [] [SOME ck,SOME ca,SOME cc, SOME cx, SOME ct]) (ac_dvd_eq))
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: 14479
diff changeset
   362
               in ([], fn [] => [th1,(prove_elementar sg "lf" (HOLogic.mk_eq (snd (qe_get_terms th1) ,rs)))] MRS trans)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   363
                        
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: 14479
diff changeset
   364
               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: 14479
diff changeset
   365
              end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   366
  else ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] 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: 14479
diff changeset
   367
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: 14479
diff changeset
   368
 |( 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: 14479
diff changeset
   369
  |( Const ("op &",_) $ p $ q) => ([p,q], 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: 14479
diff changeset
   370
  |( Const ("op |",_) $ p $ q) =>([p,q], fn [th1,th2] => [th1,th2] MRS qe_disjI)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   371
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   372
  |_ => ([], fn [] => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] refl);
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   373
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   374
fun proof_of_adjustcoeffeq sg x l = thm_of sg (decomp_adjustcoeffeq sg x l);
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   375
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   376
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   377
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: 14479
diff changeset
   378
(*==================================================*)
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: 14479
diff changeset
   379
(*   Finding rho for modd_minusinfinity             *)
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: 14479
diff changeset
   380
(*==================================================*)
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: 14479
diff changeset
   381
fun rho_for_modd_minf x dlcm sg fm1 =
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: 14479
diff changeset
   382
let
15661
9ef583b08647 reverted renaming of Some/None in comments and strings;
wenzelm
parents: 15531
diff changeset
   383
    (*Some certified Terms*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   384
    
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   385
   val ctrue = cterm_of sg HOLogic.true_const
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   386
   val cfalse = cterm_of sg HOLogic.false_const
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   387
   val fm = norm_zero_one fm1
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   388
  in  case fm1 of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   389
      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   390
         if (x=y) andalso (c1= zero) andalso (c2= one) then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   391
           else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   392
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   393
      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   394
  	   if (is_arith_rel fm) andalso (x=y) andalso (c1= zero) andalso (c2= one) 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   395
	   then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   396
	 	 else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)) 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   397
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   398
      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   399
           if (y=x) andalso (c1 = zero) then 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   400
            if (pm1 = one) then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_minf)) else
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   401
	     (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_minf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   402
	    else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   403
  
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   404
      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   405
         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   406
			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   407
	 	      in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_minf)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   408
		      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   409
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   410
      |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   411
      c $ y ) $ z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   412
         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   413
			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   414
	 	      in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_minf)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   415
		      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   416
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   417
		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   418
    
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   419
   |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_minf)
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: 14479
diff changeset
   420
   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: 14479
diff changeset
   421
(*=========================================================================*)
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: 14479
diff changeset
   422
(*=========================================================================*)
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: 14479
diff changeset
   423
fun rho_for_eq_minf x dlcm  sg fm1 =  
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: 14479
diff changeset
   424
   let
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   425
   val fm = norm_zero_one fm1
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   426
    in  case fm1 of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   427
      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   428
         if  (x=y) andalso (c1=zero) andalso (c2=one) 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   429
	   then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_minf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   430
           else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   431
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   432
      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   433
  	   if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   434
	     then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_minf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   435
	     else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf)) 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   436
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   437
      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   438
           if (y=x) andalso (c1 =zero) then 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   439
            if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_minf)) else
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   440
	     (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_minf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   441
	    else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   442
      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   443
         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   444
	 		  val cz = cterm_of sg (norm_zero_one z)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   445
	 	      in(instantiate' [] [SOME cd,  SOME cz] (not_dvd_eq_minf)) 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   446
		      end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   447
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   448
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   449
		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   450
      |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   451
         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   452
	 		  val cz = cterm_of sg (norm_zero_one z)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   453
	 	      in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   454
		      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   455
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   456
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   457
      		
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   458
    |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_minf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   459
 end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   460
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: 14479
diff changeset
   461
(*=====================================================*)
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: 14479
diff changeset
   462
(*=====================================================*)
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: 14479
diff changeset
   463
(*=========== minf proofs with the compact version==========*)
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: 14479
diff changeset
   464
fun decomp_minf_eq x dlcm sg t =  case t 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: 14479
diff changeset
   465
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_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: 14479
diff changeset
   466
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_minf_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: 14479
diff changeset
   467
   |_ => ([],fn [] => rho_for_eq_minf x dlcm sg t);
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   468
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: 14479
diff changeset
   469
fun decomp_minf_modd x dlcm sg t = case t 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: 14479
diff changeset
   470
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_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: 14479
diff changeset
   471
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_minf_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: 14479
diff changeset
   472
   |_ => ([],fn [] => rho_for_modd_minf x dlcm sg t);
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   473
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: 14479
diff changeset
   474
(* -------------------------------------------------------------*)
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: 14479
diff changeset
   475
(*                    Finding rho for pinf_modd                 *)
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: 14479
diff changeset
   476
(* -------------------------------------------------------------*)
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: 14479
diff changeset
   477
fun rho_for_modd_pinf x dlcm sg fm1 = 
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: 14479
diff changeset
   478
let
15661
9ef583b08647 reverted renaming of Some/None in comments and strings;
wenzelm
parents: 15531
diff changeset
   479
    (*Some certified Terms*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   480
    
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   481
  val ctrue = cterm_of sg HOLogic.true_const
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   482
  val cfalse = cterm_of sg HOLogic.false_const
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   483
  val fm = norm_zero_one fm1
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   484
 in  case fm1 of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   485
      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   486
         if ((x=y) andalso (c1= zero) andalso (c2= one))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   487
	 then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   488
         else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   489
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   490
      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   491
  	if ((is_arith_rel fm) andalso (x = y) andalso (c1 = zero)  andalso (c2 = one)) 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   492
	then (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   493
	else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   494
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   495
      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   496
        if ((y=x) andalso (c1 = zero)) then 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   497
          if (pm1 = one) 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   498
	  then (instantiate' [SOME cboolT] [SOME ctrue] (fm_modd_pinf)) 
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   499
	  else (instantiate' [SOME cboolT] [SOME cfalse] (fm_modd_pinf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   500
	else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   501
  
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   502
      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   503
         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   504
			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   505
	 	      in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS(((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (not_dvd_modd_pinf)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   506
		      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   507
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   508
      |(Const("Divides.op dvd",_)$ d $ (db as (Const ("op +",_) $ (Const ("op *",_) $
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   509
      c $ y ) $ z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   510
         if y=x then  let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   511
			  val fm2 = HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   512
	 	      in(instantiate' [] [SOME cz ] ((((prove_elementar sg "ss" fm2)) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1) ) RS (dvd_modd_pinf)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   513
		      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   514
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   515
		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   516
    
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   517
   |_ => instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_modd_pinf)
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: 14479
diff changeset
   518
   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: 14479
diff changeset
   519
(* -------------------------------------------------------------*)
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: 14479
diff changeset
   520
(*                    Finding rho for pinf_eq                 *)
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: 14479
diff changeset
   521
(* -------------------------------------------------------------*)
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: 14479
diff changeset
   522
fun rho_for_eq_pinf x dlcm sg fm1 = 
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: 14479
diff changeset
   523
  let
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   524
					val fm = norm_zero_one fm1
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   525
    in  case fm1 of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   526
      (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   527
         if  (x=y) andalso (c1=zero) andalso (c2=one) 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   528
	   then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (neq_eq_pinf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   529
           else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   530
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   531
      |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z)) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   532
  	   if (is_arith_rel fm) andalso (x=y) andalso ((c1=zero) orelse (c1 = norm_zero_one zero)) andalso ((c2=one) orelse (c1 = norm_zero_one one))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   533
	     then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (eq_eq_pinf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   534
	     else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf)) 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   535
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   536
      |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   537
           if (y=x) andalso (c1 =zero) then 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   538
            if pm1 = one then (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (le_eq_pinf)) else
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   539
	     (instantiate' [] [SOME (cterm_of sg (norm_zero_one z))] (len_eq_pinf))
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   540
	    else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   541
      |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   542
         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   543
	 		  val cz = cterm_of sg (norm_zero_one z)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   544
	 	      in(instantiate' [] [SOME cd,  SOME cz] (not_dvd_eq_pinf)) 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   545
		      end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   546
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   547
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   548
		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   549
      |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   550
         if y=x then  let val cd = cterm_of sg (norm_zero_one d)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   551
	 		  val cz = cterm_of sg (norm_zero_one z)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   552
	 	      in(instantiate' [] [SOME cd, SOME cz ] (dvd_eq_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   553
		      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   554
		else (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   555
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   556
      		
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   557
    |_ => (instantiate' [SOME cboolT] [SOME (cterm_of sg fm)] (fm_eq_pinf))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   558
 end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   559
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   560
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: 14479
diff changeset
   561
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: 14479
diff changeset
   562
fun  minf_proof_of_c sg x dlcm t =
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: 14479
diff changeset
   563
  let val minf_eqth   = thm_of sg (decomp_minf_eq x dlcm sg) t
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: 14479
diff changeset
   564
      val minf_moddth = thm_of sg (decomp_minf_modd x dlcm sg) t
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: 14479
diff changeset
   565
  in (minf_eqth, minf_moddth)
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: 14479
diff changeset
   566
end;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   567
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: 14479
diff changeset
   568
(*=========== pinf proofs with the compact version==========*)
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: 14479
diff changeset
   569
fun decomp_pinf_eq x dlcm sg t = case t 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: 14479
diff changeset
   570
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_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: 14479
diff changeset
   571
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS eq_pinf_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: 14479
diff changeset
   572
   |_ =>([],fn [] => rho_for_eq_pinf x dlcm sg t) ;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   573
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: 14479
diff changeset
   574
fun decomp_pinf_modd x dlcm sg t =  case t 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: 14479
diff changeset
   575
   Const ("op &",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_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: 14479
diff changeset
   576
   |Const ("op |",_) $ p $q => ([p,q],fn [th1,th2] => [th1,th2] MRS modd_pinf_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: 14479
diff changeset
   577
   |_ => ([],fn [] => rho_for_modd_pinf x dlcm sg t);
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: 14479
diff changeset
   578
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: 14479
diff changeset
   579
fun  pinf_proof_of_c sg x dlcm t =
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: 14479
diff changeset
   580
  let val pinf_eqth   = thm_of sg (decomp_pinf_eq x dlcm sg) t
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: 14479
diff changeset
   581
      val pinf_moddth = thm_of sg (decomp_pinf_modd x dlcm sg) t
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: 14479
diff changeset
   582
  in (pinf_eqth,pinf_moddth)
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: 14479
diff changeset
   583
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: 14479
diff changeset
   584
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   585
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   586
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   587
(* Here we generate the theorem for the Bset Property in the simple direction*)
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: 14479
diff changeset
   588
(* It is just an instantiation*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   589
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   590
(*
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: 14479
diff changeset
   591
fun bsetproof_of sg (x as Free(xn,xT)) fm bs dlcm   = 
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: 14479
diff changeset
   592
  let
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: 14479
diff changeset
   593
    val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
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: 14479
diff changeset
   594
    val cdlcm = cterm_of sg dlcm
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: 14479
diff changeset
   595
    val cB = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one bs))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   596
  in instantiate' [] [SOME cdlcm,SOME cB, SOME cp] (bst_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: 14479
diff changeset
   597
end;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   598
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: 14479
diff changeset
   599
fun asetproof_of sg (x as Free(xn,xT)) fm ast dlcm = 
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: 14479
diff changeset
   600
  let
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: 14479
diff changeset
   601
    val cp = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
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: 14479
diff changeset
   602
    val cdlcm = cterm_of sg dlcm
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: 14479
diff changeset
   603
    val cA = cterm_of sg (list_to_set HOLogic.intT (map norm_zero_one ast))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   604
  in instantiate' [] [SOME cdlcm,SOME cA, SOME cp] (ast_thm)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   605
end;
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: 14479
diff changeset
   606
*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   607
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   608
(* For the generation of atomic Theorems*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   609
(* Prove the premisses on runtime and then make RS*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   610
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   611
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: 14479
diff changeset
   612
(*========= this is rho ============*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   613
fun generate_atomic_not_bst_p sg (x as Free(xn,xT)) fm dlcm B at = 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   614
  let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   615
    val cdlcm = cterm_of sg dlcm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   616
    val cB = cterm_of sg B
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   617
    val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   618
    val cat = cterm_of sg (norm_zero_one at)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   619
  in
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   620
  case at of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   621
   (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   622
      if  (x=y) andalso (c1=zero) andalso (c2=one) 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   623
	 then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   624
	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   625
		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   626
	 in  (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_bst_p_ne)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   627
	 end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   628
         else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   629
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   630
   |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   631
     if (is_arith_rel at) andalso (x=y)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   632
	then let val bst_z = norm_zero_one (linear_neg (linear_add [] z (mk_numeral 1)))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   633
	         in let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ bst_z $ B)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   634
	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (bst_z,Const("op -",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   635
		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   636
	 in  (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_bst_p_eq)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   637
	 end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   638
       end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   639
         else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   640
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   641
   |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   642
        if (y=x) andalso (c1 =zero) then 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   643
        if pm1 = one then 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   644
	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ B)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   645
              val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   646
	  in  (instantiate' [] [SOME cfma,  SOME cdlcm]([th1,th2] MRS (not_bst_p_gt)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   647
	    end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   648
	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   649
	      in (instantiate' [] [SOME cfma, SOME cB,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_bst_p_lt)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   650
	      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   651
      else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   652
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   653
   |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   654
      if y=x then  
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   655
           let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   656
	       val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   657
 	     in (instantiate' []  [SOME cfma, SOME cB,SOME cz] (th1 RS (not_bst_p_ndvd)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   658
	     end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   659
      else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   660
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   661
   |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   662
       if y=x then  
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   663
	 let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   664
	     val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   665
 	    in (instantiate' []  [SOME cfma,SOME cB,SOME cz] (th1 RS (not_bst_p_dvd)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   666
	  end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   667
      else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   668
      		
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   669
   |_ => (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cB,SOME cat] (not_bst_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   670
      		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   671
    end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   672
    
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: 14479
diff changeset
   673
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   674
(* ------------------------------------------------------------------------- *)    
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   675
(* Main interpretation function for this backwards dirction*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   676
(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   677
(*Help Function*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   678
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   679
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: 14479
diff changeset
   680
(*==================== Proof with the compact version   *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   681
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: 14479
diff changeset
   682
fun decomp_nbstp sg x dlcm B fm t = case t 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: 14479
diff changeset
   683
   Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_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: 14479
diff changeset
   684
  |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_bst_p_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: 14479
diff changeset
   685
  |_ => ([], fn [] => generate_atomic_not_bst_p sg x fm dlcm B t);
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: 14479
diff changeset
   686
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: 14479
diff changeset
   687
fun not_bst_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm B t =
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: 14479
diff changeset
   688
  let 
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: 14479
diff changeset
   689
       val th =  thm_of sg (decomp_nbstp sg x dlcm (list_to_set xT (map norm_zero_one B)) fm) t
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   690
      val fma = absfree (xn,xT, norm_zero_one fm)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   691
  in let val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   692
     in [th,th1] MRS (not_bst_p_Q_elim)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   693
     end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   694
  end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   695
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   696
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   697
(* ------------------------------------------------------------------------- *)    
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   698
(* Protokol interpretation function for the backwards direction for cooper's Theorem*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   699
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   700
(* For the generation of atomic Theorems*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   701
(* Prove the premisses on runtime and then make RS*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   702
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   703
(*========= this is rho ============*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   704
fun generate_atomic_not_ast_p sg (x as Free(xn,xT)) fm dlcm A at = 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   705
  let
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   706
    val cdlcm = cterm_of sg dlcm
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   707
    val cA = cterm_of sg A
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   708
    val cfma = cterm_of sg (absfree (xn,xT,(norm_zero_one fm)))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   709
    val cat = cterm_of sg (norm_zero_one at)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   710
  in
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   711
  case at of 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   712
   (Const ("Not", _) $ (Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $ (Const ("op +", _) $(Const ("op *",_) $ c2 $ y) $z))) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   713
      if  (x=y) andalso (c1=zero) andalso (c2=one) 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   714
	 then let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one (linear_cmul ~1 z)) $ A)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   715
	          val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq ((norm_zero_one (linear_cmul ~1 z)),Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one  z)))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   716
		  val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   717
	 in  (instantiate' [] [SOME cfma]([th3,th1,th2] MRS (not_ast_p_ne)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   718
	 end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   719
         else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   720
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   721
   |(Const("op =",Type ("fun",[Type ("IntDef.int", []),_])) $ c1 $(Const ("op +", T) $(Const ("op *",_) $ c2 $ y) $z)) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   722
     if (is_arith_rel at) andalso (x=y)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   723
	then let val ast_z = norm_zero_one (linear_sub [] one z )
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   724
	         val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ ast_z $ A)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   725
	         val th2 =  prove_elementar sg "ss" (HOLogic.mk_eq (ast_z,Const("op +",T) $ (Const("uminus",HOLogic.intT --> HOLogic.intT) $(norm_zero_one z)) $ (Const("1",HOLogic.intT))))
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   726
		 val th3 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   727
	 in  (instantiate' [] [SOME cfma] ([th3,th1,th2] MRS (not_ast_p_eq)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   728
       end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   729
         else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   730
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   731
   |(Const("op <",_) $ c1 $(Const ("op +", _) $(Const ("op *",_) $ pm1 $ y ) $ z )) =>
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   732
        if (y=x) andalso (c1 =zero) then 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   733
        if pm1 = (mk_numeral ~1) then 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   734
	  let val th1 = prove_elementar sg "ss" (Const ("op :",HOLogic.intT --> (HOLogic.mk_setT HOLogic.intT) --> HOLogic.boolT) $ (norm_zero_one z) $ A)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   735
              val th2 =  prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   736
	  in  (instantiate' [] [SOME cfma]([th2,th1] MRS (not_ast_p_lt)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   737
	    end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   738
	 else let val th1 = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (Const("0",HOLogic.intT),dlcm))
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   739
	      in (instantiate' [] [SOME cfma, SOME cA,SOME (cterm_of sg (norm_zero_one z))] (th1 RS (not_ast_p_gt)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   740
	      end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   741
      else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   742
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   743
   |Const ("Not",_) $ (Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   744
      if y=x then  
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   745
           let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   746
	       val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   747
 	     in (instantiate' []  [SOME cfma, SOME cA,SOME cz] (th1 RS (not_ast_p_ndvd)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   748
	     end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   749
      else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   750
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   751
   |(Const("Divides.op dvd",_)$ d $ (Const ("op +",_) $ (Const ("op *",_) $ c $ y ) $ z)) => 
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   752
       if y=x then  
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   753
	 let val cz = cterm_of sg (norm_zero_one z)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   754
	     val th1 = (prove_elementar sg "ss"  (HOLogic.mk_binrel "op =" (HOLogic.mk_binop "Divides.op mod" (dlcm,d),norm_zero_one zero))) RS (((zdvd_iff_zmod_eq_0)RS sym) RS iffD1)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   755
 	    in (instantiate' []  [SOME cfma,SOME cA,SOME cz] (th1 RS (not_ast_p_dvd)))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   756
	  end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   757
      else (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   758
      		
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   759
   |_ => (instantiate' [] [SOME cfma,  SOME cdlcm, SOME cA,SOME cat] (not_ast_p_fm))
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   760
      		
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   761
    end;
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: 14479
diff changeset
   762
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: 14479
diff changeset
   763
(* ------------------------------------------------------------------------ *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   764
(* Main interpretation function for this backwards dirction*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   765
(* if atomic do generate atomis formulae else Construct theorems and then make RS with the construction theorems*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   766
(*Help Function*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   767
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   768
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: 14479
diff changeset
   769
fun decomp_nastp sg x dlcm A fm t = case t 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: 14479
diff changeset
   770
   Const("op &",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_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: 14479
diff changeset
   771
  |Const("op |",_) $ ls $ rs => ([ls,rs],fn [th1,th2] => [th1,th2] MRS not_ast_p_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: 14479
diff changeset
   772
  |_ => ([], fn [] => generate_atomic_not_ast_p sg x fm dlcm A t);
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   773
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: 14479
diff changeset
   774
fun not_ast_p_proof_of_c sg (x as Free(xn,xT)) fm dlcm A t =
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: 14479
diff changeset
   775
  let 
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: 14479
diff changeset
   776
       val th =  thm_of sg (decomp_nastp sg x dlcm (list_to_set xT (map norm_zero_one A)) fm) t
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   777
      val fma = absfree (xn,xT, norm_zero_one fm)
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: 14479
diff changeset
   778
  in let val th1 =  prove_elementar sg "ss"  (HOLogic.mk_eq (fma,fma))
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: 14479
diff changeset
   779
     in [th,th1] MRS (not_ast_p_Q_elim)
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: 14479
diff changeset
   780
     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: 14479
diff changeset
   781
  end;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   782
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   783
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: 14479
diff changeset
   784
(* -------------------------------*)
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: 14479
diff changeset
   785
(* Finding rho and beta for evalc *)
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: 14479
diff changeset
   786
(* -------------------------------*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   787
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: 14479
diff changeset
   788
fun rho_for_evalc sg at = case at 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: 14479
diff changeset
   789
    (Const (p,_) $ s $ t) =>(  
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: 14479
diff changeset
   790
    case assoc (operations,p) of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   791
        SOME f => 
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: 14479
diff changeset
   792
           ((if (f ((dest_numeral s),(dest_numeral t))) 
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: 14479
diff changeset
   793
             then prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)) 
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: 14479
diff changeset
   794
             else prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const)))  
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   795
		   handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   796
        | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] 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: 14479
diff changeset
   797
     |Const("Not",_)$(Const (p,_) $ s $ t) =>(  
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: 14479
diff changeset
   798
       case assoc (operations,p) of 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   799
         SOME f => 
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: 14479
diff changeset
   800
           ((if (f ((dest_numeral s),(dest_numeral t))) 
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: 14479
diff changeset
   801
             then prove_elementar sg "ss" (HOLogic.mk_eq(at, HOLogic.false_const))  
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: 14479
diff changeset
   802
             else prove_elementar sg "ss" (HOLogic.mk_eq(at,HOLogic.true_const)))  
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   803
		      handle _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl) 
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   804
         | _ => instantiate' [SOME cboolT] [SOME (cterm_of sg at)] refl ) 
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   805
     | _ =>   instantiate' [SOME cboolT] [SOME (cterm_of sg at)] 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: 14479
diff changeset
   806
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: 14479
diff changeset
   807
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: 14479
diff changeset
   808
(*=========================================================*)
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: 14479
diff changeset
   809
fun decomp_evalc sg t = case t 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: 14479
diff changeset
   810
   (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: 14479
diff changeset
   811
   |(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: 14479
diff changeset
   812
   |(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: 14479
diff changeset
   813
   |(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: 14479
diff changeset
   814
   |_ => ([], fn [] => rho_for_evalc sg t);
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: 14479
diff changeset
   815
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: 14479
diff changeset
   816
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: 14479
diff changeset
   817
fun proof_of_evalc sg fm = thm_of sg (decomp_evalc sg) fm;
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: 14479
diff changeset
   818
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: 14479
diff changeset
   819
(*==================================================*)
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: 14479
diff changeset
   820
(*     Proof of linform with the compact model      *)
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: 14479
diff changeset
   821
(*==================================================*)
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: 14479
diff changeset
   822
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: 14479
diff changeset
   823
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: 14479
diff changeset
   824
fun decomp_linform sg vars t = case t 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: 14479
diff changeset
   825
   (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: 14479
diff changeset
   826
   |(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: 14479
diff changeset
   827
   |(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: 14479
diff changeset
   828
   |(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: 14479
diff changeset
   829
   |(Const("Not",_)$p) => ([p],fn [th] => th RS qe_Not)
15164
5d7c96e0f9dc m dvd t where m is non numeral is now catched!
chaieb
parents: 15123
diff changeset
   830
   |(Const("Divides.op dvd",_)$d$r) => 
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   831
     if is_numeral d then ([], fn [] => (prove_elementar sg "lf" (HOLogic.mk_eq (r, lint vars r))) RS (instantiate' [] [NONE , NONE, SOME (cterm_of sg d)](linearize_dvd)))
15164
5d7c96e0f9dc m dvd t where m is non numeral is now catched!
chaieb
parents: 15123
diff changeset
   832
     else (warning "Nonlinear Term --- Non numeral leftside at dvd";
5d7c96e0f9dc m dvd t where m is non numeral is now catched!
chaieb
parents: 15123
diff changeset
   833
       raise COOPER)
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: 14479
diff changeset
   834
   |_ => ([], fn [] => prove_elementar sg "lf" (HOLogic.mk_eq (t, linform vars t)));
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: 14479
diff changeset
   835
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: 14479
diff changeset
   836
fun proof_of_linform sg vars f = thm_of sg (decomp_linform sg vars) f;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   837
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   838
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   839
(* Interpretaion of Protocols of the cooper procedure : minusinfinity version*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   840
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   841
fun coopermi_proof_of sg (x as Free(xn,xT)) fm B dlcm =
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   842
  (* Get the Bset 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: 14479
diff changeset
   843
  let val (minf_eqth, minf_moddth) = minf_proof_of_c sg x dlcm fm 
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   844
      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
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: 14479
diff changeset
   845
      val nbstpthm = not_bst_p_proof_of_c sg x fm dlcm B fm
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: 14479
diff changeset
   846
  in (dpos,minf_eqth,nbstpthm,minf_moddth)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   847
end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   848
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   849
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   850
(* Interpretaion of Protocols of the cooper procedure : plusinfinity version *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   851
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   852
fun cooperpi_proof_of sg (x as Free(xn,xT)) fm A dlcm =
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: 14479
diff changeset
   853
  let val (pinf_eqth,pinf_moddth) = pinf_proof_of_c sg x dlcm fm
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   854
      val dpos = prove_elementar sg "ss" (HOLogic.mk_binrel "op <" (zero,dlcm));
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: 14479
diff changeset
   855
      val nastpthm = not_ast_p_proof_of_c sg x fm dlcm A fm
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: 14479
diff changeset
   856
  in (dpos,pinf_eqth,nastpthm,pinf_moddth)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   857
end;
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   858
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   859
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   860
(* Interpretaion of Protocols of the cooper procedure : full version*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   861
(* ------------------------------------------------------------------------- *)
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: 14479
diff changeset
   862
fun cooper_thm sg s (x as Free(xn,xT)) cfm dlcm ast bst= case s 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: 14479
diff changeset
   863
  "pi" => let val (dpsthm,pinf_eqth,nbpth,pinf_moddth) = cooperpi_proof_of sg x cfm ast dlcm 
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: 14479
diff changeset
   864
	      in [dpsthm,pinf_eqth,nbpth,pinf_moddth] MRS (cppi_eq)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   865
           end
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: 14479
diff changeset
   866
  |"mi" => let val (dpsthm,minf_eqth,nbpth,minf_moddth) = coopermi_proof_of sg x cfm bst dlcm
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: 14479
diff changeset
   867
	       in [dpsthm,minf_eqth,nbpth,minf_moddth] MRS (cpmi_eq)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   868
                end
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   869
 |_ => error "parameter error";
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   870
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   871
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   872
(* This function should evoluate to the end prove Procedure for one quantifier elimination for Presburger arithmetic*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   873
(* It shoud be plugged in the qfnp argument of the quantifier elimination proof function*)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   874
(* ------------------------------------------------------------------------- *)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   875
15165
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   876
(* val (timef:(unit->thm) -> thm,prtime,time_reset) = gen_timer();*)
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   877
(* val (timef2:(unit->thm) -> thm,prtime2,time_reset2) = gen_timer(); *)
15164
5d7c96e0f9dc m dvd t where m is non numeral is now catched!
chaieb
parents: 15123
diff changeset
   878
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: 14479
diff changeset
   879
fun cooper_prv sg (x as Free(xn,xT)) efm = let 
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   880
   (* lfm_thm : efm = linearized form of efm*)
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: 14479
diff changeset
   881
   val lfm_thm = proof_of_linform sg [xn] efm
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   882
   (*efm2 is the linearized form of efm *) 
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: 14479
diff changeset
   883
   val efm2 = snd(qe_get_terms lfm_thm)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   884
   (* l is the lcm of all coefficients of x *)
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: 14479
diff changeset
   885
   val l = formlcm x efm2
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   886
   (*ac_thm: efm = efm2 with adjusted coefficients of x *)
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   887
   val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   888
   (* fm is efm2 with adjusted coefficients of x *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   889
   val fm = snd (qe_get_terms ac_thm)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   890
  (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   891
   val  cfm = unitycoeff x fm
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   892
   (*afm is fm where c*x is replaced by 1*x or -1*x *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   893
   val afm = adjustcoeff x l fm
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   894
   (* P = %x.afm*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   895
   val P = absfree(xn,xT,afm)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   896
   (* This simpset allows the elimination of the sets in bex {1..d} *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   897
   val ss = presburger_ss addsimps
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   898
     [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   899
   (* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   900
   val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   901
   (* e_ac_thm : Ex x. efm = EX x. fm*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   902
   val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   903
   (* A and B set of the formula*)
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: 14479
diff changeset
   904
   val A = aset x cfm
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: 14479
diff changeset
   905
   val B = bset x cfm
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   906
   (* the divlcm (delta) of the formula*)
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: 14479
diff changeset
   907
   val dlcm = mk_numeral (divlcm x cfm)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   908
   (* Which set is smaller to generate the (hoepfully) shorter proof*)
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: 14479
diff changeset
   909
   val cms = if ((length A) < (length B )) then "pi" else "mi"
15165
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   910
(*   val _ = if cms = "pi" then writeln "Plusinfinity" else writeln "Minusinfinity"*)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   911
   (* synthesize the proof of cooper's theorem*)
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   912
    (* cp_thm: EX x. cfm = Q*)
15165
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   913
   val cp_thm =  cooper_thm sg cms x cfm dlcm A B
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   914
   (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   915
   (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*)
15165
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   916
(*
15164
5d7c96e0f9dc m dvd t where m is non numeral is now catched!
chaieb
parents: 15123
diff changeset
   917
   val _ = prth cp_thm
5d7c96e0f9dc m dvd t where m is non numeral is now catched!
chaieb
parents: 15123
diff changeset
   918
   val _ = writeln "Expanding the bounded EX..."
15165
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   919
*)
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   920
   val exp_cp_thm = refl RS (simplify ss (cp_thm RSN (2,trans)))
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   921
(*
a1e84e86c583 corrected
chaieb
parents: 15164
diff changeset
   922
   val _ = writeln "Expanded" *)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   923
   (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   924
   val (lsuth,rsuth) = qe_get_terms (uth)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   925
   (* lseacth = EX x. efm; rseacth = EX x. fm*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   926
   val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   927
   (* lscth = EX x. cfm; rscth = Q' *)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   928
   val (lscth,rscth) = qe_get_terms (exp_cp_thm)
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   929
   (* u_c_thm: EX x. P(l*x) = Q'*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   930
   val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
14877
28084696907f More readable code.
chaieb
parents: 14758
diff changeset
   931
   (* result: EX x. efm = Q'*)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   932
 in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
   933
   end
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: 14479
diff changeset
   934
|cooper_prv _ _ _ =  error "Parameters format";
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: 14479
diff changeset
   935
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   936
(* **************************************** *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   937
(*    An Other Version of cooper proving    *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   938
(*     by giving a withness for EX          *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   939
(* **************************************** *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   940
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   941
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   942
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   943
fun cooper_prv_w sg (x as Free(xn,xT)) efm = let 
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   944
   (* lfm_thm : efm = linearized form of efm*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   945
   val lfm_thm = proof_of_linform sg [xn] efm
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   946
   (*efm2 is the linearized form of efm *) 
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   947
   val efm2 = snd(qe_get_terms lfm_thm)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   948
   (* l is the lcm of all coefficients of x *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   949
   val l = formlcm x efm2
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   950
   (*ac_thm: efm = efm2 with adjusted coefficients of x *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   951
   val ac_thm = [lfm_thm , (proof_of_adjustcoeffeq sg x l efm2)] MRS trans
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   952
   (* fm is efm2 with adjusted coefficients of x *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   953
   val fm = snd (qe_get_terms ac_thm)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   954
  (* cfm is l dvd x & fm' where fm' is fm where l*x is replaced by x*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   955
   val  cfm = unitycoeff x fm
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   956
   (*afm is fm where c*x is replaced by 1*x or -1*x *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   957
   val afm = adjustcoeff x l fm
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   958
   (* P = %x.afm*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   959
   val P = absfree(xn,xT,afm)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   960
   (* This simpset allows the elimination of the sets in bex {1..d} *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   961
   val ss = presburger_ss addsimps
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   962
     [simp_from_to] delsimps [P_eqtrue, P_eqfalse, bex_triv, insert_iff]
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   963
   (* uth : EX x.P(l*x) = EX x. l dvd x & P x*)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   964
   val uth = instantiate' [] [SOME (cterm_of sg P) , SOME (cterm_of sg (mk_numeral l))] (unity_coeff_ex)
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   965
   (* e_ac_thm : Ex x. efm = EX x. fm*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   966
   val e_ac_thm = (forall_intr (cterm_of sg x) ac_thm) COMP (qe_exI)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   967
   (* lsuth = EX.P(l*x) ; rsuth = EX x. l dvd x & P x*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   968
   val (lsuth,rsuth) = qe_get_terms (uth)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   969
   (* lseacth = EX x. efm; rseacth = EX x. fm*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   970
   val (lseacth,rseacth) = qe_get_terms(e_ac_thm)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   971
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   972
   val (w,rs) = cooper_w [] cfm
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   973
   val exp_cp_thm =  case w of 
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   974
     (* FIXME - e_ac_thm just tipped to test syntactical correctness of the program!!!!*)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15165
diff changeset
   975
    SOME n =>  e_ac_thm (* Prove cfm (n) and use exI and then Eq_TrueI*)
15107
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   976
   |_ => let 
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   977
    (* A and B set of the formula*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   978
    val A = aset x cfm
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   979
    val B = bset x cfm
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   980
    (* the divlcm (delta) of the formula*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   981
    val dlcm = mk_numeral (divlcm x cfm)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   982
    (* Which set is smaller to generate the (hoepfully) shorter proof*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   983
    val cms = if ((length A) < (length B )) then "pi" else "mi"
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   984
    (* synthesize the proof of cooper's theorem*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   985
     (* cp_thm: EX x. cfm = Q*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   986
    val cp_thm = cooper_thm sg cms x cfm dlcm A B
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   987
     (* Exxpand the right hand side to get rid of EX j : {1..d} to get a huge disjunction*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   988
    (* exp_cp_thm: EX x.cfm = Q' , where Q' is a simplified version of Q*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   989
    in refl RS (simplify ss (cp_thm RSN (2,trans)))
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   990
    end
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   991
   (* lscth = EX x. cfm; rscth = Q' *)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   992
   val (lscth,rscth) = qe_get_terms (exp_cp_thm)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   993
   (* u_c_thm: EX x. P(l*x) = Q'*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   994
   val  u_c_thm = [([uth,prove_elementar sg "ss" (HOLogic.mk_eq (rsuth,lscth))] MRS trans),exp_cp_thm] MRS trans
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   995
   (* result: EX x. efm = Q'*)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   996
 in  ([e_ac_thm,[(prove_elementar sg "ss" (HOLogic.mk_eq (rseacth,lsuth))),u_c_thm] MRS trans] MRS trans)
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   997
   end
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   998
|cooper_prv_w _ _ _ =  error "Parameters format";
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
   999
f233706d9fce oracle corrected
chaieb
parents: 14981
diff changeset
  1000
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1001
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: 14479
diff changeset
  1002
fun decomp_cnnf sg lfnp P = 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: 14479
diff changeset
  1003
     Const ("op &",_) $ p $q => ([p,q] , 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: 14479
diff changeset
  1004
   |Const ("op |",_) $ p $q => ([p,q] , 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: 14479
diff changeset
  1005
   |Const ("Not",_) $ (Const("Not",_) $ p) => ([p], fn [th] => th RS nnf_nn)
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: 14479
diff changeset
  1006
   |Const("Not",_) $ (Const(opn,T) $ p $ q) => 
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: 14479
diff changeset
  1007
     if opn = "op |" 
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: 14479
diff changeset
  1008
      then case (p,q) 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: 14479
diff changeset
  1009
         (A as (Const ("op &",_) $ r $ s),B as (Const ("op &",_) $ r1 $ t)) =>
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: 14479
diff changeset
  1010
          if r1 = negate r 
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: 14479
diff changeset
  1011
          then  ([r,HOLogic.Not$s,r1,HOLogic.Not$t],fn [th1_1,th1_2,th2_1,th2_2] => [[th1_1,th1_1] MRS qe_conjI,[th2_1,th2_2] MRS qe_conjI] MRS nnf_sdj)
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1012
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: 14479
diff changeset
  1013
          else ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj)
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: 14479
diff changeset
  1014
        |(_,_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] => [th1,th2] MRS nnf_ndj)
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: 14479
diff changeset
  1015
      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: 14479
diff changeset
  1016
         case (opn,T) 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: 14479
diff changeset
  1017
           ("op &",_) => ([HOLogic.Not $ p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_ncj )
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: 14479
diff changeset
  1018
           |("op -->",_) => ([p,HOLogic.Not $ q ], fn [th1,th2] =>[th1,th2] MRS nnf_nim )
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: 14479
diff changeset
  1019
           |("op =",Type ("fun",[Type ("bool", []),_])) => 
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: 14479
diff changeset
  1020
           ([HOLogic.conj $ p $ (HOLogic.Not $ q),HOLogic.conj $ (HOLogic.Not $ p) $ q], fn [th1,th2] => [th1,th2] MRS nnf_neq)
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: 14479
diff changeset
  1021
            |(_,_) => ([], fn [] => lfnp 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: 14479
diff changeset
  1022
)
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: 14479
diff changeset
  1023
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: 14479
diff changeset
  1024
   |(Const ("op -->",_) $ p $ q) => ([HOLogic.Not$p,q], fn [th1,th2] => [th1,th2] MRS nnf_im)
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: 14479
diff changeset
  1025
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: 14479
diff changeset
  1026
   |(Const ("op =", Type ("fun",[Type ("bool", []),_])) $ p $ q) =>
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: 14479
diff changeset
  1027
     ([HOLogic.conj $ p $ q,HOLogic.conj $ (HOLogic.Not $ p) $ (HOLogic.Not $ q) ], fn [th1,th2] =>[th1,th2] MRS nnf_eq )
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: 14479
diff changeset
  1028
   |_ => ([], fn [] => lfnp P);
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1029
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1030
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1031
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1032
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: 14479
diff changeset
  1033
fun proof_of_cnnf sg p lfnp = 
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: 14479
diff changeset
  1034
 let val th1 = thm_of sg (decomp_cnnf sg lfnp) 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: 14479
diff changeset
  1035
     val rs = 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: 14479
diff changeset
  1036
     val th2 = prove_elementar sg "ss" (HOLogic.mk_eq(rs,simpl rs))
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: 14479
diff changeset
  1037
  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: 14479
diff changeset
  1038
  end;
13876
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1039
68f4ed8311ac New decision procedure for Presburger arithmetic.
berghofe
parents:
diff changeset
  1040
end;
14920
a7525235e20f An oracle is built in. The tactic will not generate any proofs any more, if the quick_and_dirty flag is set on.
chaieb
parents: 14877
diff changeset
  1041