src/Tools/Argo/argo_cdcl.ML
author wenzelm
Sun, 21 Mar 2021 23:15:55 +0100
changeset 73461 067c23324784
parent 63960 3daf02070be5
permissions -rw-r--r--
clarified symbol names, notably relevant for Z_Notation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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_cdcl.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
Propositional satisfiability solver in the style of conflict-driven
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     5
clause-learning (CDCL). It features:
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     6
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     7
 * conflict analysis and clause learning based on the first unique implication point
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     8
 * nonchronological backtracking
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     9
 * dynamic variable ordering (VSIDS)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    10
 * restarting
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    11
 * polarity caching
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    12
 * propagation via two watched literals
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    13
 * special propagation of binary clauses 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    14
 * minimizing learned clauses
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    15
 * support for external knowledge
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    16
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    17
These features might be added:
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    18
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    19
 * pruning of unnecessary learned clauses
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    20
 * rebuilding the variable heap
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    21
 * aligning the restart level with the decision heuristics: keep decisions that would
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    22
   be recovered instead of backjumping to level 0
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    23
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    24
The implementation 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
    25
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    26
  Niklas E'en and Niklas S"orensson. An Extensible SAT-solver. In Enrico
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    27
  Giunchiglia and Armando Tacchella, editors, Theory and Applications of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    28
  Satisfiability Testing. Volume 2919 of Lecture Notes in Computer
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    29
  Science, pages 502-518. Springer, 2003.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    30
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    31
  Niklas S"orensson and Armin Biere. Minimizing Learned Clauses. In
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    32
  Oliver Kullmann, editor, Theory and Applications of Satisfiability
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    33
  Testing. Volume 5584 of Lecture Notes in Computer Science,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    34
  pages 237-243. Springer, 2009.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    35
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    36
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    37
signature ARGO_CDCL =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    38
sig
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    39
  (* types *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    40
  type 'a explain = Argo_Lit.literal -> 'a -> Argo_Cls.clause * 'a
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    41
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    42
  (* context *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    43
  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
    44
  val 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
    45
  val assignment_of: context -> Argo_Lit.literal -> bool option
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    46
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    47
  (* enriching the context *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    48
  val add_atom: Argo_Term.term -> 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
    49
  val add_axiom: Argo_Cls.clause -> context -> int * context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    50
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    51
  (* main operations *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    52
  val assume: 'a explain -> Argo_Lit.literal -> context -> 'a ->
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    53
    Argo_Cls.clause option * context * 'a
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 propagate: context -> Argo_Common.literal Argo_Common.implied * context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    55
  val decide: context -> context option
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    56
  val analyze: 'a explain -> Argo_Cls.clause -> context -> 'a -> int * context * 'a
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    57
  val restart: context -> int * context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    58
end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    59
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    60
structure Argo_Cdcl: ARGO_CDCL =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    61
struct
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
(* basic types and operations *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    64
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    65
type 'a explain = Argo_Lit.literal -> 'a -> Argo_Cls.clause * 'a
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    66
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    67
datatype reason =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    68
  Level0 of Argo_Proof.proof |
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    69
  Decided of int * int * (bool * reason) Argo_Termtab.table |
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    70
  Implied of int * int * (Argo_Lit.literal * reason) list * Argo_Proof.proof |
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    71
  External 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
    72
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    73
fun level_of (Level0 _) = 0
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    74
  | level_of (Decided (l, _, _)) = l
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    75
  | level_of (Implied (l, _, _, _)) = l
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    76
  | level_of (External l) = l
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    77
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    78
type justified = Argo_Lit.literal * reason
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    79
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    80
type watches = Argo_Cls.clause list * Argo_Cls.clause list
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    81
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    82
fun get_watches wts t = Argo_Termtab.lookup wts t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    83
fun map_watches f t wts = Argo_Termtab.map_default (t, ([], [])) f wts
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    84
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    85
fun map_lit_watches f (Argo_Lit.Pos t) = map_watches (apsnd 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
    86
  | map_lit_watches f (Argo_Lit.Neg t) = map_watches (apfst 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
    87
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    88
fun watches_of wts (Argo_Lit.Pos t) = (case get_watches wts t of SOME (ws, _) => ws | NONE => [])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    89
  | watches_of wts (Argo_Lit.Neg t) = (case get_watches wts t of SOME (_, ws) => ws | NONE => [])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    90
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    91
fun attach cls lit = map_lit_watches (cons cls) lit
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    92
fun detach cls lit = map_lit_watches (remove Argo_Cls.eq_clause cls) lit
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    93
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
(* literal values *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    96
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    97
fun raw_val_of vals lit = Argo_Termtab.lookup vals (Argo_Lit.term_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
    98
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    99
fun val_of vals (Argo_Lit.Pos t) = Argo_Termtab.lookup vals t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   100
  | val_of vals (Argo_Lit.Neg t) = Option.map (apfst not) (Argo_Termtab.lookup vals t)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   101
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   102
fun value_of vals (Argo_Lit.Pos t) = Option.map fst (Argo_Termtab.lookup vals t)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   103
  | value_of vals (Argo_Lit.Neg t) = Option.map (not o fst) (Argo_Termtab.lookup vals t)
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
fun justified vals lit = Option.map (pair lit o snd) (raw_val_of vals lit)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   106
fun the_reason_of vals lit = snd (the (raw_val_of vals lit))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   107
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   108
fun assign (Argo_Lit.Pos t) r = Argo_Termtab.update (t, (true, r))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   109
  | assign (Argo_Lit.Neg t) r = Argo_Termtab.update (t, (false, r))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   110
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   111
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   112
(* context *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   113
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   114
type trail = int * justified list (* the trail height and the sequence of assigned literals *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   115
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   116
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
   117
  units: Argo_Common.literal list, (* the literals that await propagation *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   118
  level: int, (* the decision level *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   119
  trail: int * justified list, (* the trail height and the sequence of assigned literals *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   120
  vals: (bool * reason) Argo_Termtab.table, (* mapping of terms to polarity and reason *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   121
  wts: watches Argo_Termtab.table, (* clauses watched by terms *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   122
  heap: Argo_Heap.heap, (* max-priority heap for decision heuristics *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   123
  clss: Argo_Cls.table, (* information about clauses *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   124
  prf: Argo_Proof.context} (* the 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
   125
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   126
fun mk_context units level trail vals wts heap clss prf: context =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   127
  {units=units, level=level, trail=trail, vals=vals, wts=wts, heap=heap, clss=clss, prf=prf}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   128
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   129
val context =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   130
  mk_context [] 0 (0, []) Argo_Termtab.empty Argo_Termtab.empty Argo_Heap.heap
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   131
    Argo_Cls.table Argo_Proof.cdcl_context
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
fun drop_levels n (Decided (l, h, vals)) trail heap =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   134
      if l = n + 1 then ((h, trail), vals, heap) else drop_literal n trail heap
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   135
  | drop_levels n _ tr heap = drop_literal n tr heap
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
and drop_literal n ((lit, r) :: trail) heap = drop_levels n r trail (Argo_Heap.insert lit heap)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   138
  | drop_literal _ [] _ = raise Fail "bad trail"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   139
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   140
fun backjump_to new_level (cx as {level, trail=(_, tr), wts, heap, clss, prf, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   141
  if new_level >= level then (0, cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   142
  else
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   143
    let val (trail, vals, heap) = drop_literal (Integer.max 0 new_level) tr heap
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   144
    in (level - new_level, mk_context [] new_level trail vals wts heap clss prf) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   145
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
(* proofs *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   148
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   149
fun tag_clause (lits, p) prf = Argo_Proof.mk_clause lits p prf |>> pair lits
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   150
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   151
fun level0_unit_proof (lit, Level0 p') (p, prf) = Argo_Proof.mk_unit_res lit p p' prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   152
  | level0_unit_proof _ _ = raise Fail "bad reason"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   153
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   154
fun level0_unit_proofs lrs p prf = fold level0_unit_proof lrs (p, prf)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   155
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   156
fun unsat ({vals, prf, ...}: context) (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
   157
  let val lrs = map (fn lit => (lit, the_reason_of vals lit)) lits
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   158
  in Argo_Proof.unsat (fst (level0_unit_proofs lrs p prf)) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   159
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   160
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   161
(* literal operations *)
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
fun push lit p reason prf ({units, level, trail=(h, tr), vals, wts, heap, clss, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   164
  let val vals = assign lit reason vals
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   165
  in mk_context ((lit, p) :: units) level (h + 1, (lit, reason) :: tr) vals wts heap clss prf end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   166
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   167
fun push_level0 lit p lrs (cx as {prf, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   168
  let val (p, prf) = level0_unit_proofs lrs p prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   169
  in push lit (SOME p) (Level0 p) prf 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
   170
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   171
fun push_implied lit p lrs (cx as {level, trail=(h, _), prf, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   172
  if level > 0 then push lit NONE (Implied (level, h, lrs, p)) prf cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   173
  else push_level0 lit p lrs cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   174
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   175
fun push_decided lit (cx as {level, trail=(h, _), vals, prf, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   176
  push lit NONE (Decided (level, h, vals)) prf cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   177
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   178
fun assignment_of ({vals, ...}: context) = value_of vals
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   179
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   180
fun replace_watches old new cls ({units, level, trail, vals, wts, heap, clss, prf}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   181
  mk_context units level trail vals (attach cls new (detach cls old wts)) heap clss prf
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
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   184
(* clause operations *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   185
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   186
fun as_clause cls ({units, level, trail, vals, wts, heap, clss, prf}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   187
  let val (cls, prf) = tag_clause cls prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   188
  in (cls, mk_context units level trail vals wts heap clss prf) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   189
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   190
fun note_watches ([_, _], _) _ clss = clss
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   191
  | note_watches cls lp clss = Argo_Cls.put_watches cls lp clss
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   192
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   193
fun attach_clause lit1 lit2 (cls as (lits, _)) cx =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   194
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   195
    val {units, level, trail, vals, wts, heap, clss, prf}: context = cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   196
    val wts = attach cls lit1 (attach cls lit2 wts)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   197
    val clss = note_watches cls (lit1, lit2) clss
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   198
  in mk_context units level trail vals wts (fold Argo_Heap.count lits heap) clss prf end
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
fun change_watches _ (false, _, _) cx = cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   201
  | change_watches cls (true, l1, l2) ({units, level, trail, vals, wts, heap, clss, prf}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   202
      mk_context units level trail vals wts heap (Argo_Cls.put_watches cls (l1, l2) clss) prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   203
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   204
fun add_asserting lit lit' (cls as (_, p)) lrs cx =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   205
  attach_clause lit lit' cls (push_implied lit p lrs cx)
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
(*
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   208
  When learning a non-unit clause, the context is backtracked to the highest decision level
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   209
  of the assigned literals.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   210
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   211
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   212
fun learn_clause _ ([lit], p) cx = backjump_to 0 cx ||> push_level0 lit p []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   213
  | learn_clause lrs (cls as (lits, _)) cx =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   214
      let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   215
        fun max_level (l, r) (ll as (_, lvl)) = if level_of r > lvl then (l, level_of r) else ll
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   216
        val (lit, lvl) = fold max_level lrs (hd lits, 0)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   217
      in backjump_to lvl cx ||> add_asserting (hd lits) lit cls lrs end
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
  An axiom with one unassigned literal and all remaining literals being assigned to
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   221
  false is asserting. An axiom with all literals assigned to false on level 0 makes the
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   222
  context unsatisfiable. An axiom with all literals assigned to false on higher levels
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   223
  causes backjumping before the highest level, and then the axiom might be asserting if
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   224
  only one literal is unassigned on that level.
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
fun min lit i NONE = SOME (lit, i)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   228
  | min lit i (SOME (lj as (_, j))) = SOME (if i < j then (lit, i) else lj)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   229
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   230
fun level_ord ((_, r1), (_, r2)) = int_ord (level_of r2, level_of r1)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   231
fun add_max lr lrs = Ord_List.insert level_ord lr lrs
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 part [] [] t us fs = (t, us, fs)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   234
  | part (NONE :: vs) (l :: ls) t us fs = part vs ls t (l :: us) fs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   235
  | part (SOME (true, r) :: vs) (l :: ls) t us fs = part vs ls (min l (level_of r) t) us fs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   236
  | part (SOME (false, r) :: vs) (l :: ls) t us fs = part vs ls t us (add_max (l, r) fs)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   237
  | part _ _ _ _ _ = raise Fail "mismatch between values and literals"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   238
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   239
fun backjump_add (lit, r) (lit', r') cls lrs cx =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   240
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   241
    val add =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   242
      if level_of r = level_of r' then attach_clause lit lit' cls
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   243
      else add_asserting lit lit' cls lrs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   244
  in backjump_to (level_of r - 1) cx ||> add end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   245
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   246
fun analyze_axiom vs (cls as (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
   247
  (case part vs lits NONE [] [] of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   248
    (SOME (lit, lvl), [], []) =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   249
      if lvl > 0 then backjump_to 0 cx ||> push_implied lit p [] else (0, cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   250
  | (SOME (lit, lvl), [], (lit', _) :: _) => (0, cx |> (lvl > 0) ? attach_clause lit lit' cls)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   251
  | (SOME (lit, lvl), lit' :: _, _) => (0, cx |> (lvl > 0) ? attach_clause lit lit' cls)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   252
  | (NONE, [], (_, Level0 _) :: _) => unsat cx cls
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   253
  | (NONE, [], [(lit, _)]) => backjump_to 0 cx ||> push_implied lit p []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   254
  | (NONE, [], lrs as (lr :: lr' :: _)) => backjump_add lr lr' cls lrs cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   255
  | (NONE, [lit], []) => backjump_to 0 cx ||> push_implied lit p []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   256
  | (NONE, [lit], lrs as (lit', _) :: _) => (0, add_asserting lit lit' cls lrs cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   257
  | (NONE, lit1 :: lit2 :: _, _) => (0, attach_clause lit1 lit2 cls cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   258
  | _ => raise Fail "bad clause")
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   259
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   260
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   261
(* enriching the context *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   262
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   263
fun add_atom t ({units, level, trail, vals, wts, heap, clss, prf}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   264
  let val heap = Argo_Heap.insert (Argo_Lit.Pos t) heap
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   265
  in mk_context units level trail vals wts heap clss prf end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   266
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   267
fun add_axiom ([], p) _ = Argo_Proof.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
   268
  | add_axiom (cls as (lits, _)) (cx as {vals, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   269
      if has_duplicates Argo_Lit.eq_lit lits then raise Fail "clause with duplicate literals"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   270
      else if has_duplicates Argo_Lit.dual_lit lits then (0, cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   271
      else analyze_axiom (map (val_of vals) lits) (as_clause cls cx)
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
(* external knowledge *)
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 assume explain lit (cx as {level, vals, prf, ...}: context) x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   277
  (case value_of vals lit of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   278
    SOME true => (NONE, cx, x)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   279
  | SOME false => 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   280
      let val (cls, x) = explain lit x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   281
      in if level = 0 then unsat cx cls else (SOME cls, cx, x) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   282
  | NONE =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   283
      if level = 0 then
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   284
        let val ((lits, p), x) = explain lit x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   285
        in (NONE, push_level0 lit p (map_filter (justified vals) lits) cx, x) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   286
      else (NONE, push lit NONE (External level) prf cx, x))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   287
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   288
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   289
(* propagation *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   290
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   291
exception CONFLICT of Argo_Cls.clause * context
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 order_lits_by lit (l1, l2) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   294
  if Argo_Lit.eq_id (l1, lit) then (true, l2, l1) else (false, l1, l2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   295
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   296
fun prop_binary (_, implied_lit, other_lit) (cls as (_, p)) (cx as {level, vals, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   297
  (case value_of vals implied_lit of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   298
    NONE => push_implied implied_lit p [(other_lit, the_reason_of vals other_lit)] cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   299
  | SOME true => cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   300
  | SOME false => if level = 0 then unsat cx cls else raise CONFLICT (cls, cx))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   301
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   302
datatype next = Lit of Argo_Lit.literal | None of justified list
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   303
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   304
fun with_non_false f l (SOME (false, r)) lrs = f ((l, r) :: lrs)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   305
  | with_non_false _ l _ _ = Lit l
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   306
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   307
fun first_non_false _ _ [] lrs = None lrs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   308
  | first_non_false vals lit (l :: ls) lrs =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   309
      if Argo_Lit.eq_lit (l, lit) then first_non_false vals lit ls lrs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   310
      else with_non_false (first_non_false vals lit ls) l (val_of vals l) lrs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   311
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   312
fun prop_nary (lp as (_, lit1, lit2)) (cls as (lits, p)) (cx as {level, vals, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   313
  let val v = value_of vals lit1
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   314
  in
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   315
    if v = SOME true then change_watches cls lp cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   316
    else
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   317
      (case first_non_false vals lit1 lits [] of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   318
        Lit lit2' => change_watches cls (true, lit1, lit2') (replace_watches lit2 lit2' cls cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   319
      | None lrs =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   320
          if v = NONE then push_implied lit1 p lrs (change_watches cls lp cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   321
          else if level = 0 then unsat cx cls
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   322
          else raise CONFLICT (cls, change_watches cls lp cx))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   323
  end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   324
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   325
fun prop_cls lit (cls as ([l1, l2], _)) cx = prop_binary (order_lits_by lit (l1, l2)) cls cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   326
  | prop_cls lit cls (cx as {clss, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   327
      prop_nary (order_lits_by lit (Argo_Cls.get_watches clss cls)) cls cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   328
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 prop_lit (lp as (lit, _)) (lps, cx as {wts, ...}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   330
  (lp :: lps, fold (prop_cls lit) (watches_of wts lit) cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   331
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 prop lps (cx as {units=[], ...}: context) = (Argo_Common.Implied (rev lps), cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   333
  | prop lps ({units, level, trail, vals, wts, heap, clss, prf}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   334
      fold_rev prop_lit units (lps, mk_context [] level trail vals wts heap clss prf) |-> prop
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   335
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   336
fun propagate cx = prop [] cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   337
  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   338
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   339
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   340
(* decisions *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   341
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   342
(*
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   343
  Decisions are based on an activity heuristics. The most active variable that is
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   344
  still unassigned is chosen.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   345
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   346
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   347
fun decide ({units, level, trail, vals, wts, heap, clss, prf}: context) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   348
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   349
    fun check NONE = NONE
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   350
      | check (SOME (lit, heap)) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   351
          if Argo_Termtab.defined vals (Argo_Lit.term_of lit) then check (Argo_Heap.extract heap)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   352
          else SOME (push_decided lit (mk_context units (level + 1) trail vals wts heap clss prf))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   353
  in check (Argo_Heap.extract heap) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   354
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   355
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   356
(* conflict analysis and clause learning *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   357
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   358
(*
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   359
  Learned clauses often contain literals that are redundant, because they are
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   360
  subsumed by other literals of the clause. By analyzing the implication graph beyond
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   361
  the unique implication point, such redundant literals can be identified and hence
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   362
  removed from the learned clause. Only literals occurring in the learned clause and
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   363
  their reasons need to be analyzed.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   364
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   365
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   366
exception ESSENTIAL of unit
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   367
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   368
fun history_ord ((h1, lit1, _), (h2, lit2, _)) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   369
  if h1 < 0 andalso h2 < 0 then int_ord (apply2 Argo_Lit.signed_id_of (lit1, lit2))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   370
  else int_ord (h2, h1)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   371
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   372
fun rec_redundant stop (lit, Implied (lvl, h, lrs, p)) lps =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   373
      if stop lit lvl then lps
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   374
      else fold (rec_redundant stop) lrs ((h, lit, p) :: lps)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   375
  | rec_redundant stop (lit, Decided (lvl, _, _)) lps =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   376
      if stop lit lvl then lps
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   377
      else raise ESSENTIAL ()
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   378
  | rec_redundant _ (lit, Level0 p) lps = ((~1, lit, p) :: lps)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   379
  | rec_redundant _ _ _ = raise ESSENTIAL ()
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   380
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   381
fun redundant stop (lr as (lit, Implied (_, h, lrs, p))) (lps, essential_lrs) = (
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   382
      (fold (rec_redundant stop) lrs ((h, lit, p) :: lps), essential_lrs)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   383
      handle ESSENTIAL () => (lps, lr :: essential_lrs))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   384
  | redundant _ lr (lps, essential_lrs) = (lps, lr :: essential_lrs)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   385
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   386
fun resolve_step (_, l, p') (p, prf) = Argo_Proof.mk_unit_res l p p' prf
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 reduce lrs p prf =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   389
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   390
    val lits = map fst lrs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   391
    val levels = fold (insert (op =) o level_of o snd) lrs []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   392
    fun stop lit level =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   393
      if member Argo_Lit.eq_lit lits lit then true
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   394
      else if member (op =) levels level then false
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   395
      else raise ESSENTIAL ()
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   396
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   397
    val (lps, lrs) = fold (redundant stop) lrs ([], [])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   398
  in (lrs, fold resolve_step (sort_distinct history_ord lps) (p, prf)) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   399
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   400
(*
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   401
  Literals that are candidates for the learned lemma are marked and unmarked while
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   402
  traversing backwards through the trail. The last remaining marked literal is the first
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   403
  unique implication point.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   404
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   405
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   406
fun unmark lit ms = remove Argo_Lit.eq_id lit ms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   407
fun marked ms lit = member Argo_Lit.eq_id ms lit
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   408
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
  Whenever an implication is recorded, the reason for the false literals 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
   411
  asserting clause are known. It is reasonable to store this justification list as part
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   412
  of the implication reason. Consequently, the implementation of conflict analysis can
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   413
  benefit from this information, which does not need to be re-computed here.
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
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   416
fun justification_for _ _ _ (Implied (_, _, lrs, p)) x = (lrs, p, x)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   417
  | justification_for explain vals lit (External _) x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   418
      let val ((lits, p), x) = explain lit x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   419
      in (map_filter (justified vals) lits, p, x) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   420
  | justification_for _ _ _ _ _ = raise Fail "bad reason"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   421
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   422
fun first_lit pred ((lr as (lit, _)) :: lrs) = if pred lit then (lr, lrs) else first_lit pred lrs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   423
  | first_lit _ _ = raise Empty
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   424
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   425
(*
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   426
  Beginning from the conflicting clause, the implication graph is traversed to the first
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   427
  unique implication point. This breadth-first search is controlled by the topological order of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   428
  the trail, which is traversed backwards. While traversing through the trail, the conflict
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   429
  literals of lower levels are collected to form the conflict lemma together with the unique
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   430
  implication point. Conflict literals assigned on level 0 are excluded from the conflict lemma.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   431
  Conflict literals assigned on the current level are candidates for the first unique
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   432
  implication point.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   433
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   434
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   435
fun analyze explain cls (cx as {level, trail, vals, wts, heap, clss, prf, ...}: context) x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   436
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   437
    fun from_clause [] trail ms lrs h p prf x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   438
          from_trail (first_lit (marked ms) trail) ms lrs h p prf x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   439
      | from_clause ((lit, r) :: clause_lrs) trail ms lrs h p prf x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   440
          from_reason r lit clause_lrs trail ms lrs h p prf x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   441
 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   442
    and from_reason (Level0 p') lit clause_lrs trail ms lrs h p prf x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   443
          let val (p, prf) = Argo_Proof.mk_unit_res lit p p' prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   444
          in from_clause clause_lrs trail ms lrs h p prf x end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   445
      | from_reason r lit clause_lrs trail ms lrs h p prf x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   446
          if level_of r = level then
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   447
            if marked ms lit then from_clause clause_lrs trail ms lrs h p prf x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   448
            else from_clause clause_lrs trail (lit :: ms) lrs (Argo_Heap.increase lit h) p prf x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   449
          else
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   450
            let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   451
              val (lrs, h) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   452
                if AList.defined Argo_Lit.eq_id lrs lit then (lrs, h)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   453
                else ((lit, r) :: lrs, Argo_Heap.increase lit h)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   454
            in from_clause clause_lrs trail ms lrs h p prf x end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   455
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   456
    and from_trail ((lit, _), _) [_] lrs h p prf x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   457
          let val (lrs, (p, prf)) = reduce lrs p prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   458
          in (Argo_Lit.negate lit :: map fst lrs, lrs, h, p, prf, x) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   459
      | from_trail ((lit, r), trail) ms lrs h p prf x =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   460
          let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   461
            val (clause_lrs, p', x) = justification_for explain vals lit r x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   462
            val (p, prf) = Argo_Proof.mk_unit_res lit p' p prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   463
          in from_clause clause_lrs trail (unmark lit ms) lrs h p prf x end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   464
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   465
    val (ls, p) = cls
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   466
    val lrs = if level = 0 then unsat cx cls else map (fn l => (l, the_reason_of vals l)) ls
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   467
    val (lits, lrs, heap, p, prf, x) = from_clause lrs (snd trail) [] [] heap p prf x
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   468
    val heap = Argo_Heap.decay heap
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   469
    val (levels, cx) = learn_clause lrs (lits, p) (mk_context [] level trail vals wts heap clss prf)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   470
  in (levels, cx, x) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   471
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   472
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   473
(* restarting *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   474
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   475
fun restart cx = backjump_to 0 cx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   476
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   477
end