author | wenzelm |
Sat, 22 Oct 2022 16:56:17 +0200 | |
changeset 76361 | 3b9f36ef7365 |
parent 72966 | f931a2a68ab8 |
permissions | -rw-r--r-- |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
1 |
(* Title: Tools/Argo/argo_proof.ML |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
2 |
Author: Sascha Boehme |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
3 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
4 |
The proof language of the Argo solver. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
5 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
6 |
Proofs trace the inferences of the solver. They can be used to check unsatisfiability results. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
7 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
8 |
The proof language is inspired by: |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
9 |
|
72966 | 10 |
Leonardo de Moura and Nikolaj Bjørner. Proofs and Refutations, and Z3. In |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
11 |
Proceedings of the LPAR 2008 Workshops, Knowledge Exchange: Automated Provers and Proof |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
12 |
Assistants, and the 7th International Workshop on the Implementation of Logics, |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
13 |
volume 418 of CEUR Workshop Proceedings. CEUR-WS.org, 2008. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
14 |
*) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
15 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
16 |
signature ARGO_PROOF = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
17 |
sig |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
18 |
(* types *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
19 |
type proof_id |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
20 |
datatype tautology = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
21 |
Taut_And_1 of int | Taut_And_2 of int * int | Taut_Or_1 of int * int | Taut_Or_2 of int | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
22 |
Taut_Iff_1 | Taut_Iff_2 | Taut_Iff_3 | Taut_Iff_4 | Taut_Ite_Then | Taut_Ite_Else |
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
23 |
datatype side = Left | Right |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
24 |
datatype inequality = Le | Lt |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
25 |
datatype rewrite = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
26 |
Rewr_Not_True | Rewr_Not_False | Rewr_Not_Not | Rewr_Not_And of int | Rewr_Not_Or of int | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
27 |
Rewr_Not_Iff | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
28 |
Rewr_And_False of int | Rewr_And_Dual of int * int | Rewr_And_Sort of int * int list list | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
29 |
Rewr_Or_True of int | Rewr_Or_Dual of int * int | Rewr_Or_Sort of int * int list list | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
30 |
Rewr_Iff_True | Rewr_Iff_False | Rewr_Iff_Not_Not | Rewr_Iff_Refl | Rewr_Iff_Symm | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
31 |
Rewr_Iff_Dual | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
32 |
Rewr_Imp | Rewr_Ite_Prop | Rewr_Ite_True | Rewr_Ite_False | Rewr_Ite_Eq | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
33 |
Rewr_Eq_Refl | Rewr_Eq_Symm | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
34 |
Rewr_Neg | Rewr_Add of (Rat.rat * int option) list * (Rat.rat * int option) list | Rewr_Sub | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
35 |
Rewr_Mul_Nums of Rat.rat * Rat.rat | Rewr_Mul_Zero | Rewr_Mul_One | Rewr_Mul_Comm | |
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
36 |
Rewr_Mul_Assoc of side | Rewr_Mul_Sum of side | Rewr_Mul_Div of side | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
37 |
Rewr_Div_Zero | Rewr_Div_One | Rewr_Div_Nums of Rat.rat * Rat.rat | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
38 |
Rewr_Div_Num of side * Rat.rat | Rewr_Div_Mul of side * Rat.rat | Rewr_Div_Div of side | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
39 |
Rewr_Div_Sum | Rewr_Min_Eq | Rewr_Min_Lt | Rewr_Min_Gt | Rewr_Max_Eq | Rewr_Max_Lt | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
40 |
Rewr_Max_Gt | Rewr_Abs | Rewr_Eq_Nums of bool | Rewr_Eq_Sub | Rewr_Eq_Le | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
41 |
Rewr_Ineq_Nums of inequality * bool | Rewr_Ineq_Add of inequality * Rat.rat | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
42 |
Rewr_Ineq_Sub of inequality | Rewr_Ineq_Mul of inequality * Rat.rat | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
43 |
Rewr_Not_Ineq of inequality |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
44 |
datatype conv = |
66301
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
45 |
Keep_Conv | Then_Conv of conv * conv | Args_Conv of Argo_Expr.kind * conv list | |
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
46 |
Rewr_Conv of rewrite |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
47 |
datatype rule = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
48 |
Axiom of int | Taut of tautology * Argo_Expr.expr | Conjunct of int * int | Rewrite of conv | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
49 |
Hyp of int * Argo_Expr.expr | Clause of int list | Lemma of int list | Unit_Res of int | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
50 |
Refl of Argo_Expr.expr | Symm | Trans | Cong | Subst | Linear_Comb |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
51 |
type proof |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
52 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
53 |
(* equalities and orders *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
54 |
val eq_proof_id: proof_id * proof_id -> bool |
70586 | 55 |
val proof_id_ord: proof_id ord |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
56 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
57 |
(* conversion constructors *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
58 |
val keep_conv: conv |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
59 |
val mk_then_conv: conv -> conv -> conv |
66301
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
60 |
val mk_args_conv: Argo_Expr.kind -> conv list -> conv |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
61 |
val mk_rewr_conv: rewrite -> conv |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
62 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
63 |
(* context *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
64 |
type context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
65 |
val cdcl_context: context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
66 |
val cc_context: context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
67 |
val simplex_context: context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
68 |
val solver_context: context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
69 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
70 |
(* proof constructors *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
71 |
val mk_axiom: int -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
72 |
val mk_taut: tautology -> Argo_Expr.expr -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
73 |
val mk_conj: int -> int -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
74 |
val mk_rewrite: conv -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
75 |
val mk_hyp: Argo_Lit.literal -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
76 |
val mk_clause: Argo_Lit.literal list -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
77 |
val mk_lemma: Argo_Lit.literal list -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
78 |
val mk_unit_res: Argo_Lit.literal -> proof -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
79 |
val mk_refl: Argo_Term.term -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
80 |
val mk_symm: proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
81 |
val mk_trans: proof -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
82 |
val mk_cong: proof -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
83 |
val mk_subst: proof -> proof -> proof -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
84 |
val mk_linear_comb: proof list -> context -> proof * context |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
85 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
86 |
(* proof destructors *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
87 |
val id_of: proof -> proof_id |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
88 |
val dest: proof -> proof_id * rule * proof list |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
89 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
90 |
(* string representations *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
91 |
val string_of_proof_id: proof_id -> string |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
92 |
val string_of_taut: tautology -> string |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
93 |
val string_of_rule: rule -> string |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
94 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
95 |
(* unsatisfiability *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
96 |
exception UNSAT of proof |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
97 |
val unsat: proof -> 'a (* raises UNSAT *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
98 |
end |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
99 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
100 |
structure Argo_Proof: ARGO_PROOF = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
101 |
struct |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
102 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
103 |
(* types *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
104 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
105 |
datatype tautology = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
106 |
Taut_And_1 of int | Taut_And_2 of int * int | Taut_Or_1 of int * int | Taut_Or_2 of int | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
107 |
Taut_Iff_1 | Taut_Iff_2 | Taut_Iff_3 | Taut_Iff_4 | Taut_Ite_Then | Taut_Ite_Else |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
108 |
|
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
109 |
datatype side = Left | Right |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
110 |
|
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
111 |
datatype inequality = Le | Lt |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
112 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
113 |
datatype rewrite = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
114 |
Rewr_Not_True | Rewr_Not_False | Rewr_Not_Not | Rewr_Not_And of int | Rewr_Not_Or of int | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
115 |
Rewr_Not_Iff | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
116 |
Rewr_And_False of int | Rewr_And_Dual of int * int | Rewr_And_Sort of int * int list list | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
117 |
Rewr_Or_True of int | Rewr_Or_Dual of int * int | Rewr_Or_Sort of int * int list list | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
118 |
Rewr_Iff_True | Rewr_Iff_False | Rewr_Iff_Not_Not | Rewr_Iff_Refl | Rewr_Iff_Symm | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
119 |
Rewr_Iff_Dual | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
120 |
Rewr_Imp | Rewr_Ite_Prop | Rewr_Ite_True | Rewr_Ite_False | Rewr_Ite_Eq | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
121 |
Rewr_Eq_Refl | Rewr_Eq_Symm | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
122 |
Rewr_Neg | Rewr_Add of (Rat.rat * int option) list * (Rat.rat * int option) list | Rewr_Sub | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
123 |
Rewr_Mul_Nums of Rat.rat * Rat.rat | Rewr_Mul_Zero | Rewr_Mul_One | Rewr_Mul_Comm | |
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
124 |
Rewr_Mul_Assoc of side | Rewr_Mul_Sum of side | Rewr_Mul_Div of side | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
125 |
Rewr_Div_Zero | Rewr_Div_One | Rewr_Div_Nums of Rat.rat * Rat.rat | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
126 |
Rewr_Div_Num of side * Rat.rat | Rewr_Div_Mul of side * Rat.rat | Rewr_Div_Div of side | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
127 |
Rewr_Div_Sum | Rewr_Min_Eq | Rewr_Min_Lt | Rewr_Min_Gt | Rewr_Max_Eq | Rewr_Max_Lt | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
128 |
Rewr_Max_Gt | Rewr_Abs | Rewr_Eq_Nums of bool | Rewr_Eq_Sub | Rewr_Eq_Le | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
129 |
Rewr_Ineq_Nums of inequality * bool | Rewr_Ineq_Add of inequality * Rat.rat | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
130 |
Rewr_Ineq_Sub of inequality | Rewr_Ineq_Mul of inequality * Rat.rat | |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
131 |
Rewr_Not_Ineq of inequality |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
132 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
133 |
datatype conv = |
66301
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
134 |
Keep_Conv | Then_Conv of conv * conv | Args_Conv of Argo_Expr.kind * conv list | |
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
135 |
Rewr_Conv of rewrite |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
136 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
137 |
datatype rule = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
138 |
Axiom of int | Taut of tautology * Argo_Expr.expr | Conjunct of int * int | Rewrite of conv | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
139 |
Hyp of int * Argo_Expr.expr | Clause of int list | Lemma of int list | Unit_Res of int | |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
140 |
Refl of Argo_Expr.expr | Symm | Trans | Cong | Subst | Linear_Comb |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
141 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
142 |
(* |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
143 |
Proof identifiers are intentially hidden to prevent that functions outside of this structure |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
144 |
are able to build proofs. Proof can hence only be built by the functions provided by |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
145 |
this structure. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
146 |
*) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
147 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
148 |
datatype proof_id = Cdcl of int | Cc of int | Simplex of int | Solver of int |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
149 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
150 |
datatype proof = Proof of proof_id * rule * proof list |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
151 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
152 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
153 |
(* internal functions *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
154 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
155 |
val proof_id_card = 4 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
156 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
157 |
fun raw_proof_id (Cdcl i) = i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
158 |
| raw_proof_id (Cc i) = i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
159 |
| raw_proof_id (Simplex i) = i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
160 |
| raw_proof_id (Solver i) = i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
161 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
162 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
163 |
(* equalities and orders *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
164 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
165 |
fun int_of_proof_id (Cdcl _) = 0 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
166 |
| int_of_proof_id (Cc _) = 1 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
167 |
| int_of_proof_id (Simplex _) = 2 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
168 |
| int_of_proof_id (Solver _) = 3 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
169 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
170 |
fun eq_proof_id (Cdcl i1, Cdcl i2) = (i1 = i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
171 |
| eq_proof_id (Cc i1, Cc i2) = (i1 = i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
172 |
| eq_proof_id (Simplex i1, Simplex i2) = (i1 = i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
173 |
| eq_proof_id (Solver i1, Solver i2) = (i1 = i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
174 |
| eq_proof_id _ = false |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
175 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
176 |
fun proof_id_ord (Cdcl i1, Cdcl i2) = int_ord (i1, i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
177 |
| proof_id_ord (Cc i1, Cc i2) = int_ord (i1, i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
178 |
| proof_id_ord (Simplex i1, Simplex i2) = int_ord (i1, i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
179 |
| proof_id_ord (Solver i1, Solver i2) = int_ord (i1, i2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
180 |
| proof_id_ord (id1, id2) = int_ord (int_of_proof_id id1, int_of_proof_id id2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
181 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
182 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
183 |
(* conversion constructors *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
184 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
185 |
val keep_conv = Keep_Conv |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
186 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
187 |
fun mk_then_conv Keep_Conv c = c |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
188 |
| mk_then_conv c Keep_Conv = c |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
189 |
| mk_then_conv c1 c2 = Then_Conv (c1, c2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
190 |
|
66301
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
191 |
fun mk_args_conv k cs = |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
192 |
if forall (fn Keep_Conv => true | _ => false) cs then Keep_Conv |
66301
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
193 |
else Args_Conv (k, cs) |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
194 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
195 |
fun mk_rewr_conv r = Rewr_Conv r |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
196 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
197 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
198 |
(* context *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
199 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
200 |
(* |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
201 |
The proof context stores the next unused identifier. Incidentally, the same type as |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
202 |
for the proof identifier can be used as context. Every proof-producing module of the |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
203 |
solver has its own proof identifier domain to ensure globally unique identifiers |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
204 |
without sharing a single proof context. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
205 |
*) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
206 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
207 |
type context = proof_id |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
208 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
209 |
val cdcl_context = Cdcl 0 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
210 |
val cc_context = Cc 0 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
211 |
val simplex_context = Simplex 0 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
212 |
val solver_context = Solver 0 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
213 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
214 |
fun next_id (id as Cdcl i) = (id, Cdcl (i + 1)) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
215 |
| next_id (id as Cc i) = (id, Cc (i + 1)) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
216 |
| next_id (id as Simplex i) = (id, Simplex (i + 1)) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
217 |
| next_id (id as Solver i) = (id, Solver (i + 1)) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
218 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
219 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
220 |
(* proof destructors *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
221 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
222 |
fun id_of (Proof (id, _, _)) = id |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
223 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
224 |
fun dest (Proof p) = p |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
225 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
226 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
227 |
(* proof constructors *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
228 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
229 |
fun mk_proof r ps cx = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
230 |
let val (id, cx) = next_id cx |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
231 |
in (Proof (id, r, ps), cx) end |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
232 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
233 |
fun mk_axiom i = mk_proof (Axiom i) [] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
234 |
fun mk_taut t e = mk_proof (Taut (t, e)) [] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
235 |
fun mk_conj i n p = mk_proof (Conjunct (i, n)) [p] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
236 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
237 |
fun mk_rewrite Keep_Conv p cx = (p, cx) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
238 |
| mk_rewrite c p cx = mk_proof (Rewrite c) [p] cx |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
239 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
240 |
fun mk_hyp lit = mk_proof (Hyp (Argo_Lit.signed_id_of lit, Argo_Lit.signed_expr_of lit)) [] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
241 |
fun mk_clause lits p cx = mk_proof (Clause (map Argo_Lit.signed_id_of lits)) [p] cx |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
242 |
fun mk_lemma lits p = mk_proof (Lemma (map Argo_Lit.signed_id_of lits)) [p] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
243 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
244 |
(* |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
245 |
Replay of unit-resolution steps can be optimized if all premises follow a specific form. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
246 |
Therefore, each premise is checked if it is in clausal form. |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
247 |
*) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
248 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
249 |
fun check_clause (p as Proof (_, Clause _, _)) = p |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
250 |
| check_clause (p as Proof (_, Lemma _, _)) = p |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
251 |
| check_clause (p as Proof (_, Unit_Res _, _)) = p |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
252 |
| check_clause _ = raise Fail "bad clause proof" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
253 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
254 |
fun mk_unit t p1 p2 = mk_proof (Unit_Res (Argo_Term.id_of t)) (map check_clause [p1, p2]) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
255 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
256 |
fun mk_unit_res (Argo_Lit.Pos t) p1 p2 = mk_unit t p1 p2 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
257 |
| mk_unit_res (Argo_Lit.Neg t) p1 p2 = mk_unit t p2 p1 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
258 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
259 |
fun mk_refl t = mk_proof (Refl (Argo_Term.expr_of t)) [] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
260 |
fun mk_symm p = mk_proof Symm [p] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
261 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
262 |
fun mk_trans (Proof (_, Refl _, _)) p2 = pair p2 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
263 |
| mk_trans p1 (Proof (_, Refl _, _)) = pair p1 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
264 |
| mk_trans p1 p2 = mk_proof Trans [p1, p2] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
265 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
266 |
fun mk_cong p1 p2 = mk_proof Cong [p1, p2] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
267 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
268 |
fun mk_subst p1 (Proof (_, Refl _, _)) (Proof (_, Refl _, _)) = pair p1 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
269 |
| mk_subst p1 p2 p3 = mk_proof Subst [p1, p2, p3] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
270 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
271 |
fun mk_linear_comb ps = mk_proof Linear_Comb ps |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
272 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
273 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
274 |
(* string representations *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
275 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
276 |
fun string_of_proof_id id = string_of_int (proof_id_card * raw_proof_id id + int_of_proof_id id) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
277 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
278 |
fun string_of_list l r f xs = enclose l r (space_implode ", " (map f xs)) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
279 |
fun parens f xs = string_of_list "(" ")" f xs |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
280 |
fun brackets f xs = string_of_list "[" "]" f xs |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
281 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
282 |
fun string_of_taut (Taut_And_1 n) = "and " ^ string_of_int n |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
283 |
| string_of_taut (Taut_And_2 (i, n)) = "and " ^ parens string_of_int [i, n] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
284 |
| string_of_taut (Taut_Or_1 (i, n)) = "or " ^ parens string_of_int [i, n] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
285 |
| string_of_taut (Taut_Or_2 n) = "or " ^ string_of_int n |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
286 |
| string_of_taut Taut_Iff_1 = "(p1 == p2) | p1 | p2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
287 |
| string_of_taut Taut_Iff_2 = "(p1 == p2) | ~p1 | ~p2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
288 |
| string_of_taut Taut_Iff_3 = "~(p1 == p2) | ~p1 | p2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
289 |
| string_of_taut Taut_Iff_4 = "~(p1 == p2) | p1 | ~p2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
290 |
| string_of_taut Taut_Ite_Then = "~p | (ite p t1 t2) = t1" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
291 |
| string_of_taut Taut_Ite_Else = "p | (ite p t1 t2) = t2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
292 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
293 |
fun string_of_rewr Rewr_Not_True = "~T = F" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
294 |
| string_of_rewr Rewr_Not_False = "~F = T" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
295 |
| string_of_rewr Rewr_Not_Not = "~~p = p" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
296 |
| string_of_rewr (Rewr_Not_And n) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
297 |
"~(and [" ^ string_of_int n ^ "]) = (or [" ^ string_of_int n ^ "])" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
298 |
| string_of_rewr (Rewr_Not_Or n) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
299 |
"~(or [" ^ string_of_int n ^ "]) = (and [" ^ string_of_int n ^ "])" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
300 |
| string_of_rewr Rewr_Not_Iff = "~(p1 == p2) = (~p1 == ~p2)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
301 |
| string_of_rewr (Rewr_And_False i) = "(and ... F(" ^ string_of_int i ^ ") ...) = F" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
302 |
| string_of_rewr (Rewr_And_Dual (i1, i2)) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
303 |
"(and ... p(" ^ string_of_int i1 ^ ") ... ~p(" ^ string_of_int i2 ^ ") ...) = F" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
304 |
| string_of_rewr (Rewr_And_Sort (n, iss)) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
305 |
"(and [" ^ string_of_int n ^ "]) = " ^ |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
306 |
"(and " ^ brackets (brackets string_of_int) iss ^ ")" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
307 |
| string_of_rewr (Rewr_Or_True i) = "(or ... T(" ^ string_of_int i ^ ") ...) = T" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
308 |
| string_of_rewr (Rewr_Or_Dual (i1, i2)) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
309 |
"(or ... p(" ^ string_of_int i1 ^ ") ... ~p(" ^ string_of_int i2 ^ ") ...) = T" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
310 |
| string_of_rewr (Rewr_Or_Sort (n, iss)) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
311 |
"(or [" ^ string_of_int n ^ "]) = " ^ |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
312 |
"(or " ^ brackets (brackets string_of_int) iss ^ ")" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
313 |
| string_of_rewr Rewr_Iff_True = "(p == T) = p" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
314 |
| string_of_rewr Rewr_Iff_False = "(p == F) = ~p" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
315 |
| string_of_rewr Rewr_Iff_Not_Not = "(~p1 == ~p2) = (p1 == p2)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
316 |
| string_of_rewr Rewr_Iff_Refl = "(p == p) = T" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
317 |
| string_of_rewr Rewr_Iff_Symm = "(p1 == p2) = (p2 == p1)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
318 |
| string_of_rewr Rewr_Iff_Dual = "(p == ~p) = F" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
319 |
| string_of_rewr Rewr_Imp = "(p1 --> p2) = (~p1 | p2)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
320 |
| string_of_rewr Rewr_Ite_Prop = "(if p1 p2 p2) = ((~p1 | p2) & (p1 | p3) & (p2 | p3))" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
321 |
| string_of_rewr Rewr_Ite_True = "(if T t1 t2) = t1" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
322 |
| string_of_rewr Rewr_Ite_False = "(if F t1 t2) = t2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
323 |
| string_of_rewr Rewr_Ite_Eq = "(if p t t) = t" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
324 |
| string_of_rewr Rewr_Eq_Refl = "(e = e) = T" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
325 |
| string_of_rewr Rewr_Eq_Symm = "(e1 = e2) = (e2 = e1)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
326 |
| string_of_rewr Rewr_Neg = "-e = -1 * e" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
327 |
| string_of_rewr (Rewr_Add (p1, p2)) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
328 |
let |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
329 |
fun string_of_monom (n, NONE) = Rat.string_of_rat n |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
330 |
| string_of_monom (n, SOME i) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
331 |
(if n = @1 then "" else Rat.string_of_rat n ^ " * ") ^ "e" ^ string_of_int i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
332 |
fun string_of_polynom ms = space_implode " + " (map string_of_monom ms) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
333 |
in string_of_polynom p1 ^ " = " ^ string_of_polynom p2 end |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
334 |
| string_of_rewr Rewr_Sub = "e1 - e2 = e1 + -1 * e2" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
335 |
| string_of_rewr (Rewr_Mul_Nums (n1, n2)) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
336 |
Rat.string_of_rat n1 ^ " * " ^ Rat.string_of_rat n2 ^ " = " ^ Rat.string_of_rat (n1 * n2) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
337 |
| string_of_rewr Rewr_Mul_Zero = "0 * e = 0" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
338 |
| string_of_rewr Rewr_Mul_One = "1 * e = e" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
339 |
| string_of_rewr Rewr_Mul_Comm = "e1 * e2 = e2 * e1" |
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
340 |
| string_of_rewr (Rewr_Mul_Assoc Left) = "(e1 * e2) * e3 = e1 * (e2 * e3)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
341 |
| string_of_rewr (Rewr_Mul_Assoc Right) = "e1 * (n * e2) = (e1 * n) * e2" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
342 |
| string_of_rewr (Rewr_Mul_Sum Left) = "(e1 + ... + em) * e = e1 * e + ... em * e" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
343 |
| string_of_rewr (Rewr_Mul_Sum Right) = "e * (e1 + ... + em) = e * e1 + ... e * em" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
344 |
| string_of_rewr (Rewr_Mul_Div Left) = "(e1 / e2) * e3 = (e1 * e3) / e2" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
345 |
| string_of_rewr (Rewr_Mul_Div Right) = "e1 * (e2 / * e3) = (e1 * e2) / e3" |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
346 |
| string_of_rewr Rewr_Div_Zero = "0 / e = 0" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
347 |
| string_of_rewr Rewr_Div_One = "e / 1 = e" |
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
348 |
| string_of_rewr (Rewr_Div_Nums (n1, n2)) = |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
349 |
Rat.string_of_rat n1 ^ " / " ^ Rat.string_of_rat n2 ^ " = " ^ Rat.string_of_rat (n1 / n2) |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
350 |
| string_of_rewr (Rewr_Div_Num (Left, n)) = |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
351 |
Rat.string_of_rat n ^ " / e = " ^ Rat.string_of_rat n ^ " * (1 / e)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
352 |
| string_of_rewr (Rewr_Div_Num (Right, n)) = |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
353 |
"e / " ^ Rat.string_of_rat n ^ " = " ^ Rat.string_of_rat (Rat.inv n) ^ " * e" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
354 |
| string_of_rewr (Rewr_Div_Mul (Left, n)) = |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
355 |
"(" ^ Rat.string_of_rat n ^ " * e1) / e2 = " ^ Rat.string_of_rat n ^ " * (e1 / e2)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
356 |
| string_of_rewr (Rewr_Div_Mul (Right, n)) = |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
357 |
"e1 / (" ^ Rat.string_of_rat n ^ " * e2) = " ^ Rat.string_of_rat (Rat.inv n) ^ " * (e1 / e2)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
358 |
| string_of_rewr (Rewr_Div_Div Left) = "(e1 / e2) / e3 = e1 / (e2 * e3)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
359 |
| string_of_rewr (Rewr_Div_Div Right) = "e1 / (e2 / e3) = (e1 * e3) / e2" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
360 |
| string_of_rewr Rewr_Div_Sum = "(e1 + ... + em) / e = e1 / e + ... + em / e" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
361 |
| string_of_rewr Rewr_Min_Eq = "min e e = e" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
362 |
| string_of_rewr Rewr_Min_Lt = "min e1 e2 = (if e1 <= e2 then e1 else e2)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
363 |
| string_of_rewr Rewr_Min_Gt = "min e1 e2 = (if e2 <= e1 then e2 else e1)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
364 |
| string_of_rewr Rewr_Max_Eq = "max e e = e" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
365 |
| string_of_rewr Rewr_Max_Lt = "max e1 e2 = (if e1 < e2 then e2 else e1)" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
366 |
| string_of_rewr Rewr_Max_Gt = "max e1 e2 = (if e2 < e1 then e1 else e2)" |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
367 |
| string_of_rewr Rewr_Abs = "abs e = (if 0 <= e then e else -e)" |
64927
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
368 |
| string_of_rewr (Rewr_Eq_Nums true) = "(n1 = n2) = true" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
369 |
| string_of_rewr (Rewr_Eq_Nums false) = "(n1 ~= n2) = false" |
a5a09855e424
less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents:
63960
diff
changeset
|
370 |
| string_of_rewr Rewr_Eq_Sub = "(e1 = e2) = (e1 - e2 = 0)" |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
371 |
| string_of_rewr Rewr_Eq_Le = "(e1 = e2) = (and (e1 <= e2) (e2 <= e1))" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
372 |
| string_of_rewr (Rewr_Ineq_Nums (Le, true)) = "(n1 <= n2) = true" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
373 |
| string_of_rewr (Rewr_Ineq_Nums (Le, false)) = "(n1 <= n2) = false" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
374 |
| string_of_rewr (Rewr_Ineq_Nums (Lt, true)) = "(n1 < n2) = true" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
375 |
| string_of_rewr (Rewr_Ineq_Nums (Lt, false)) = "(n1 < n2) = false" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
376 |
| string_of_rewr (Rewr_Ineq_Add (Le, _)) = "(e1 <= e2) = (e1 + n <= e2 + n)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
377 |
| string_of_rewr (Rewr_Ineq_Add (Lt, _)) = "(e1 < e2) = (e1 + n < e2 + n)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
378 |
| string_of_rewr (Rewr_Ineq_Sub Le) = "(e1 <= e2) = (e1 - e2 <= 0)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
379 |
| string_of_rewr (Rewr_Ineq_Sub Lt) = "(e1 < e2) = (e1 - e2 < 0)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
380 |
| string_of_rewr (Rewr_Ineq_Mul (Le, _)) = "(e1 <= e2) = (n * e1 <= n * e2)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
381 |
| string_of_rewr (Rewr_Ineq_Mul (Lt, _)) = "(e1 < e2) = (n * e1 < n * e2)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
382 |
| string_of_rewr (Rewr_Not_Ineq Le) = "~(e1 <= e2) = (e2 < e1)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
383 |
| string_of_rewr (Rewr_Not_Ineq Lt) = "~(e1 < e2) = (e2 <= e1)" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
384 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
385 |
fun flatten_then_conv (Then_Conv (c1, c2)) = flatten_then_conv c1 @ flatten_then_conv c2 |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
386 |
| flatten_then_conv c = [c] |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
387 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
388 |
fun string_of_conv Keep_Conv = "_" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
389 |
| string_of_conv (c as Then_Conv _) = |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
390 |
space_implode " then " (map (enclose "(" ")" o string_of_conv) (flatten_then_conv c)) |
66301
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
391 |
| string_of_conv (Args_Conv (k, cs)) = |
8a6a89d6cf2b
more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents:
64927
diff
changeset
|
392 |
"args " ^ Argo_Expr.string_of_kind k ^ " " ^ brackets string_of_conv cs |
63960
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
393 |
| string_of_conv (Rewr_Conv r) = string_of_rewr r |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
394 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
395 |
fun string_of_rule (Axiom i) = "axiom " ^ string_of_int i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
396 |
| string_of_rule (Taut (t, _)) = "tautology: " ^ string_of_taut t |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
397 |
| string_of_rule (Conjunct (i, n)) = "conjunct " ^ string_of_int i ^ " of " ^ string_of_int n |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
398 |
| string_of_rule (Rewrite c) = "rewrite: " ^ string_of_conv c |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
399 |
| string_of_rule (Hyp (i, _)) = "hypothesis " ^ string_of_int i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
400 |
| string_of_rule (Clause is) = "clause " ^ brackets string_of_int is |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
401 |
| string_of_rule (Lemma is) = "lemma " ^ brackets string_of_int is |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
402 |
| string_of_rule (Unit_Res i) = "unit-resolution " ^ string_of_int i |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
403 |
| string_of_rule (Refl _) = "reflexivity" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
404 |
| string_of_rule Symm = "symmetry" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
405 |
| string_of_rule Trans = "transitivity" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
406 |
| string_of_rule Cong = "congruence" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
407 |
| string_of_rule Subst = "substitution" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
408 |
| string_of_rule Linear_Comb = "linear-combination" |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
409 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
410 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
411 |
(* unsatisfiability *) |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
412 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
413 |
exception UNSAT of proof |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
414 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
415 |
fun unsat p = raise UNSAT p |
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
416 |
|
3daf02070be5
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff
changeset
|
417 |
end |