new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
authorboehmes
Thu, 29 Sep 2016 20:54:44 +0200
changeset 63960 3daf02070be5
parent 63959 f77dca1abf1b
child 63961 2fd9656c4c82
new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
src/HOL/Argo.thy
src/HOL/ROOT
src/HOL/Real.thy
src/HOL/Tools/Argo/argo_real.ML
src/HOL/Tools/Argo/argo_tactic.ML
src/HOL/ex/Argo_Examples.thy
src/Tools/Argo/argo_cc.ML
src/Tools/Argo/argo_cdcl.ML
src/Tools/Argo/argo_clausify.ML
src/Tools/Argo/argo_cls.ML
src/Tools/Argo/argo_common.ML
src/Tools/Argo/argo_core.ML
src/Tools/Argo/argo_expr.ML
src/Tools/Argo/argo_heap.ML
src/Tools/Argo/argo_lit.ML
src/Tools/Argo/argo_proof.ML
src/Tools/Argo/argo_rewr.ML
src/Tools/Argo/argo_simplex.ML
src/Tools/Argo/argo_solver.ML
src/Tools/Argo/argo_term.ML
src/Tools/Argo/argo_thy.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Argo.thy	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,27 @@
+(*  Title:      HOL/Argo.thy
+    Author:     Sascha Boehme
+*)
+
+theory Argo
+imports HOL
+begin
+
+ML_file "~~/src/Tools/Argo/argo_expr.ML"
+ML_file "~~/src/Tools/Argo/argo_term.ML"
+ML_file "~~/src/Tools/Argo/argo_lit.ML"
+ML_file "~~/src/Tools/Argo/argo_proof.ML"
+ML_file "~~/src/Tools/Argo/argo_rewr.ML"
+ML_file "~~/src/Tools/Argo/argo_cls.ML"
+ML_file "~~/src/Tools/Argo/argo_common.ML"
+ML_file "~~/src/Tools/Argo/argo_cc.ML"
+ML_file "~~/src/Tools/Argo/argo_simplex.ML"
+ML_file "~~/src/Tools/Argo/argo_thy.ML"
+ML_file "~~/src/Tools/Argo/argo_heap.ML"
+ML_file "~~/src/Tools/Argo/argo_cdcl.ML"
+ML_file "~~/src/Tools/Argo/argo_core.ML"
+ML_file "~~/src/Tools/Argo/argo_clausify.ML"
+ML_file "~~/src/Tools/Argo/argo_solver.ML"
+
+ML_file "Tools/Argo/argo_tactic.ML"
+
+end
--- a/src/HOL/ROOT	Thu Sep 29 18:52:34 2016 +0200
+++ b/src/HOL/ROOT	Thu Sep 29 20:54:44 2016 +0200
@@ -626,6 +626,7 @@
     Sudoku
     Code_Timing
     Perm_Fragments
+    Argo_Examples
   theories [skip_proofs = false]
     Meson_Test
   document_files "root.bib" "root.tex"
--- a/src/HOL/Real.thy	Thu Sep 29 18:52:34 2016 +0200
+++ b/src/HOL/Real.thy	Thu Sep 29 20:54:44 2016 +0200
@@ -10,7 +10,7 @@
 section \<open>Development of the Reals using Cauchy Sequences\<close>
 
 theory Real
-imports Rat
+imports Rat Argo
 begin
 
 text \<open>
@@ -1807,4 +1807,9 @@
   for x y :: real
   by auto
 
+
+subsection \<open>Setup for Argo\<close>
+
+ML_file "Tools/Argo/argo_real.ML"
+
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Argo/argo_real.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,297 @@
+(*  Title:      HOL/Tools/argo_real.ML
+    Author:     Sascha Boehme
+
+Extension of the Argo tactic for the reals.
+*)
+
+structure Argo_Real: sig end =
+struct
+
+(* translating input terms *)
+
+fun trans_type _ @{typ Real.real} tcx = SOME (Argo_Expr.Real, tcx)
+  | trans_type _ _ _ = NONE
+
+fun trans_term f (@{const Groups.uminus_class.uminus (real)} $ t) tcx =
+      tcx |> f t |>> Argo_Expr.mk_neg |> SOME
+  | trans_term f (@{const Groups.plus_class.plus (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_add2 |> SOME
+  | trans_term f (@{const Groups.minus_class.minus (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_sub |> SOME
+  | trans_term f (@{const Groups.times_class.times (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_mul |> SOME
+  | trans_term f (@{const Rings.divide_class.divide (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_div |> SOME
+  | trans_term f (@{const Orderings.ord_class.min (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_min |> SOME
+  | trans_term f (@{const Orderings.ord_class.max (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_max |> SOME
+  | trans_term f (@{const Groups.abs_class.abs (real)} $ t) tcx =
+      tcx |> f t |>> Argo_Expr.mk_abs |> SOME
+  | trans_term f (@{const Orderings.ord_class.less (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_lt |> SOME
+  | trans_term f (@{const Orderings.ord_class.less_eq (real)} $ t1 $ t2) tcx =
+      tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_le |> SOME
+  | trans_term _ t tcx =
+      (case try HOLogic.dest_number t of
+        SOME (@{typ Real.real}, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx)
+      | _ => NONE)
+
+
+(* reverse translation *)
+
+fun mk_plus t1 t2 = @{const Groups.plus_class.plus (real)} $ t1 $ t2
+fun mk_sum ts = uncurry (fold_rev mk_plus) (split_last ts)
+fun mk_times t1 t2 = @{const Groups.times_class.times (real)} $ t1 $ t2
+fun mk_divide t1 t2 = @{const Rings.divide_class.divide (real)} $ t1 $ t2
+fun mk_le t1 t2 = @{const Orderings.ord_class.less_eq (real)} $ t1 $ t2
+fun mk_lt t1 t2 = @{const Orderings.ord_class.less (real)} $ t1 $ t2
+
+fun mk_real_num i = HOLogic.mk_number @{typ Real.real} i
+
+fun mk_number n =
+  let val (p, q) = Rat.dest n
+  in if q = 1 then mk_real_num p else mk_divide (mk_real_num p) (mk_real_num q) end
+
+fun term_of _ (Argo_Expr.E (Argo_Expr.Num n, _)) = SOME (mk_number n)
+  | term_of f (Argo_Expr.E (Argo_Expr.Neg, [e])) =
+      SOME (@{const Groups.uminus_class.uminus (real)} $ f e)
+  | term_of f (Argo_Expr.E (Argo_Expr.Add, es)) = SOME (mk_sum (map f es))
+  | term_of f (Argo_Expr.E (Argo_Expr.Sub, [e1, e2])) =
+      SOME (@{const Groups.minus_class.minus (real)} $ f e1 $ f e2)
+  | term_of f (Argo_Expr.E (Argo_Expr.Mul, [e1, e2])) = SOME (mk_times (f e1) (f e2))
+  | term_of f (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) = SOME (mk_divide (f e1) (f e2))
+  | term_of f (Argo_Expr.E (Argo_Expr.Min, [e1, e2])) =
+      SOME (@{const Orderings.ord_class.min (real)} $ f e1 $ f e2)
+  | term_of f (Argo_Expr.E (Argo_Expr.Max, [e1, e2])) =
+      SOME (@{const Orderings.ord_class.max (real)} $ f e1 $ f e2)
+  | term_of f (Argo_Expr.E (Argo_Expr.Abs, [e])) = SOME (@{const Groups.abs_class.abs (real)} $ f e)
+  | term_of f (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (mk_le (f e1) (f e2))
+  | term_of f (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (mk_lt (f e1) (f e2))
+  | term_of _ _ = NONE
+
+
+(* proof replay for rewrite steps *)
+
+fun mk_rewr thm = thm RS @{thm eq_reflection}
+
+fun by_simp ctxt t = 
+  let fun prove {context, ...} = HEADGOAL (Simplifier.simp_tac context)
+  in Goal.prove ctxt [] [] (HOLogic.mk_Trueprop t) prove end
+
+fun prove_num_pred ctxt n =
+  by_simp ctxt (uncurry mk_lt (apply2 mk_number (if @0 < n then (@0, n) else (n, @0))))
+
+fun simp_conv ctxt t = Conv.rewr_conv (mk_rewr (by_simp ctxt t))
+
+fun nums_conv mk f ctxt n m =
+  simp_conv ctxt (HOLogic.mk_eq (mk (mk_number n) (mk_number m), mk_number (f (n, m))))
+
+val add_nums_conv = nums_conv mk_plus (op +)
+val mul_nums_conv = nums_conv mk_times (op *)
+val div_nums_conv = nums_conv mk_divide (op /)
+
+fun cmp_nums_conv ctxt b ct =
+  let val t = if b then @{const HOL.True} else @{const HOL.False}
+  in simp_conv ctxt (HOLogic.mk_eq (Thm.term_of ct, t)) ct end
+
+local
+
+fun is_add2 (@{const Groups.plus_class.plus (real)} $ _ $ _) = true
+  | is_add2 _ = false
+
+fun is_add3 (@{const Groups.plus_class.plus (real)} $ _ $ t) = is_add2 t
+  | is_add3 _ = false
+
+val flatten_thm = mk_rewr @{lemma "(a::real) + b + c = a + (b + c)" by simp}
+  
+fun flatten_conv ct =
+  if is_add2 (Thm.term_of ct) then Argo_Tactic.flatten_conv flatten_conv flatten_thm ct
+  else Conv.all_conv ct
+
+val swap_conv = Conv.rewrs_conv (map mk_rewr @{lemma 
+  "(a::real) + (b + c) = b + (a + c)"
+  "(a::real) + b = b + a"
+  by simp_all})
+
+val assoc_conv = Conv.rewr_conv (mk_rewr @{lemma "(a::real) + (b + c) = (a + b) + c" by simp})
+
+val norm_monom_thm = mk_rewr @{lemma "1 * (a::real) = a" by simp}
+fun norm_monom_conv n = if n = @1 then Conv.rewr_conv norm_monom_thm else Conv.all_conv
+
+val add2_thms = map mk_rewr @{lemma
+  "n * (a::real) + m * a = (n + m) * a"
+  "n * (a::real) + a = (n + 1) * a"
+  "(a::real) + m * a = (1 + m) * a"
+  "(a::real) + a = (1 + 1) * a"
+  by algebra+}
+
+val add3_thms = map mk_rewr @{lemma
+  "n * (a::real) + (m * a + b) = (n + m) * a + b"
+  "n * (a::real) + (a + b) = (n + 1) * a + b"
+  "(a::real) + (m * a + b) = (1 + m) * a + b"
+  "(a::real) + (a + b) = (1 + 1) * a + b"
+  by algebra+}
+
+fun choose_conv cv2 cv3 ct = if is_add3 (Thm.term_of ct) then cv3 ct else cv2 ct
+
+fun join_num_conv ctxt n m =
+  let val conv = add_nums_conv ctxt n m
+  in choose_conv conv (assoc_conv then_conv Conv.arg1_conv conv) end
+
+fun join_monom_conv ctxt n m =
+  let
+    val conv = Conv.arg1_conv (add_nums_conv ctxt n m) then_conv norm_monom_conv (n + m)
+    fun seq_conv thms cv = Conv.rewrs_conv thms then_conv cv
+  in choose_conv (seq_conv add2_thms conv) (seq_conv add3_thms (Conv.arg1_conv conv)) end
+
+fun join_conv NONE = join_num_conv
+  | join_conv (SOME _) = join_monom_conv
+
+fun bubble_down_conv _ _ [] ct = Conv.no_conv ct
+  | bubble_down_conv _ _ [_] ct = Conv.all_conv ct
+  | bubble_down_conv ctxt i ((m1 as (n1, i1)) :: (m2 as (n2, i2)) :: ms) ct =
+      if i1 = i then
+        if i2 = i then
+          (join_conv i ctxt n1 n2 then_conv bubble_down_conv ctxt i ((n1 + n2, i) :: ms)) ct
+        else (swap_conv then_conv Conv.arg_conv (bubble_down_conv ctxt i (m1 :: ms))) ct
+      else Conv.arg_conv (bubble_down_conv ctxt i (m2 :: ms)) ct
+
+fun drop_var i ms = filter_out (fn (_, i') => i' = i) ms
+
+fun permute_conv _ [] [] ct = Conv.all_conv ct
+  | permute_conv ctxt (ms as ((_, i) :: _)) [] ct =
+      (bubble_down_conv ctxt i ms then_conv permute_conv ctxt (drop_var i ms) []) ct
+  | permute_conv ctxt ms1 ms2 ct =
+      let val (ms2', (_, i)) = split_last ms2
+      in (bubble_down_conv ctxt i ms1 then_conv permute_conv ctxt (drop_var i ms1) ms2') ct end
+
+val no_monom_conv = Conv.rewr_conv (mk_rewr @{lemma "0 * (a::real) = 0" by simp})
+
+val norm_sum_conv = Conv.rewrs_conv (map mk_rewr @{lemma
+  "0 * (a::real) + b = b"
+  "(a::real) + 0 * b = a"
+  "0 + (a::real) = a"
+  "(a::real) + 0 = a"
+  by simp_all})
+
+fun drop0_conv ct =
+  if is_add2 (Thm.term_of ct) then
+    ((norm_sum_conv then_conv drop0_conv) else_conv Conv.arg_conv drop0_conv) ct
+  else Conv.try_conv no_monom_conv ct
+
+fun full_add_conv ctxt ms1 ms2 =
+  if eq_list (op =) (ms1, ms2) then flatten_conv
+  else flatten_conv then_conv permute_conv ctxt ms1 ms2 then_conv drop0_conv
+
+in
+
+fun add_conv ctxt (ms1, ms2 as [(n, NONE)]) =
+      if n = @0 then full_add_conv ctxt ms1 [] else full_add_conv ctxt ms1 ms2
+  | add_conv ctxt (ms1, ms2) = full_add_conv ctxt ms1 ms2
+
+end
+
+val mul_sum_thm = mk_rewr @{lemma "(x::real) * (y + z) = x * y + x * z" by (rule ring_distribs)}
+fun mul_sum_conv ct =
+  Conv.try_conv (Conv.rewr_conv mul_sum_thm then_conv Conv.binop_conv mul_sum_conv) ct
+
+fun var_of_add_cmp (_ $ _ $ (_ $ _ $ (_ $ _ $ Var v))) = v
+  | var_of_add_cmp t = raise TERM ("var_of_add_cmp", [t])
+
+fun add_cmp_conv ctxt n thm =
+  let val v = var_of_add_cmp (Thm.prop_of thm)
+  in Conv.rewr_conv (Thm.instantiate ([], [(v, Thm.cterm_of ctxt (mk_number n))]) thm) end
+
+fun mul_cmp_conv ctxt n pos_thm neg_thm =
+  let val thm = if @0 < n then pos_thm else neg_thm
+  in Conv.rewr_conv (prove_num_pred ctxt n RS thm) end
+
+val neg_thm = mk_rewr @{lemma "-(x::real) = -1 * x" by simp}
+val sub_thm = mk_rewr @{lemma "(x::real) - y = x + -1 * y" by simp}
+val mul_zero_thm = mk_rewr @{lemma "0 * (x::real) = 0" by (rule mult_zero_left)}
+val mul_one_thm = mk_rewr @{lemma "1 * (x::real) = x" by (rule mult_1)}
+val mul_comm_thm = mk_rewr @{lemma "(x::real) * y = y * x" by simp}
+val mul_assoc_thm = mk_rewr @{lemma "(x::real) * (y * z) = (x * y) * z" by simp}
+val div_zero_thm = mk_rewr @{lemma "0 / (x::real) = 0" by simp}
+val div_one_thm = mk_rewr @{lemma "(x::real) / 1 = x" by simp}
+val div_mul_thm = mk_rewr @{lemma "(x::real) / y = x * (1 / y)" by simp}
+val div_inv_thm = mk_rewr @{lemma "(x::real) / y = (1 / y) * x" by simp}
+val div_left_thm = mk_rewr @{lemma "((x::real) * y) / z = x * (y / z)" by simp}
+val div_right_thm = mk_rewr @{lemma "(x::real) / (y * z) = (1 / y) * (x / z)" by simp}
+val min_thm = mk_rewr @{lemma "min (x::real) y = (if x <= y then x else y)" by (rule min_def)}
+val max_thm = mk_rewr @{lemma "max (x::real) y = (if x <= y then y else x)" by (rule max_def)}
+val abs_thm = mk_rewr @{lemma "abs (x::real) = (if 0 <= x then x else -x)" by simp}
+val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x <= y) & (y <= x))" by (rule order_eq_iff)}
+val add_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x + n <= y + n)" by simp}
+val add_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x + n < y + n)" by simp}
+val sub_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x - y <= 0)" by simp}
+val sub_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x - y < 0)" by simp}
+val pos_mul_le_thm = mk_rewr @{lemma "0 < n ==> ((x::real) <= y) = (n * x <= n * y)" by simp}
+val neg_mul_le_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) <= y) = (n * y <= n * x)" by simp}
+val pos_mul_lt_thm = mk_rewr @{lemma "0 < n ==> ((x::real) < y) = (n * x < n * y)" by simp}
+val neg_mul_lt_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) < y) = (n * y < n * x)" by simp}
+val not_le_thm = mk_rewr @{lemma "(~((x::real) <= y)) = (y < x)" by auto}
+val not_lt_thm = mk_rewr @{lemma "(~((x::real) < y)) = (y <= x)" by auto}
+
+fun replay_rewr _ Argo_Proof.Rewr_Neg = Conv.rewr_conv neg_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Add ps) = add_conv ctxt ps
+  | replay_rewr _ Argo_Proof.Rewr_Sub = Conv.rewr_conv sub_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Zero = Conv.rewr_conv mul_zero_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_One = Conv.rewr_conv mul_one_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Mul_Nums (n, m)) = mul_nums_conv ctxt n m
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Comm = Conv.rewr_conv mul_comm_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Assoc = Conv.rewr_conv mul_assoc_thm
+  | replay_rewr _ Argo_Proof.Rewr_Mul_Sum = mul_sum_conv
+  | replay_rewr ctxt (Argo_Proof.Rewr_Div_Nums (n, m)) = div_nums_conv ctxt n m
+  | replay_rewr _ Argo_Proof.Rewr_Div_Zero = Conv.rewr_conv div_zero_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_One = Conv.rewr_conv div_one_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Mul = Conv.rewr_conv div_mul_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Inv = Conv.rewr_conv div_inv_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Left = Conv.rewr_conv div_left_thm
+  | replay_rewr _ Argo_Proof.Rewr_Div_Right = Conv.rewr_conv div_right_thm
+  | replay_rewr _ Argo_Proof.Rewr_Min = Conv.rewr_conv min_thm
+  | replay_rewr _ Argo_Proof.Rewr_Max = Conv.rewr_conv max_thm
+  | replay_rewr _ Argo_Proof.Rewr_Abs = Conv.rewr_conv abs_thm
+  | replay_rewr _ Argo_Proof.Rewr_Eq_Le = Conv.rewr_conv eq_le_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Nums (_, b)) = cmp_nums_conv ctxt b
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Le, n)) =
+      add_cmp_conv ctxt n add_le_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Lt, n)) =
+      add_cmp_conv ctxt n add_lt_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Le) = Conv.rewr_conv sub_le_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Lt) = Conv.rewr_conv sub_lt_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Le, n)) =
+      mul_cmp_conv ctxt n pos_mul_le_thm neg_mul_le_thm
+  | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Lt, n)) =
+      mul_cmp_conv ctxt n pos_mul_lt_thm neg_mul_lt_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) = Conv.rewr_conv not_le_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) = Conv.rewr_conv not_lt_thm
+  | replay_rewr _ _ = Conv.no_conv
+
+
+(* proof replay *)
+
+val combine_thms = @{lemma
+  "(a::real) < b ==> c < d ==> a + c < b + d"
+  "(a::real) < b ==> c <= d ==> a + c < b + d"
+  "(a::real) <= b ==> c < d ==> a + c < b + d"
+  "(a::real) <= b ==> c <= d ==> a + c <= b + d"
+  by auto}
+
+fun combine thm1 thm2 = hd (Argo_Tactic.discharges thm2 (Argo_Tactic.discharges thm1 combine_thms))
+
+fun replay _ _ Argo_Proof.Linear_Comb prems = SOME (uncurry (fold_rev combine) (split_last prems))
+  | replay _ _ _ _ = NONE
+
+
+(* real extension of the Argo solver *)
+
+val _ = Theory.setup (Argo_Tactic.add_extension {
+  trans_type = trans_type,
+  trans_term = trans_term,
+  term_of = term_of,
+  replay_rewr = replay_rewr,
+  replay = replay})
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/Argo/argo_tactic.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,730 @@
+(*  Title:      HOL/Tools/argo_tactic.ML
+    Author:     Sascha Boehme
+
+HOL method and tactic for the Argo solver.
+*)
+
+signature ARGO_TACTIC =
+sig
+  val trace: string Config.T
+  val timeout: real Config.T
+
+  (* extending the tactic *)
+  type trans_context =
+    Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table
+  type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context
+  type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option
+  type extension = {
+    trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans',
+    trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans',
+    term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option,
+    replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv,
+    replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option}
+  val add_extension: extension -> theory -> theory
+
+  (* proof utilities *)
+  val discharges: thm -> thm list -> thm list
+  val flatten_conv: conv -> thm -> conv
+
+  (* interface to the tactic as well as the underlying checker and prover *)
+  datatype result = Satisfiable of term -> bool option | Unsatisfiable
+  val check: term list -> Proof.context -> result * Proof.context
+  val prove: thm list -> Proof.context -> thm option * Proof.context
+  val argo_tac: Proof.context -> thm list -> int -> tactic
+end
+
+structure Argo_Tactic: ARGO_TACTIC =
+struct
+
+(* readable fresh names for terms *)
+
+fun fresh_name n = Name.variant (case Long_Name.base_name n of "" => "x" | n' => n')
+
+fun fresh_type_name (Type (n, _)) = fresh_name n
+  | fresh_type_name (TFree (n, _)) = fresh_name n
+  | fresh_type_name (TVar ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i)
+
+fun fresh_term_name (Const (n, _)) = fresh_name n
+  | fresh_term_name (Free (n, _)) = fresh_name n
+  | fresh_term_name (Var ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i)
+  | fresh_term_name _ = fresh_name ""
+
+
+(* tracing *)
+
+datatype mode = None | Basic | Full
+
+fun string_of_mode None = "none"
+  | string_of_mode Basic = "basic"
+  | string_of_mode Full = "full"
+
+fun requires_mode None = []
+  | requires_mode Basic = [Basic, Full]
+  | requires_mode Full = [Full]
+
+val trace = Attrib.setup_config_string @{binding argo_trace} (K (string_of_mode None))
+
+fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode
+
+fun output mode ctxt msg = if allows_mode ctxt mode then Output.tracing ("Argo: " ^ msg) else ()
+val tracing = output Basic
+val full_tracing = output Full
+
+fun with_mode mode ctxt f = if allows_mode ctxt mode then f ctxt else ()
+val with_tracing = with_mode Basic
+val with_full_tracing = with_mode Full
+
+
+(* timeout *)
+
+val timeout = Attrib.setup_config_real @{binding argo_timeout} (K 10.0)
+
+fun time_of_timeout ctxt = Time.fromReal (Config.get ctxt timeout)
+
+fun with_timeout ctxt f x = Timeout.apply (time_of_timeout ctxt) f x
+
+
+(* extending the tactic *)
+
+type trans_context =
+  Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table
+
+type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context
+type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option
+
+type extension = {
+  trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans',
+  trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans',
+  term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option,
+  replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv,
+  replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option}
+
+fun eq_extension ((serial1, _), (serial2, _)) = (serial1 = serial2)
+
+structure Extensions = Theory_Data
+(
+  type T = (serial * extension) list
+  val empty = []
+  val extend = I
+  val merge = merge eq_extension 
+)
+
+fun add_extension ext = Extensions.map (insert eq_extension (serial (), ext))
+fun get_extensions ctxt = Extensions.get (Proof_Context.theory_of ctxt)
+fun apply_first ctxt f = get_first (fn (_, e) => f e) (get_extensions ctxt)
+
+fun ext_trans sel ctxt f x tcx = apply_first ctxt (fn ext => sel ext f x tcx)
+
+val ext_trans_type = ext_trans (fn {trans_type, ...}: extension => trans_type)
+val ext_trans_term = ext_trans (fn {trans_term, ...}: extension => trans_term)
+
+fun ext_term_of ctxt f e = apply_first ctxt (fn {term_of, ...}: extension => term_of f e)
+
+fun ext_replay_rewr ctxt r = 
+  get_extensions ctxt
+  |> map (fn (_, {replay_rewr, ...}: extension) => replay_rewr ctxt r)
+  |> Conv.first_conv
+
+fun ext_replay cprop_of ctxt rule prems =
+  (case apply_first ctxt (fn {replay, ...}: extension => replay cprop_of ctxt rule prems) of
+    SOME thm => thm
+  | NONE => raise THM ("failed to replay " ^ quote (Argo_Proof.string_of_rule rule), 0, []))
+
+
+(* translating input terms *)
+
+fun add_new_type T (names, types, terms) =
+  let
+    val (n, names') = fresh_type_name T names
+    val ty = Argo_Expr.Type n
+  in (ty, (names', Typtab.update (T, ty) types, terms)) end
+
+fun add_type T (tcx as (_, types, _)) =
+  (case Typtab.lookup types T of
+    SOME ty => (ty, tcx)
+  | NONE => add_new_type T tcx)
+
+fun trans_type _ @{typ HOL.bool} = pair Argo_Expr.Bool
+  | trans_type ctxt (Type (@{type_name "fun"}, [T1, T2])) =
+      trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func
+  | trans_type ctxt T = (fn tcx =>
+      (case ext_trans_type ctxt (trans_type ctxt) T tcx of
+        SOME result => result
+      | NONE => add_type T tcx))
+
+fun add_new_term ctxt t T tcx =
+  let
+    val (ty, (names, types, terms)) = trans_type ctxt T tcx
+    val (n, names') = fresh_term_name t names
+    val c = (n, ty)
+  in (Argo_Expr.mk_con c, (names', types, Termtab.update (t, c) terms)) end
+
+fun add_term ctxt t (tcx as (_, _, terms)) =
+  (case Termtab.lookup terms t of
+    SOME c => (Argo_Expr.mk_con c, tcx)
+  | NONE => add_new_term ctxt t (Term.fastype_of t) tcx)
+
+fun mk_eq @{typ HOL.bool} = Argo_Expr.mk_iff
+  | mk_eq _ = Argo_Expr.mk_eq
+
+fun trans_term _ @{const HOL.True} = pair Argo_Expr.true_expr
+  | trans_term _ @{const HOL.False} = pair Argo_Expr.false_expr
+  | trans_term ctxt (@{const HOL.Not} $ t) = trans_term ctxt t #>> Argo_Expr.mk_not
+  | trans_term ctxt (@{const HOL.conj} $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2
+  | trans_term ctxt (@{const HOL.disj} $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2
+  | trans_term ctxt (@{const HOL.implies} $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp
+  | trans_term ctxt (Const (@{const_name HOL.If}, _) $ t1 $ t2 $ t3) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>>
+      (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3)
+  | trans_term ctxt (Const (@{const_name HOL.eq}, T) $ t1 $ t2) =
+      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq (Term.domain_type T))
+  | trans_term ctxt (t as (t1 $ t2)) = (fn tcx =>
+      (case ext_trans_term ctxt (trans_term ctxt) t tcx of
+        SOME result => result
+      | NONE => tcx |> trans_term ctxt t1 ||>> trans_term ctxt t2 |>> uncurry Argo_Expr.mk_app))
+  | trans_term ctxt t = (fn tcx =>
+      (case ext_trans_term ctxt (trans_term ctxt) t tcx of
+        SOME result => result
+      | NONE => add_term ctxt t tcx))
+
+fun translate ctxt prop = trans_term ctxt (HOLogic.dest_Trueprop prop)
+
+
+(* invoking the solver *)
+
+type data = {
+  names: Name.context,
+  types: Argo_Expr.typ Typtab.table,
+  terms: (string * Argo_Expr.typ) Termtab.table,
+  argo: Argo_Solver.context}
+
+fun mk_data names types terms argo: data = {names=names, types=types, terms=terms, argo=argo}
+val empty_data = mk_data Name.context Typtab.empty Termtab.empty Argo_Solver.context
+
+structure Solver_Data = Proof_Data
+(
+  type T = data option
+  fun init _ = SOME empty_data
+)
+
+datatype ('m, 'p) solver_result = Model of 'm | Proof of 'p
+
+fun raw_solve es argo = Model (Argo_Solver.assert es argo)
+  handle Argo_Proof.UNSAT proof => Proof proof
+
+fun value_of terms model t =
+  (case Termtab.lookup terms t of
+    SOME c => model c
+  | _ => NONE)
+
+fun trace_props props ctxt =
+  tracing ctxt (Pretty.string_of (Pretty.big_list "using these propositions:"
+    (map (Syntax.pretty_term ctxt) props)))
+
+fun trace_result ctxt ({elapsed, ...}: Timing.timing) msg =
+  tracing ctxt ("found a " ^ msg ^ " in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms")
+
+fun solve props ctxt =
+  (case Solver_Data.get ctxt of
+    NONE => error "bad Argo solver context"
+  | SOME {names, types, terms, argo} =>
+      let
+        val _ = with_tracing ctxt (trace_props props)
+        val (es, (names', types', terms')) = fold_map (translate ctxt) props (names, types, terms)
+        val _ = tracing ctxt "starting the prover"
+      in
+        (case Timing.timing (raw_solve es) argo of
+          (time, Proof proof) =>
+            let val _ = trace_result ctxt time "proof"
+            in (Proof (terms', proof), Solver_Data.put NONE ctxt) end
+        | (time, Model argo') =>
+            let
+              val _ = trace_result ctxt time "model"
+              val model = value_of terms' (Argo_Solver.model_of argo')
+            in (Model model, Solver_Data.put (SOME (mk_data names' types' terms' argo')) ctxt) end)
+      end)
+
+
+(* reverse translation *)
+
+structure Contab = Table(type key = string * Argo_Expr.typ val ord = Argo_Expr.con_ord)
+
+fun mk_nary f ts = uncurry (fold_rev (curry f)) (split_last ts)
+
+fun mk_nary' d _ [] = d
+  | mk_nary' _ f ts = mk_nary f ts
+
+fun mk_ite t1 t2 t3 =
+  let
+    val T = Term.fastype_of t2
+    val ite = Const (@{const_name HOL.If}, [@{typ HOL.bool}, T, T] ---> T)
+  in ite $ t1 $ t2 $ t3 end
+
+fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = @{const HOL.True}
+  | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = @{const HOL.False}
+  | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e)
+  | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) =
+      mk_nary' @{const HOL.True} HOLogic.mk_conj (map (term_of cx) es)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) =
+      mk_nary' @{const HOL.False} HOLogic.mk_disj (map (term_of cx) es)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) =
+      HOLogic.mk_imp (term_of cx e1, term_of cx e2)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) =
+      HOLogic.mk_eq (term_of cx e1, term_of cx e2)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) =
+      mk_ite (term_of cx e1) (term_of cx e2) (term_of cx e3)
+  | term_of cx (Argo_Expr.E (Argo_Expr.Eq, [e1, e2])) =
+      HOLogic.mk_eq (term_of cx e1, term_of cx e2)
+  | term_of cx (Argo_Expr.E (Argo_Expr.App, [e1, e2])) =
+      term_of cx e1 $ term_of cx e2
+  | term_of (_, cons) (Argo_Expr.E (Argo_Expr.Con (c as (n, _)), _)) =
+      (case Contab.lookup cons c of
+        SOME t => t
+      | NONE => error ("Unexpected expression named " ^ quote n))
+  | term_of (cx as (ctxt, _)) e =
+      (case ext_term_of ctxt (term_of cx) e of
+        SOME t => t
+      | NONE => raise Fail "bad expression")
+
+fun as_prop ct = Thm.apply @{cterm HOL.Trueprop} ct
+
+fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e)
+fun cprop_of ctxt cons e = as_prop (cterm_of ctxt cons e)
+
+
+(* generic proof tools *)
+
+fun discharge thm rule = thm INCR_COMP rule
+fun discharge2 thm1 thm2 rule = discharge thm2 (discharge thm1 rule)
+fun discharges thm rules = [thm] RL rules
+
+fun under_assumption f ct =
+  let val cprop = as_prop ct
+  in Thm.implies_intr cprop (f (Thm.assume cprop)) end
+
+fun instantiate cv ct = Thm.instantiate ([], [(Term.dest_Var (Thm.term_of cv), ct)])
+
+
+(* proof replay for tautologies *)
+
+fun prove_taut ctxt ns t = Goal.prove ctxt ns [] (HOLogic.mk_Trueprop t)
+  (fn {context, ...} => HEADGOAL (Classical.fast_tac context))
+
+fun with_frees ctxt n mk =
+  let
+    val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1))
+    val ts = map (Free o rpair @{typ bool}) ns
+    val t = mk_nary HOLogic.mk_disj (mk ts)
+  in prove_taut ctxt ns t end
+
+fun taut_and1_term ts = mk_nary HOLogic.mk_conj ts :: map HOLogic.mk_not ts
+fun taut_and2_term i ts = [HOLogic.mk_not (mk_nary HOLogic.mk_conj ts), nth ts i]
+fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)]
+fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts
+
+val iff_1_taut = @{lemma "P = Q | P | Q" by fast}
+val iff_2_taut = @{lemma "P = Q | (~P) | (~Q)" by fast}
+val iff_3_taut = @{lemma "~(P = Q) | (~P) | Q" by fast}
+val iff_4_taut = @{lemma "~(P = Q) | P | (~Q)" by fast}
+val ite_then_taut = @{lemma "~P | (if P then t else u) = t" by auto}
+val ite_else_taut = @{lemma "P | (if P then t else u) = u" by auto}
+
+fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term
+  | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i)
+  | taut_rule_of ctxt (Argo_Proof.Taut_Or_1 (i, n)) = with_frees ctxt n (taut_or1_term i)
+  | taut_rule_of ctxt (Argo_Proof.Taut_Or_2 n) = with_frees ctxt n taut_or2_term
+  | taut_rule_of _ Argo_Proof.Taut_Iff_1 = iff_1_taut
+  | taut_rule_of _ Argo_Proof.Taut_Iff_2 = iff_2_taut
+  | taut_rule_of _ Argo_Proof.Taut_Iff_3 = iff_3_taut
+  | taut_rule_of _ Argo_Proof.Taut_Iff_4 = iff_4_taut
+  | taut_rule_of _ Argo_Proof.Taut_Ite_Then = ite_then_taut
+  | taut_rule_of _ Argo_Proof.Taut_Ite_Else = ite_else_taut
+
+fun replay_taut ctxt k ct =
+  let val rule = taut_rule_of ctxt k
+  in Thm.instantiate (Thm.match (Thm.cprop_of rule, ct)) rule end
+
+
+(* proof replay for conjunct extraction *)
+
+fun replay_conjunct 0 1 thm = thm
+  | replay_conjunct 0 _ thm = discharge thm @{thm conjunct1}
+  | replay_conjunct 1 2 thm = discharge thm @{thm conjunct2}
+  | replay_conjunct i n thm = replay_conjunct (i - 1) (n - 1) (discharge thm @{thm conjunct2})
+
+
+(* proof replay for rewrite steps *)
+
+fun mk_rewr thm = thm RS @{thm eq_reflection}
+
+fun not_nary_conv rule i ct =
+  if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct
+  else Conv.all_conv ct
+
+val flatten_and_thm = @{lemma "(P1 & P2) & P3 == P1 & P2 & P3" by simp}
+val flatten_or_thm = @{lemma "(P1 | P2) | P3 == P1 | P2 | P3" by simp}
+
+fun flatten_conv cv rule ct = (
+  Conv.try_conv (Conv.arg_conv cv) then_conv
+  Conv.try_conv (Conv.rewr_conv rule then_conv cv)) ct
+
+fun flat_conj_conv ct =
+  (case Thm.term_of ct of
+    @{const HOL.conj} $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct
+  | _ => Conv.all_conv ct)
+
+fun flat_disj_conv ct =
+  (case Thm.term_of ct of
+    @{const HOL.disj} $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct
+  | _ => Conv.all_conv ct)
+
+fun explode rule1 rule2 thm =
+  explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm]
+val explode_conj = explode @{thm conjunct1} @{thm conjunct2}
+val explode_ndis = explode @{lemma "~(P | Q) ==> ~P" by auto} @{lemma "~(P | Q) ==> ~Q" by auto}
+
+fun pick_false i thms = nth thms i
+
+fun pick_dual rule (i1, i2) thms =
+  rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1]
+val pick_dual_conj = pick_dual @{lemma "~P ==> P ==> False" by auto}
+val pick_dual_ndis = pick_dual @{lemma "~P ==> P ==> ~True" by auto}
+
+fun join thm0 rule is thms =
+  let
+    val l = length thms
+    val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is []
+  in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end
+
+val join_conj = join @{lemma "True" by auto} @{lemma "P ==> Q ==> P & Q" by auto}
+val join_ndis = join @{lemma "~False" by auto} @{lemma "~P ==> ~Q ==> ~(P | Q)" by auto}
+
+val false_thm = @{lemma "False ==> P" by auto}
+val ntrue_thm = @{lemma "~True ==> P" by auto}
+val iff_conj_thm = @{lemma "(P ==> Q) ==> (Q ==> P) ==> P = Q" by auto}
+val iff_ndis_thm = @{lemma "(~P ==> ~Q) ==> (~Q ==> ~P) ==> P = Q" by auto}
+
+fun iff_intro rule lf rf ct =
+  let
+    val lhs = under_assumption lf ct
+    val rhs = rf (Thm.dest_arg (snd (Thm.dest_implies (Thm.cprop_of lhs))))
+  in mk_rewr (discharge2 lhs rhs rule) end
+
+fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct
+fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply @{cterm HOL.Not} ct)
+
+fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1))
+fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g))
+val sort_conj = sort_nary with_conj join_conj explode_conj
+val sort_ndis = sort_nary with_ndis join_ndis explode_ndis 
+
+val not_true_thm = mk_rewr @{lemma "(~True) = False" by auto}
+val not_false_thm = mk_rewr @{lemma "(~False) = True" by auto}
+val not_not_thm = mk_rewr @{lemma "(~~P) = P" by auto}
+val not_and_thm = mk_rewr @{lemma "(~(P & Q)) = (~P | ~Q)" by auto}
+val not_or_thm = mk_rewr @{lemma "(~(P | Q)) = (~P & ~Q)" by auto}
+val not_iff_thms = map mk_rewr
+  @{lemma "(~((~P) = Q)) = (P = Q)" "(~(P = (~Q))) = (P = Q)" "(~(P = Q)) = ((~P) = Q)" by auto}
+val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto}
+val iff_false_thms = map mk_rewr @{lemma "(False = P) = (~P)" "(P = False) = (~P)" by auto}
+val iff_not_not_thm = mk_rewr @{lemma "((~P) = (~Q)) = (P = Q)" by auto}
+val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto}
+val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto}
+val iff_dual_thms = map mk_rewr @{lemma "(P = (~P)) = False" "((~P) = P) = False" by auto}
+val imp_thm = mk_rewr @{lemma "(P --> Q) = (~P | Q)" by auto}
+val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((~P | Q) & (P | R) & (Q | R))" by auto}
+val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto}
+val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto}
+val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto}
+val eq_refl_thm = mk_rewr @{lemma "(t = t) = True" by auto}
+val eq_symm_thm = mk_rewr @{lemma "(t1 = t2) = (t2 = t1)" by auto}
+
+fun replay_rewr _ Argo_Proof.Rewr_Not_True = Conv.rewr_conv not_true_thm
+  | replay_rewr _ Argo_Proof.Rewr_Not_False = Conv.rewr_conv not_false_thm
+  | replay_rewr _ Argo_Proof.Rewr_Not_Not = Conv.rewr_conv not_not_thm
+  | replay_rewr _ (Argo_Proof.Rewr_Not_And i) = not_nary_conv not_and_thm i
+  | replay_rewr _ (Argo_Proof.Rewr_Not_Or i) = not_nary_conv not_or_thm i
+  | replay_rewr _ Argo_Proof.Rewr_Not_Iff = Conv.rewrs_conv not_iff_thms
+  | replay_rewr _ (Argo_Proof.Rewr_And_False i) = with_conj (pick_false i) (K false_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_And_Dual ip) = with_conj (pick_dual_conj ip) (K false_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_And_Sort is) = flat_conj_conv then_conv sort_conj is
+  | replay_rewr _ (Argo_Proof.Rewr_Or_True i) = with_ndis (pick_false i) (K ntrue_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_Or_Dual ip) = with_ndis (pick_dual_ndis ip) (K ntrue_thm)
+  | replay_rewr _ (Argo_Proof.Rewr_Or_Sort is) = flat_disj_conv then_conv sort_ndis is
+  | replay_rewr _ Argo_Proof.Rewr_Iff_True = Conv.rewrs_conv iff_true_thms
+  | replay_rewr _ Argo_Proof.Rewr_Iff_False = Conv.rewrs_conv iff_false_thms
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Not_Not = Conv.rewr_conv iff_not_not_thm
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Refl = Conv.rewr_conv iff_refl_thm
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Symm = Conv.rewr_conv iff_symm_thm
+  | replay_rewr _ Argo_Proof.Rewr_Iff_Dual  = Conv.rewrs_conv iff_dual_thms
+  | replay_rewr _ Argo_Proof.Rewr_Imp = Conv.rewr_conv imp_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_Prop = Conv.rewr_conv ite_prop_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_True = Conv.rewr_conv ite_true_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_False = Conv.rewr_conv ite_false_thm
+  | replay_rewr _ Argo_Proof.Rewr_Ite_Eq = Conv.rewr_conv ite_eq_thm
+  | replay_rewr _ Argo_Proof.Rewr_Eq_Refl = Conv.rewr_conv eq_refl_thm
+  | replay_rewr _ Argo_Proof.Rewr_Eq_Symm = Conv.rewr_conv eq_symm_thm
+  | replay_rewr ctxt r = ext_replay_rewr ctxt r
+
+fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
+
+fun replay_conv _ Argo_Proof.Keep_Conv ct = Conv.all_conv ct
+  | replay_conv ctxt (Argo_Proof.Then_Conv (c1, c2)) ct = 
+      (replay_conv ctxt c1 then_conv replay_conv ctxt c2) ct
+  | replay_conv ctxt (Argo_Proof.Args_Conv cs) ct = replay_args_conv ctxt cs ct
+  | replay_conv ctxt (Argo_Proof.Rewr_Conv r) ct = replay_rewr ctxt r ct
+
+and replay_args_conv _ [] ct = Conv.all_conv ct
+  | replay_args_conv ctxt [c] ct = Conv.arg_conv (replay_conv ctxt c) ct
+  | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct
+  | replay_args_conv ctxt (c :: cs) ct =
+      (case Term.head_of (Thm.term_of ct) of
+        Const (@{const_name HOL.If}, _) =>
+          let val (cs', c') = split_last cs
+          in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end
+      | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct)
+
+fun replay_rewrite ctxt c thm = Conv.fconv_rule (HOLogic.Trueprop_conv (replay_conv ctxt c)) thm
+
+
+(* proof replay for clauses *)
+
+val prep_clause_rule = @{lemma "P ==> ~P ==> False" by fast}
+val extract_lit_rule = @{lemma "(~(P | Q) ==> False) ==> ~P ==> ~Q ==> False" by fast}
+
+fun add_lit i thm lits =
+  let val ct = Thm.cprem_of thm 1
+  in (Thm.implies_elim thm (Thm.assume ct), (i, ct) :: lits) end
+
+fun extract_lits [] _ = error "Bad clause"
+  | extract_lits [i] (thm, lits) = add_lit i thm lits
+  | extract_lits (i :: is) (thm, lits) =
+      extract_lits is (add_lit i (discharge thm extract_lit_rule) lits)
+
+fun lit_ord ((l1, _), (l2, _)) = int_ord (abs l1, abs l2)
+
+fun replay_with_lits [] thm lits = (thm, lits)
+  | replay_with_lits is thm lits =
+      extract_lits is (discharge thm prep_clause_rule, lits)
+      ||> Ord_List.make lit_ord
+
+fun replay_clause is thm = replay_with_lits is thm []
+
+
+(* proof replay for unit resolution *)
+
+val unit_rule = @{lemma "(P ==> False) ==> (~P ==> False) ==> False" by fast}
+val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1))
+val bogus_ct = @{cterm HOL.True}
+
+fun replay_unit_res lit (pthm, plits) (nthm, nlits) =
+  let
+    val plit = the (AList.lookup (op =) plits lit)
+    val nlit = the (AList.lookup (op =) nlits (~lit))
+    val prune = Ord_List.remove lit_ord (lit, bogus_ct)
+  in
+    unit_rule
+    |> instantiate unit_rule_var (Thm.dest_arg plit)
+    |> Thm.elim_implies (Thm.implies_intr plit pthm)
+    |> Thm.elim_implies (Thm.implies_intr nlit nthm)
+    |> rpair (Ord_List.union lit_ord (prune nlits) (prune plits))
+  end
+
+
+(* proof replay for hypothesis *)
+
+val dneg_rule = @{lemma "~~P ==> P" by auto}
+
+fun replay_hyp i ct =
+  if i < 0 then (Thm.assume ct, [(~i, ct)])
+  else
+    let val cu = as_prop (Thm.apply @{cterm HOL.Not} (Thm.apply @{cterm HOL.Not} (Thm.dest_arg ct)))
+    in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end
+
+
+(* proof replay for lemma *)
+
+fun replay_lemma is (thm, lits) = replay_with_lits is thm lits
+
+
+(* proof replay for reflexivity *)
+
+val refl_rule = @{thm refl}
+val refl_rule_var = Thm.dest_arg1 (Thm.dest_arg (Thm.cprop_of refl_rule))
+
+fun replay_refl ct = Thm.instantiate (Thm.match (refl_rule_var, ct)) refl_rule
+
+
+(* proof replay for symmetry *)
+
+val symm_rules = @{lemma "a = b ==> b = a" "~(a = b) ==> ~(b = a)" by simp_all}
+
+fun replay_symm thm = hd (discharges thm symm_rules)
+
+
+(* proof replay for transitivity *)
+
+val trans_rules = @{lemma
+  "~(a = b) ==> b = c ==> ~(a = c)"
+  "a = b ==> ~(b = c) ==> ~(a = c)"
+  "a = b ==> b = c ==> a = c"
+  by simp_all}
+
+fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules))
+
+
+(* proof replay for congruence *)
+
+fun replay_cong thm1 thm2 = discharge thm2 (discharge thm1 @{thm cong})
+
+
+(* proof replay for substitution *)
+
+val subst_rule1 = @{lemma "~(p a) ==> p = q ==> a = b ==> ~(q b)" by simp}
+val subst_rule2 = @{lemma "p a ==> p = q ==> a = b ==> q b" by simp}
+
+fun replay_subst thm1 thm2 thm3 =
+  subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3]
+
+
+(* proof replay *)
+
+structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord)
+
+val unclausify_rule1 = @{lemma "(~P ==> False) ==> P" by auto}
+val unclausify_rule2 = @{lemma "(P ==> False) ==> ~P" by auto}
+
+fun unclausify (thm, lits) ls =
+  (case (Thm.prop_of thm, lits) of
+    (@{const HOL.Trueprop} $ @{const HOL.False}, [(_, ct)]) =>
+      let val thm = Thm.implies_intr ct thm
+      in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end
+  | _ => (thm, Ord_List.union lit_ord lits ls))
+
+fun with_thms f tps = fold_map unclausify tps [] |>> f
+
+fun bad_premises () = raise Fail "bad number of premises"
+fun with_thms1 f = with_thms (fn [thm] => f thm | _ => bad_premises ())
+fun with_thms2 f = with_thms (fn [thm1, thm2] => f thm1 thm2 | _ => bad_premises ())
+fun with_thms3 f = with_thms (fn [thm1, thm2, thm3] => f thm1 thm2 thm3 | _ => bad_premises ())
+
+fun replay_rule (ctxt, cons, facts) prems rule =
+  (case rule of
+    Argo_Proof.Axiom i => (nth facts i, [])
+  | Argo_Proof.Taut (k, concl) => (replay_taut ctxt k (cprop_of ctxt cons concl), [])
+  | Argo_Proof.Conjunct (i, n) => with_thms1 (replay_conjunct i n) prems
+  | Argo_Proof.Rewrite c => with_thms1 (replay_rewrite ctxt c) prems
+  | Argo_Proof.Clause is => replay_clause is (fst (hd prems))
+  | Argo_Proof.Unit_Res i => replay_unit_res i (hd prems) (hd (tl prems))
+  | Argo_Proof.Hyp (i, concl) => replay_hyp i (cprop_of ctxt cons concl)
+  | Argo_Proof.Lemma is => replay_lemma is (hd prems)
+  | Argo_Proof.Refl concl => (replay_refl (cterm_of ctxt cons concl), [])
+  | Argo_Proof.Symm => with_thms1 replay_symm prems
+  | Argo_Proof.Trans => with_thms2 replay_trans prems
+  | Argo_Proof.Cong => with_thms2 replay_cong prems
+  | Argo_Proof.Subst => with_thms3 replay_subst prems
+  | _ => with_thms (ext_replay (cprop_of ctxt cons) ctxt rule) prems)
+
+fun with_cache f proof thm_cache =
+  (case Thm_Cache.lookup thm_cache (Argo_Proof.id_of proof) of
+    SOME thm => (thm, thm_cache)
+  | NONE =>
+      let val (thm, thm_cache') = f proof thm_cache
+      in (thm, Thm_Cache.update (Argo_Proof.id_of proof, thm) thm_cache') end)
+
+fun trace_step ctxt proof_id rule proofs = with_full_tracing ctxt (fn ctxt' =>
+  let
+    val id = Argo_Proof.string_of_proof_id proof_id
+    val ids = map (Argo_Proof.string_of_proof_id o Argo_Proof.id_of) proofs
+    val rule_string = Argo_Proof.string_of_rule rule
+  in full_tracing ctxt' ("  " ^ id ^ " <- " ^ space_implode " " ids ^ " . " ^ rule_string) end)
+
+fun replay_bottom_up (env as (ctxt, _, _)) proof thm_cache =
+  let
+    val (proof_id, rule, proofs) = Argo_Proof.dest proof
+    val (prems, thm_cache) = fold_map (with_cache (replay_bottom_up env)) proofs thm_cache
+    val _ = trace_step ctxt proof_id rule proofs
+  in (replay_rule env prems rule, thm_cache) end
+
+fun replay_proof env proof = with_cache (replay_bottom_up env) proof Thm_Cache.empty
+
+fun replay ctxt terms facts proof =
+  let
+    val env = (ctxt, Termtab.fold (Contab.update o swap) terms Contab.empty, facts)
+    val _ = tracing ctxt "replaying the proof"
+    val ({elapsed=t, ...}, ((thm, _), _)) = Timing.timing (replay_proof env) proof
+    val _ = tracing ctxt ("replayed the proof in " ^ string_of_int (Time.toMilliseconds t) ^ " ms")
+  in thm end
+
+
+(* normalizing goals *)
+
+fun instantiate_elim_rule thm =
+  let val ct = Drule.strip_imp_concl (Thm.cprop_of thm)
+  in
+    (case Thm.term_of ct of
+      @{const HOL.Trueprop} $ Var (_, @{typ HOL.bool}) =>
+        instantiate (Thm.dest_arg ct) @{cterm HOL.False} thm
+    | Var _ => instantiate ct @{cprop HOL.False} thm
+    | _ => thm)
+  end
+
+fun atomize_conv ctxt ct =
+  (case Thm.term_of ct of
+    @{const HOL.Trueprop} $ _ => Conv.all_conv
+  | @{const Pure.imp} $ _ $ _ =>
+      Conv.binop_conv (atomize_conv ctxt) then_conv
+      Conv.rewr_conv @{thm atomize_imp}
+  | Const (@{const_name Pure.eq}, _) $ _ $ _ =>
+      Conv.binop_conv (atomize_conv ctxt) then_conv
+      Conv.rewr_conv @{thm atomize_eq}
+  | Const (@{const_name Pure.all}, _) $ Abs _ =>
+      Conv.binder_conv (atomize_conv o snd) ctxt then_conv
+      Conv.rewr_conv @{thm atomize_all}
+  | _ => Conv.all_conv) ct
+
+fun normalize ctxt thm =
+  thm
+  |> instantiate_elim_rule
+  |> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion)
+  |> Drule.forall_intr_vars
+  |> Conv.fconv_rule (atomize_conv ctxt)
+
+
+(* prover, tactic and method *)
+
+datatype result = Satisfiable of term -> bool option | Unsatisfiable
+
+fun check props ctxt =
+  (case solve props ctxt of
+    (Proof _, ctxt') => (Unsatisfiable, ctxt')
+  | (Model model, ctxt') => (Satisfiable model, ctxt'))
+
+fun prove thms ctxt =
+  let val thms' = map (normalize ctxt) thms
+  in
+    (case solve (map Thm.prop_of thms') ctxt of
+      (Model _, ctxt') => (NONE, ctxt')
+    | (Proof (terms, proof), ctxt') => (SOME (replay ctxt' terms thms' proof), ctxt'))
+  end
+
+fun argo_tac ctxt thms =
+  CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1
+    (Conv.try_conv (Conv.rewr_conv @{thm atomize_eq})))) ctxt)
+  THEN' Tactic.resolve_tac ctxt [@{thm ccontr}]
+  THEN' Subgoal.FOCUS (fn {context, prems, ...} =>
+    (case with_timeout context (prove (thms @ prems)) context of
+      (SOME thm, _) => Tactic.resolve_tac context [thm] 1
+    | (NONE, _) => Tactical.no_tac)) ctxt
+
+val _ =
+  Theory.setup (Method.setup @{binding argo}
+    (Scan.optional Attrib.thms [] >>
+      (fn thms => fn ctxt => METHOD (fn facts =>
+        HEADGOAL (argo_tac ctxt (thms @ facts)))))
+    "Applies the Argo prover")
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/ex/Argo_Examples.thy	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,1328 @@
+(*  Title:      HOL/ex/Argo_Examples.thy
+    Author:     Sascha Boehme
+*)
+
+section \<open>Argo\<close>
+
+theory Argo_Examples
+imports Complex_Main
+begin
+
+text \<open>
+  This theory is intended to showcase and test different features of the \<open>argo\<close> proof method.
+
+  The \<open>argo\<close> proof method can be applied to propositional problems, problems involving equality
+  reasoning and problems of linear real arithmetic.
+
+  The \<open>argo\<close> proof method provides two options. To specify an upper limit of the proof methods
+  run time in seconds, use the option \<open>argo_timeout\<close>. To specify the amount of output, use the
+  option \<open>argo_trace\<close> with value \<open>none\<close> for no tracing output, value \<open>basic\<close> for viewing the
+  underlying propositions and some timings, and value \<open>full\<close> for additionally inspecting the
+  proof replay steps.
+\<close>
+
+declare[[argo_trace = full]]
+
+subsection \<open>Propositional logic\<close>
+
+notepad
+begin
+  have "True" by argo
+next
+  have "~False" by argo
+next
+  fix P :: bool
+  assume "False"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "~True"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "P"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "~~P"
+  then have "P" by argo
+next
+  fix P Q R :: bool
+  assume "P & Q & R"
+  then have "R & P & Q" by argo
+next
+  fix P Q R :: bool
+  assume "P & (Q & True & R) & (Q & P) & True"
+  then have "R & P & Q" by argo
+next
+  fix P Q1 Q2 Q3 Q4 Q5 :: bool
+  assume "Q1 & (Q2 & P & Q3) & (Q4 & ~P & Q5)"
+  then have "~True" by argo
+next
+  fix P Q1 Q2 Q3  :: bool
+  assume "(Q1 & False) & Q2 & Q3"
+  then have "P::bool" by argo
+next
+  fix P Q R :: bool
+  assume "P | Q | R"
+  then have "R | P | Q" by argo
+next
+  fix P Q R :: bool
+  assume "P | (Q | False | R) | (Q | P) | False"
+  then have "R | P | Q" by argo
+next
+  fix P Q1 Q2 Q3 Q4 :: bool
+  have "(Q1 & P & Q2) --> False | (Q3 | (Q4 | P) | False)" by argo
+next
+  fix Q1 Q2 Q3 Q4 :: bool
+  have "Q1 | (Q2 | True | Q3) | Q4" by argo
+next
+  fix P Q R :: bool
+  assume "(P & Q) | (P & ~Q) | (P & R) | (P & ~R)"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "P = True"
+  then have "P" by argo
+next
+  fix P :: bool
+  assume "False = P"
+  then have "~P" by argo
+next
+  fix P Q :: bool
+  assume "P = (~P)"
+  then have "Q" by argo
+next
+  fix P :: bool
+  have "(~P) = (~P)" by argo
+next
+  fix P Q :: bool
+  assume "P" and "~Q"
+  then have "P = (~Q)" by argo
+next
+  fix P Q :: bool
+  assume "((P::bool) = Q) | (Q = P)"
+  then have "(P --> Q) & (Q --> P)" by argo
+next
+  fix P Q :: bool
+  assume "(P::bool) = Q"
+  then have "Q = P" by argo
+next
+  fix P Q R :: bool
+  assume "if P then Q else R"
+  then have "Q | R" by argo
+next
+  fix P Q :: bool
+  assume "P | Q"
+     and "P | ~Q"
+     and "~P | Q"
+     and "~P | ~Q"
+  then have "False" by argo
+next
+  fix P Q R :: bool
+  assume "P | Q | R"
+     and "P | Q | ~R"
+     and "P | ~Q | R"
+     and "P | ~Q | ~R"
+     and "~P | Q | R"
+     and "~P | Q | ~R"
+     and "~P | ~Q | R"
+     and "~P | ~Q | ~R"
+  then have "False" by argo
+next
+  fix a b c d e f g h i j k l m n p q :: bool
+  assume "(a & b | c & d) & (e & f | g & h) | (i & j | k & l) & (m & n | p & q)"
+  then have "(a & b | c & d) & (e & f | g & h) | (i & j | k & l) & (m & n | p & q)" by argo
+next
+  fix P :: bool
+  have "P=P=P=P=P=P=P=P=P=P" by argo
+next
+  fix a b c d e f p q x :: bool
+  assume "a | b | c | d"
+     and "e | f | (a & d)"
+     and "~(a | (c & ~c)) | b"
+     and "~(b & (x | ~x)) | c"
+     and "~(d | False) | c"
+     and "~(c | (~p & (p | (q & ~q))))"
+  then have "False" by argo
+end
+
+
+subsection \<open>Equality, congruence and predicates\<close>
+
+notepad
+begin
+  fix t :: "'a"
+  have "t = t" by argo
+next
+  fix t u :: "'a"
+  assume "t = u"
+  then have "u = t" by argo
+next
+  fix s t u :: "'a"
+  assume "s = t" and "t = u"
+  then have "s = u" by argo
+next
+  fix s t u v :: "'a"
+  assume "s = t" and "t = u" and "u = v" and "u = s"
+  then have "s = v" by argo
+next
+  fix s t u v w :: "'a"
+  assume "s = t" and "t = u" and "s = v" and "v = w"
+  then have "w = u" by argo
+next
+  fix s t u a b c :: "'a"
+  assume "s = t" and "t = u" and "a = b" and "b = c"
+  then have "s = a --> c = u" by argo
+next
+  fix a b c d :: "'a"
+  assume "(a = b & b = c) | (a = d & d = c)"
+  then have "a = c" by argo
+next
+  fix a b1 b2 b3 b4 c d :: "'a"
+  assume "(a = b1 & ((b1 = b2 & b2 = b3) | (b1 = b4 & b4 = b3)) & b3 = c) | (a = d & d = c)"
+  then have "a = c" by argo
+next
+  fix a b :: "'a"
+  have "(if True then a else b) = a" by argo
+next
+  fix a b :: "'a"
+  have "(if False then a else b) = b" by argo
+next
+  fix a b :: "'a"
+  have "(if \<not>True then a else b) = b" by argo
+next
+  fix a b :: "'a"
+  have "(if \<not>False then a else b) = a" by argo
+next
+  fix P :: "bool"
+  fix a :: "'a"
+  have "(if P then a else a) = a" by argo
+next
+  fix P :: "bool"
+  fix a b c :: "'a"
+  assume "P" and "a = c"
+  then have "(if P then a else b) = c" by argo
+next
+  fix P :: "bool"
+  fix a b c :: "'a"
+  assume "~P" and "b = c"
+  then have "(if P then a else b) = c" by argo
+next
+  fix P Q :: "bool"
+  fix a b c d :: "'a"
+  assume "P" and "Q" and "a = d"
+  then have "(if P then (if Q then a else b) else c) = d" by argo
+next
+  fix a b c :: "'a"
+  assume "a \<noteq> b" and "b = c"
+  then have "a \<noteq> c" by argo
+next
+  fix a b c :: "'a"
+  assume "a \<noteq> b" and "a = c"
+  then have "c \<noteq> b" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "b \<noteq> d"
+  then have "a \<noteq> c" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "d \<noteq> b"
+  then have "a \<noteq> c" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "b \<noteq> d"
+  then have "c \<noteq> a" by argo
+next
+  fix a b c d :: "'a"
+  assume "a = b" and "c = d" and "d \<noteq> b"
+  then have "c \<noteq> a" by argo
+next
+  fix a b c d e f :: "'a"
+  assume "a \<noteq> b" and "b = c" and "b = d" and "d = e" and "a = f"
+  then have "f \<noteq> e" by argo
+next
+  fix a b :: "'a" and f :: "'a \<Rightarrow> 'a"
+  assume "a = b"
+  then have "f a = f b" by argo
+next
+  fix a b c :: "'a" and f :: "'a \<Rightarrow> 'a"
+  assume "f a = f b" and "b = c"
+  then have "f a = f c" by argo
+next
+  fix a :: "'a" and f :: "'a \<Rightarrow> 'a"
+  assume "f a = a"
+  then have "f (f a) = a" by argo
+next
+  fix a b :: "'a" and f g :: "'a \<Rightarrow> 'a"
+  assume "a = b"
+  then have "g (f a) = g (f b)" by argo
+next
+  fix a b :: "'a" and f g :: "'a \<Rightarrow> 'a"
+  assume "f a = b" and "g b = a"
+  then have "f (g (f a)) = b" by argo
+next
+  fix a b :: "'a" and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  assume "a = b"
+  then have "g a b = g b a" by argo
+next
+  fix a b :: "'a" and f :: "'a \<Rightarrow> 'a" and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  assume "f a = b"
+  then have "g (f a) b = g b (f a)" by argo
+next
+  fix a b c d e g h :: "'a" and f :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  assume "c = d" and "e = c" and "e = b" and "b = h" and "f g h = d" and "f g d = a"
+  then have "a = b" by argo
+next
+  fix a b :: "'a" and P :: "'a \<Rightarrow> bool"
+  assume "P a" and "a = b"
+  then have "P b" by argo
+next
+  fix a b :: "'a" and P :: "'a \<Rightarrow> bool"
+  assume "~ P a" and "a = b"
+  then have "~ P b" by argo
+next
+  fix a b c d :: "'a" and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assume "P a b" and "a = c" and "b = d"
+  then have "P c d" by argo
+next
+  fix a b c d :: "'a" and P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  assume "~ P a b" and "a = c" and "b = d"
+  then have "~ P c d" by argo
+end
+
+
+subsection \<open>Linear real arithmetic\<close>
+
+subsubsection \<open>Negation and subtraction\<close>
+
+notepad
+begin
+  fix a b :: real
+  have
+    "-a = -1 * a"
+    "-(-a) = a"
+    "a - b = a + -1 * b"
+    "a - (-b) = a + b"
+    by argo+
+end
+
+
+subsubsection \<open>Multiplication\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(2::real) * 3 = 6"
+    "0 * a = 0"
+    "a * 0 = 0"
+    "1 * a = a"
+    "a * 1 = a"
+    "2 * a = a * 2"
+    "2 * a * 3 = 6 * a"
+    "2 * a * 3 * 5 = 30 * a"
+    "a * 0 * b = 0"
+    "a * (0 * b) = 0"
+    "a * (b * c) = (a * b) * c"
+    "a * (b * (c * d)) = ((a * b) * c) * d"
+    "a * (b + c + d) = a * b + a * c + a * d"
+    by argo+
+end
+
+
+subsubsection \<open>Division\<close>
+
+notepad
+begin
+  fix a b c :: real
+  have
+    "(6::real) / 2 = 3"
+    "a / 0 = a / 0"
+    "a / 0 <= a / 0"
+    "~(a / 0 < a / 0)"
+    "0 / a = 0"
+    "a / 1 = a"
+    "a / 3 = 1/3 * a"
+    "6 * a / 2 = 3 * a"
+    "a / ((5 * b) / 2) = 2/5 * a / b"
+    "a / (5 * (b / 2)) = 2/5 * a / b"
+    "(a / 5) * (b / 2) = 1/10 * a * b"
+    "a / (3 * b) = 1/3 * a / b"
+    "(a + b) / 5 = 1/5 * a + 1/5 * b"
+    "a / (5 * 1/5) = a"
+    by argo+
+end
+
+
+subsubsection \<open>Addition\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "a + b = b + a"
+    "a + b + c = c + b + a"
+    "a + b + c + d = d + c + b + a"
+    "a + (b + (c + d)) = ((a + b) + c) + d"
+    "(5::real) + -3 = 2"
+    "(3::real) + 5 + -1 = 7"
+    "2 + a = a + 2"
+    "a + b + a = b + 2 * a"
+    "-1 + a + -1 + 2 + b + 5/3 + b + 1/3 + 5 * b + 2/3 = 8/3 + a + 7 * b"
+    "1 + b + b + 5 * b + 3 * a + 7 + a + 2 = 10 + 4 * a + 7 * b"
+    by argo+
+end
+
+
+subsubsection \<open>Minimum and maximum\<close>
+
+notepad
+begin
+  fix a b :: real
+  have
+    "min (3::real) 5 = 3"
+    "min (5::real) 3 = 3"
+    "min (3::real) (-5) = -5"
+    "min (-5::real) 3 = -5"
+    "min a a = a"
+    "a \<le> b \<longrightarrow> min a b = a"
+    "a > b \<longrightarrow> min a b = b"
+    "min a b \<le> a"
+    "min a b \<le> b"
+    "min a b = min b a"
+    by argo+
+next
+  fix a b :: real
+  have
+    "max (3::real) 5 = 5"
+    "max (5::real) 3 = 5"
+    "max (3::real) (-5) = 3"
+    "max (-5::real) 3 = 3"
+    "max a a = a"
+    "a \<le> b \<longrightarrow> max a b = b"
+    "a > b \<longrightarrow> max a b = a"
+    "a \<le> max a b"
+    "b \<le> max a b"
+    "max a b = max b a"
+    by argo+
+next
+  fix a b :: real
+  have
+    "min a b \<le> max a b"
+    "min a b + max a b = a + b"
+    "a < b \<longrightarrow> min a b < max a b"
+    by argo+
+end
+
+
+subsubsection \<open>Absolute value\<close>
+
+notepad
+begin
+  fix a :: real
+  have
+    "abs (3::real) = 3"
+    "abs (-3::real) = 3"
+    "0 \<le> abs a"
+    "a \<le> abs a"
+    "a \<ge> 0 \<longrightarrow> abs a = a"
+    "a < 0 \<longrightarrow> abs a = -a"
+    "abs (abs a) = abs a"
+    by argo+
+end
+
+
+subsubsection \<open>Equality\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(3::real) = 3"
+    "~((3::real) = 4)"
+    "~((4::real) = 3)"
+    "3 * a = 5 --> a = 5/3"
+    "-3 * a = 5 --> -5/3 = a"
+    "5 = 3 * a --> 5/3  = a "
+    "5 = -3 * a --> a = -5/3"
+    "2 + 3 * a = 4 --> a = 2/3"
+    "4 = 2 + 3 * a --> 2/3 = a"
+    "2 + 3 * a + 5 * b + c = 4 --> 3 * a + 5 * b + c = 2"
+    "4 = 2 + 3 * a + 5 * b + c --> 2 = 3 * a + 5 * b + c"
+    "-2 * a + b + -3 * c = 7 --> -7 = 2 * a + -1 * b + 3 * c"
+    "7 = -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c = -7"
+    "-2 * a + b + -3 * c + 4 * d = 7 --> -7 = 2 * a + -1 * b + 3 * c + -4 * d"
+    "7 = -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d = -7"
+    "a + 3 * b = 5 * c + b --> a + 2 * b + -5 * c = 0"
+    by argo+
+end
+
+
+subsubsection \<open>Less-equal\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(3::real) <= 3"
+    "(3::real) <= 4"
+    "~((4::real) <= 3)"
+    "3 * a <= 5 --> a <= 5/3"
+    "-3 * a <= 5 --> -5/3 <= a"
+    "5 <= 3 * a --> 5/3  <= a "
+    "5 <= -3 * a --> a <= -5/3"
+    "2 + 3 * a <= 4 --> a <= 2/3"
+    "4 <= 2 + 3 * a --> 2/3 <= a"
+    "2 + 3 * a + 5 * b + c <= 4 --> 3 * a + 5 * b + c <= 2"
+    "4 <= 2 + 3 * a + 5 * b + c --> 2 <= 3 * a + 5 * b + c"
+    "-2 * a + b + -3 * c <= 7 --> -7 <= 2 * a + -1 * b + 3 * c"
+    "7 <= -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c <= -7"
+    "-2 * a + b + -3 * c + 4 * d <= 7 --> -7 <= 2 * a + -1 * b + 3 * c + -4 * d"
+    "7 <= -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d <= -7"
+    "a + 3 * b <= 5 * c + b --> a + 2 * b + -5 * c <= 0"
+    by argo+
+end
+
+subsubsection \<open>Less\<close>
+
+notepad
+begin
+  fix a b c d :: real
+  have
+    "(3::real) < 4"
+    "~((3::real) < 3)"
+    "~((4::real) < 3)"
+    "3 * a < 5 --> a < 5/3"
+    "-3 * a < 5 --> -5/3 < a"
+    "5 < 3 * a --> 5/3  < a "
+    "5 < -3 * a --> a < -5/3"
+    "2 + 3 * a < 4 --> a < 2/3"
+    "4 < 2 + 3 * a --> 2/3 < a"
+    "2 + 3 * a + 5 * b + c < 4 --> 3 * a + 5 * b + c < 2"
+    "4 < 2 + 3 * a + 5 * b + c --> 2 < 3 * a + 5 * b + c"
+    "-2 * a + b + -3 * c < 7 --> -7 < 2 * a + -1 * b + 3 * c"
+    "7 < -2 * a + b + -3 * c --> 2 * a + -1 * b + 3 * c < -7"
+    "-2 * a + b + -3 * c + 4 * d < 7 --> -7 < 2 * a + -1 * b + 3 * c + -4 * d"
+    "7 < -2 * a + b + -3 * c + 4 * d --> 2 * a + -1 * b + 3 * c + -4 * d < -7"
+    "a + 3 * b < 5 * c + b --> a + 2 * b + -5 * c < 0"
+    by argo+
+end
+
+
+subsubsection \<open>Other examples\<close>
+
+notepad
+begin
+  have
+    "(0::real) < 1"
+    "(47::real) + 11 < 8 * 15"
+    by argo+
+next
+  fix a :: real
+  assume "a < 3"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  assume "a <= 3"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  assume "~(3 < a)"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  assume "~(3 <= a)"
+  then have "a < 5" "a <= 5" "~(5 < a)" "~(5 <= a)" by argo+
+next
+  fix a :: real
+  have "a < 3 | a = 3 | a > 3" by argo
+next
+  fix a b :: real
+  assume "0 < a" and "a < b"
+  then have "0 < b" by argo
+next
+  fix a b :: real
+  assume "0 < a" and "a \<le> b"
+  then have "0 \<le> b" by argo
+next
+  fix a b :: real
+  assume "0 \<le> a" and "a < b"
+  then have "0 \<le> b" by argo
+next
+  fix a b :: real
+  assume "0 \<le> a" and "a \<le> b"
+  then have "0 \<le> b" by argo
+next
+  fix a b c :: real
+  assume "2 \<le> a" and "3 \<le> b" and "c \<le> 5"
+  then have "-2 * a + -3 * b + 5 * c < 13" by argo
+next
+  fix a b c :: real
+  assume "2 \<le> a" and "3 \<le> b" and "c \<le> 5"
+  then have "-2 * a + -3 * b + 5 * c \<le> 12" by argo
+next
+  fix a b :: real
+  assume "a = 2" and "b = 3"
+  then have "a + b > 5 \<or> a < b" by argo
+next
+  fix a b c :: real
+  assume "5 < b + c" and "a + c < 0" and "a > 0"
+  then have "b > 0" by argo
+next
+  fix a b c :: real
+  assume "a + b < 7" and "5 < b + c" and "a + c < 0" and "a > 0"
+  then have "0 < b \<and> b < 7" by argo
+next
+  fix a b c :: real
+  assume "a < b" and "b < c" and "c < a"
+  then have "False" by argo
+next
+  fix a b :: real
+  assume "a - 5 > b"
+  then have "b < a" by argo
+next
+  fix a b :: real
+  have "(a - b) - a = (a - a) - b" by argo
+next
+  fix n m n' m' :: real
+  have "
+    (n < m & m < n') | (n < m & m = n') | (n < n' & n' < m) |
+    (n = n' & n' < m) | (n = m & m < n') |
+    (n' < m & m < n) | (n' < m & m = n) |
+    (n' < n & n < m) | (n' = n & n < m) | (n' = m & m < n) |
+    (m < n & n < n') | (m < n & n' = n) | (m < n' & n' < n) |
+    (m = n & n < n') | (m = n' & n' < n) |
+    (n' = m & m = n)"
+    by argo
+end
+
+
+subsection \<open>Larger examples\<close>
+
+declare[[argo_trace = basic, argo_timeout = 60]]
+
+
+text \<open>Translated from TPTP problem library: PUZ015-2.006.dimacs\<close>
+
+lemma assumes 1: "~x0"
+  and 2: "~x30"
+  and 3: "~x29"
+  and 4: "~x59"
+  and 5: "x1 | x31 | x0"
+  and 6: "x2 | x32 | x1"
+  and 7: "x3 | x33 | x2"
+  and 8: "x4 | x34 | x3"
+  and 9: "x35 | x4"
+  and 10: "x5 | x36 | x30"
+  and 11: "x6 | x37 | x5 | x31"
+  and 12: "x7 | x38 | x6 | x32"
+  and 13: "x8 | x39 | x7 | x33"
+  and 14: "x9 | x40 | x8 | x34"
+  and 15: "x41 | x9 | x35"
+  and 16: "x10 | x42 | x36"
+  and 17: "x11 | x43 | x10 | x37"
+  and 18: "x12 | x44 | x11 | x38"
+  and 19: "x13 | x45 | x12 | x39"
+  and 20: "x14 | x46 | x13 | x40"
+  and 21: "x47 | x14 | x41"
+  and 22: "x15 | x48 | x42"
+  and 23: "x16 | x49 | x15 | x43"
+  and 24: "x17 | x50 | x16 | x44"
+  and 25: "x18 | x51 | x17 | x45"
+  and 26: "x19 | x52 | x18 | x46"
+  and 27: "x53 | x19 | x47"
+  and 28: "x20 | x54 | x48"
+  and 29: "x21 | x55 | x20 | x49"
+  and 30: "x22 | x56 | x21 | x50"
+  and 31: "x23 | x57 | x22 | x51"
+  and 32: "x24 | x58 | x23 | x52"
+  and 33: "x59 | x24 | x53"
+  and 34: "x25 | x54"
+  and 35: "x26 | x25 | x55"
+  and 36: "x27 | x26 | x56"
+  and 37: "x28 | x27 | x57"
+  and 38: "x29 | x28 | x58"
+  and 39: "~x1 | ~x31"
+  and 40: "~x1 | ~x0"
+  and 41: "~x31 | ~x0"
+  and 42: "~x2 | ~x32"
+  and 43: "~x2 | ~x1"
+  and 44: "~x32 | ~x1"
+  and 45: "~x3 | ~x33"
+  and 46: "~x3 | ~x2"
+  and 47: "~x33 | ~x2"
+  and 48: "~x4 | ~x34"
+  and 49: "~x4 | ~x3"
+  and 50: "~x34 | ~x3"
+  and 51: "~x35 | ~x4"
+  and 52: "~x5 | ~x36"
+  and 53: "~x5 | ~x30"
+  and 54: "~x36 | ~x30"
+  and 55: "~x6 | ~x37"
+  and 56: "~x6 | ~x5"
+  and 57: "~x6 | ~x31"
+  and 58: "~x37 | ~x5"
+  and 59: "~x37 | ~x31"
+  and 60: "~x5 | ~x31"
+  and 61: "~x7 | ~x38"
+  and 62: "~x7 | ~x6"
+  and 63: "~x7 | ~x32"
+  and 64: "~x38 | ~x6"
+  and 65: "~x38 | ~x32"
+  and 66: "~x6 | ~x32"
+  and 67: "~x8 | ~x39"
+  and 68: "~x8 | ~x7"
+  and 69: "~x8 | ~x33"
+  and 70: "~x39 | ~x7"
+  and 71: "~x39 | ~x33"
+  and 72: "~x7 | ~x33"
+  and 73: "~x9 | ~x40"
+  and 74: "~x9 | ~x8"
+  and 75: "~x9 | ~x34"
+  and 76: "~x40 | ~x8"
+  and 77: "~x40 | ~x34"
+  and 78: "~x8 | ~x34"
+  and 79: "~x41 | ~x9"
+  and 80: "~x41 | ~x35"
+  and 81: "~x9 | ~x35"
+  and 82: "~x10 | ~x42"
+  and 83: "~x10 | ~x36"
+  and 84: "~x42 | ~x36"
+  and 85: "~x11 | ~x43"
+  and 86: "~x11 | ~x10"
+  and 87: "~x11 | ~x37"
+  and 88: "~x43 | ~x10"
+  and 89: "~x43 | ~x37"
+  and 90: "~x10 | ~x37"
+  and 91: "~x12 | ~x44"
+  and 92: "~x12 | ~x11"
+  and 93: "~x12 | ~x38"
+  and 94: "~x44 | ~x11"
+  and 95: "~x44 | ~x38"
+  and 96: "~x11 | ~x38"
+  and 97: "~x13 | ~x45"
+  and 98: "~x13 | ~x12"
+  and 99: "~x13 | ~x39"
+  and 100: "~x45 | ~x12"
+  and 101: "~x45 | ~x39"
+  and 102: "~x12 | ~x39"
+  and 103: "~x14 | ~x46"
+  and 104: "~x14 | ~x13"
+  and 105: "~x14 | ~x40"
+  and 106: "~x46 | ~x13"
+  and 107: "~x46 | ~x40"
+  and 108: "~x13 | ~x40"
+  and 109: "~x47 | ~x14"
+  and 110: "~x47 | ~x41"
+  and 111: "~x14 | ~x41"
+  and 112: "~x15 | ~x48"
+  and 113: "~x15 | ~x42"
+  and 114: "~x48 | ~x42"
+  and 115: "~x16 | ~x49"
+  and 116: "~x16 | ~x15"
+  and 117: "~x16 | ~x43"
+  and 118: "~x49 | ~x15"
+  and 119: "~x49 | ~x43"
+  and 120: "~x15 | ~x43"
+  and 121: "~x17 | ~x50"
+  and 122: "~x17 | ~x16"
+  and 123: "~x17 | ~x44"
+  and 124: "~x50 | ~x16"
+  and 125: "~x50 | ~x44"
+  and 126: "~x16 | ~x44"
+  and 127: "~x18 | ~x51"
+  and 128: "~x18 | ~x17"
+  and 129: "~x18 | ~x45"
+  and 130: "~x51 | ~x17"
+  and 131: "~x51 | ~x45"
+  and 132: "~x17 | ~x45"
+  and 133: "~x19 | ~x52"
+  and 134: "~x19 | ~x18"
+  and 135: "~x19 | ~x46"
+  and 136: "~x52 | ~x18"
+  and 137: "~x52 | ~x46"
+  and 138: "~x18 | ~x46"
+  and 139: "~x53 | ~x19"
+  and 140: "~x53 | ~x47"
+  and 141: "~x19 | ~x47"
+  and 142: "~x20 | ~x54"
+  and 143: "~x20 | ~x48"
+  and 144: "~x54 | ~x48"
+  and 145: "~x21 | ~x55"
+  and 146: "~x21 | ~x20"
+  and 147: "~x21 | ~x49"
+  and 148: "~x55 | ~x20"
+  and 149: "~x55 | ~x49"
+  and 150: "~x20 | ~x49"
+  and 151: "~x22 | ~x56"
+  and 152: "~x22 | ~x21"
+  and 153: "~x22 | ~x50"
+  and 154: "~x56 | ~x21"
+  and 155: "~x56 | ~x50"
+  and 156: "~x21 | ~x50"
+  and 157: "~x23 | ~x57"
+  and 158: "~x23 | ~x22"
+  and 159: "~x23 | ~x51"
+  and 160: "~x57 | ~x22"
+  and 161: "~x57 | ~x51"
+  and 162: "~x22 | ~x51"
+  and 163: "~x24 | ~x58"
+  and 164: "~x24 | ~x23"
+  and 165: "~x24 | ~x52"
+  and 166: "~x58 | ~x23"
+  and 167: "~x58 | ~x52"
+  and 168: "~x23 | ~x52"
+  and 169: "~x59 | ~x24"
+  and 170: "~x59 | ~x53"
+  and 171: "~x24 | ~x53"
+  and 172: "~x25 | ~x54"
+  and 173: "~x26 | ~x25"
+  and 174: "~x26 | ~x55"
+  and 175: "~x25 | ~x55"
+  and 176: "~x27 | ~x26"
+  and 177: "~x27 | ~x56"
+  and 178: "~x26 | ~x56"
+  and 179: "~x28 | ~x27"
+  and 180: "~x28 | ~x57"
+  and 181: "~x27 | ~x57"
+  and 182: "~x29 | ~x28"
+  and 183: "~x29 | ~x58"
+  and 184: "~x28 | ~x58"
+  shows "False"
+    using assms
+    by argo
+
+
+text \<open>Translated from TPTP problem library: MSC007-1.008.dimacs\<close>
+
+lemma assumes 1: "x0 | x1 | x2 | x3 | x4 | x5 | x6"
+  and 2: "x7 | x8 | x9 | x10 | x11 | x12 | x13"
+  and 3: "x14 | x15 | x16 | x17 | x18 | x19 | x20"
+  and 4: "x21 | x22 | x23 | x24 | x25 | x26 | x27"
+  and 5: "x28 | x29 | x30 | x31 | x32 | x33 | x34"
+  and 6: "x35 | x36 | x37 | x38 | x39 | x40 | x41"
+  and 7: "x42 | x43 | x44 | x45 | x46 | x47 | x48"
+  and 8: "x49 | x50 | x51 | x52 | x53 | x54 | x55"
+  and 9: "~x0 | ~x7"
+  and 10: "~x0 | ~x14"
+  and 11: "~x0 | ~x21"
+  and 12: "~x0 | ~x28"
+  and 13: "~x0 | ~x35"
+  and 14: "~x0 | ~x42"
+  and 15: "~x0 | ~x49"
+  and 16: "~x7 | ~x14"
+  and 17: "~x7 | ~x21"
+  and 18: "~x7 | ~x28"
+  and 19: "~x7 | ~x35"
+  and 20: "~x7 | ~x42"
+  and 21: "~x7 | ~x49"
+  and 22: "~x14 | ~x21"
+  and 23: "~x14 | ~x28"
+  and 24: "~x14 | ~x35"
+  and 25: "~x14 | ~x42"
+  and 26: "~x14 | ~x49"
+  and 27: "~x21 | ~x28"
+  and 28: "~x21 | ~x35"
+  and 29: "~x21 | ~x42"
+  and 30: "~x21 | ~x49"
+  and 31: "~x28 | ~x35"
+  and 32: "~x28 | ~x42"
+  and 33: "~x28 | ~x49"
+  and 34: "~x35 | ~x42"
+  and 35: "~x35 | ~x49"
+  and 36: "~x42 | ~x49"
+  and 37: "~x1 | ~x8"
+  and 38: "~x1 | ~x15"
+  and 39: "~x1 | ~x22"
+  and 40: "~x1 | ~x29"
+  and 41: "~x1 | ~x36"
+  and 42: "~x1 | ~x43"
+  and 43: "~x1 | ~x50"
+  and 44: "~x8 | ~x15"
+  and 45: "~x8 | ~x22"
+  and 46: "~x8 | ~x29"
+  and 47: "~x8 | ~x36"
+  and 48: "~x8 | ~x43"
+  and 49: "~x8 | ~x50"
+  and 50: "~x15 | ~x22"
+  and 51: "~x15 | ~x29"
+  and 52: "~x15 | ~x36"
+  and 53: "~x15 | ~x43"
+  and 54: "~x15 | ~x50"
+  and 55: "~x22 | ~x29"
+  and 56: "~x22 | ~x36"
+  and 57: "~x22 | ~x43"
+  and 58: "~x22 | ~x50"
+  and 59: "~x29 | ~x36"
+  and 60: "~x29 | ~x43"
+  and 61: "~x29 | ~x50"
+  and 62: "~x36 | ~x43"
+  and 63: "~x36 | ~x50"
+  and 64: "~x43 | ~x50"
+  and 65: "~x2 | ~x9"
+  and 66: "~x2 | ~x16"
+  and 67: "~x2 | ~x23"
+  and 68: "~x2 | ~x30"
+  and 69: "~x2 | ~x37"
+  and 70: "~x2 | ~x44"
+  and 71: "~x2 | ~x51"
+  and 72: "~x9 | ~x16"
+  and 73: "~x9 | ~x23"
+  and 74: "~x9 | ~x30"
+  and 75: "~x9 | ~x37"
+  and 76: "~x9 | ~x44"
+  and 77: "~x9 | ~x51"
+  and 78: "~x16 | ~x23"
+  and 79: "~x16 | ~x30"
+  and 80: "~x16 | ~x37"
+  and 81: "~x16 | ~x44"
+  and 82: "~x16 | ~x51"
+  and 83: "~x23 | ~x30"
+  and 84: "~x23 | ~x37"
+  and 85: "~x23 | ~x44"
+  and 86: "~x23 | ~x51"
+  and 87: "~x30 | ~x37"
+  and 88: "~x30 | ~x44"
+  and 89: "~x30 | ~x51"
+  and 90: "~x37 | ~x44"
+  and 91: "~x37 | ~x51"
+  and 92: "~x44 | ~x51"
+  and 93: "~x3 | ~x10"
+  and 94: "~x3 | ~x17"
+  and 95: "~x3 | ~x24"
+  and 96: "~x3 | ~x31"
+  and 97: "~x3 | ~x38"
+  and 98: "~x3 | ~x45"
+  and 99: "~x3 | ~x52"
+  and 100: "~x10 | ~x17"
+  and 101: "~x10 | ~x24"
+  and 102: "~x10 | ~x31"
+  and 103: "~x10 | ~x38"
+  and 104: "~x10 | ~x45"
+  and 105: "~x10 | ~x52"
+  and 106: "~x17 | ~x24"
+  and 107: "~x17 | ~x31"
+  and 108: "~x17 | ~x38"
+  and 109: "~x17 | ~x45"
+  and 110: "~x17 | ~x52"
+  and 111: "~x24 | ~x31"
+  and 112: "~x24 | ~x38"
+  and 113: "~x24 | ~x45"
+  and 114: "~x24 | ~x52"
+  and 115: "~x31 | ~x38"
+  and 116: "~x31 | ~x45"
+  and 117: "~x31 | ~x52"
+  and 118: "~x38 | ~x45"
+  and 119: "~x38 | ~x52"
+  and 120: "~x45 | ~x52"
+  and 121: "~x4 | ~x11"
+  and 122: "~x4 | ~x18"
+  and 123: "~x4 | ~x25"
+  and 124: "~x4 | ~x32"
+  and 125: "~x4 | ~x39"
+  and 126: "~x4 | ~x46"
+  and 127: "~x4 | ~x53"
+  and 128: "~x11 | ~x18"
+  and 129: "~x11 | ~x25"
+  and 130: "~x11 | ~x32"
+  and 131: "~x11 | ~x39"
+  and 132: "~x11 | ~x46"
+  and 133: "~x11 | ~x53"
+  and 134: "~x18 | ~x25"
+  and 135: "~x18 | ~x32"
+  and 136: "~x18 | ~x39"
+  and 137: "~x18 | ~x46"
+  and 138: "~x18 | ~x53"
+  and 139: "~x25 | ~x32"
+  and 140: "~x25 | ~x39"
+  and 141: "~x25 | ~x46"
+  and 142: "~x25 | ~x53"
+  and 143: "~x32 | ~x39"
+  and 144: "~x32 | ~x46"
+  and 145: "~x32 | ~x53"
+  and 146: "~x39 | ~x46"
+  and 147: "~x39 | ~x53"
+  and 148: "~x46 | ~x53"
+  and 149: "~x5 | ~x12"
+  and 150: "~x5 | ~x19"
+  and 151: "~x5 | ~x26"
+  and 152: "~x5 | ~x33"
+  and 153: "~x5 | ~x40"
+  and 154: "~x5 | ~x47"
+  and 155: "~x5 | ~x54"
+  and 156: "~x12 | ~x19"
+  and 157: "~x12 | ~x26"
+  and 158: "~x12 | ~x33"
+  and 159: "~x12 | ~x40"
+  and 160: "~x12 | ~x47"
+  and 161: "~x12 | ~x54"
+  and 162: "~x19 | ~x26"
+  and 163: "~x19 | ~x33"
+  and 164: "~x19 | ~x40"
+  and 165: "~x19 | ~x47"
+  and 166: "~x19 | ~x54"
+  and 167: "~x26 | ~x33"
+  and 168: "~x26 | ~x40"
+  and 169: "~x26 | ~x47"
+  and 170: "~x26 | ~x54"
+  and 171: "~x33 | ~x40"
+  and 172: "~x33 | ~x47"
+  and 173: "~x33 | ~x54"
+  and 174: "~x40 | ~x47"
+  and 175: "~x40 | ~x54"
+  and 176: "~x47 | ~x54"
+  and 177: "~x6 | ~x13"
+  and 178: "~x6 | ~x20"
+  and 179: "~x6 | ~x27"
+  and 180: "~x6 | ~x34"
+  and 181: "~x6 | ~x41"
+  and 182: "~x6 | ~x48"
+  and 183: "~x6 | ~x55"
+  and 184: "~x13 | ~x20"
+  and 185: "~x13 | ~x27"
+  and 186: "~x13 | ~x34"
+  and 187: "~x13 | ~x41"
+  and 188: "~x13 | ~x48"
+  and 189: "~x13 | ~x55"
+  and 190: "~x20 | ~x27"
+  and 191: "~x20 | ~x34"
+  and 192: "~x20 | ~x41"
+  and 193: "~x20 | ~x48"
+  and 194: "~x20 | ~x55"
+  and 195: "~x27 | ~x34"
+  and 196: "~x27 | ~x41"
+  and 197: "~x27 | ~x48"
+  and 198: "~x27 | ~x55"
+  and 199: "~x34 | ~x41"
+  and 200: "~x34 | ~x48"
+  and 201: "~x34 | ~x55"
+  and 202: "~x41 | ~x48"
+  and 203: "~x41 | ~x55"
+  and 204: "~x48 | ~x55"
+  shows "False"
+    using assms
+    by argo
+
+
+lemma "0 \<le> (yc::real) \<and>
+       0 \<le> yd \<and> 0 \<le> yb \<and> 0 \<le> ya \<Longrightarrow>
+       0 \<le> yf \<and>
+       0 \<le> xh \<and> 0 \<le> ye \<and> 0 \<le> yg \<Longrightarrow>
+       0 \<le> yw \<and> 0 \<le> xs \<and> 0 \<le> yu \<Longrightarrow>
+       0 \<le> aea \<and> 0 \<le> aee \<and> 0 \<le> aed \<Longrightarrow>
+       0 \<le> zy \<and> 0 \<le> xz \<and> 0 \<le> zw \<Longrightarrow>
+       0 \<le> zb \<and>
+       0 \<le> za \<and> 0 \<le> yy \<and> 0 \<le> yz \<Longrightarrow>
+       0 \<le> zp \<and> 0 \<le> zo \<and> 0 \<le> yq \<Longrightarrow>
+       0 \<le> adp \<and> 0 \<le> aeb \<and> 0 \<le> aec \<Longrightarrow>
+       0 \<le> acm \<and> 0 \<le> aco \<and> 0 \<le> acn \<Longrightarrow>
+       0 \<le> abl \<Longrightarrow>
+       0 \<le> zr \<and> 0 \<le> zq \<and> 0 \<le> abh \<Longrightarrow>
+       0 \<le> abq \<and> 0 \<le> zd \<and> 0 \<le> abo \<Longrightarrow>
+       0 \<le> acd \<and>
+       0 \<le> acc \<and> 0 \<le> xi \<and> 0 \<le> acb \<Longrightarrow>
+       0 \<le> acp \<and> 0 \<le> acr \<and> 0 \<le> acq \<Longrightarrow>
+       0 \<le> xw \<and>
+       0 \<le> xr \<and> 0 \<le> xv \<and> 0 \<le> xu \<Longrightarrow>
+       0 \<le> zc \<and> 0 \<le> acg \<and> 0 \<le> ach \<Longrightarrow>
+       0 \<le> zt \<and> 0 \<le> zs \<and> 0 \<le> xy \<Longrightarrow>
+       0 \<le> ady \<and> 0 \<le> adw \<and> 0 \<le> zg \<Longrightarrow>
+       0 \<le> abd \<and>
+       0 \<le> abc \<and> 0 \<le> yr \<and> 0 \<le> abb \<Longrightarrow>
+       0 \<le> adi \<and>
+       0 \<le> x \<and> 0 \<le> adh \<and> 0 \<le> xa \<Longrightarrow>
+       0 \<le> aak \<and> 0 \<le> aai \<and> 0 \<le> aad \<Longrightarrow>
+       0 \<le> aba \<and> 0 \<le> zh \<and> 0 \<le> aay \<Longrightarrow>
+       0 \<le> abg \<and> 0 \<le> ys \<and> 0 \<le> abe \<Longrightarrow>
+       0 \<le> abs1 \<and>
+       0 \<le> yt \<and> 0 \<le> abr \<and> 0 \<le> zu \<Longrightarrow>
+       0 \<le> abv \<and>
+       0 \<le> zn \<and> 0 \<le> abw \<and> 0 \<le> zm \<Longrightarrow>
+       0 \<le> adl \<and> 0 \<le> adn \<Longrightarrow>
+       0 \<le> acf \<and> 0 \<le> aca \<Longrightarrow>
+       0 \<le> ads \<and> 0 \<le> aaq \<Longrightarrow>
+       0 \<le> ada \<Longrightarrow>
+       0 \<le> aaf \<and> 0 \<le> aac \<and> 0 \<le> aag \<Longrightarrow>
+       0 \<le> aal \<and>
+       0 \<le> acu \<and> 0 \<le> acs \<and> 0 \<le> act \<Longrightarrow>
+       0 \<le> aas \<and> 0 \<le> xb \<and> 0 \<le> aat \<Longrightarrow>
+       0 \<le> zk \<and> 0 \<le> zj \<and> 0 \<le> zi \<Longrightarrow>
+       0 \<le> ack \<and>
+       0 \<le> acj \<and> 0 \<le> xc \<and> 0 \<le> aci \<Longrightarrow>
+       0 \<le> aav \<and> 0 \<le> aah \<and> 0 \<le> xd \<Longrightarrow>
+       0 \<le> abt \<and>
+       0 \<le> xo \<and> 0 \<le> abu \<and> 0 \<le> xn \<Longrightarrow>
+       0 \<le> adc \<and>
+       0 \<le> abz \<and> 0 \<le> adc \<and> 0 \<le> abz \<Longrightarrow>
+       0 \<le> xt \<and>
+       0 \<le> zz \<and> 0 \<le> aab \<and> 0 \<le> aaa \<Longrightarrow>
+       0 \<le> adq \<and>
+       0 \<le> xl \<and> 0 \<le> adr \<and> 0 \<le> adb \<Longrightarrow>
+       0 \<le> zf \<and> 0 \<le> yh \<and> 0 \<le> yi \<Longrightarrow>
+       0 \<le> aao \<and> 0 \<le> aam \<and> 0 \<le> xe \<Longrightarrow>
+       0 \<le> abk \<and>
+       0 \<le> aby \<and> 0 \<le> abj \<and> 0 \<le> abx \<Longrightarrow>
+       0 \<le> yp \<Longrightarrow>
+       0 \<le> yl \<and> 0 \<le> yj \<and> 0 \<le> ym \<Longrightarrow>
+       0 \<le> acw \<Longrightarrow>
+       0 \<le> adk \<and>
+       0 \<le> adg \<and> 0 \<le> adj \<and> 0 \<le> adf \<Longrightarrow>
+       0 \<le> adv \<and> 0 \<le> xf \<and> 0 \<le> adu \<Longrightarrow>
+       yc + yd + yb + ya = 1 \<Longrightarrow>
+       yf + xh + ye + yg = 1 \<Longrightarrow>
+       yw + xs + yu = 1 \<Longrightarrow>
+       aea + aee + aed = 1 \<Longrightarrow>
+       zy + xz + zw = 1 \<Longrightarrow>
+       zb + za + yy + yz = 1 \<Longrightarrow>
+       zp + zo + yq = 1 \<Longrightarrow>
+       adp + aeb + aec = 1 \<Longrightarrow>
+       acm + aco + acn = 1 \<Longrightarrow>
+       abl + abl = 1 \<Longrightarrow>
+       zr + zq + abh = 1 \<Longrightarrow>
+       abq + zd + abo = 1 \<Longrightarrow>
+       acd + acc + xi + acb = 1 \<Longrightarrow>
+       acp + acr + acq = 1 \<Longrightarrow>
+       xw + xr + xv + xu = 1 \<Longrightarrow>
+       zc + acg + ach = 1 \<Longrightarrow>
+       zt + zs + xy = 1 \<Longrightarrow>
+       ady + adw + zg = 1 \<Longrightarrow>
+       abd + abc + yr + abb = 1 \<Longrightarrow>
+       adi + x + adh + xa = 1 \<Longrightarrow>
+       aak + aai + aad = 1 \<Longrightarrow>
+       aba + zh + aay = 1 \<Longrightarrow>
+       abg + ys + abe = 1 \<Longrightarrow>
+       abs1 + yt + abr + zu = 1 \<Longrightarrow>
+       abv + zn + abw + zm = 1 \<Longrightarrow>
+       adl + adn = 1 \<Longrightarrow>
+       acf + aca = 1 \<Longrightarrow>
+       ads + aaq = 1 \<Longrightarrow>
+       ada + ada = 1 \<Longrightarrow>
+       aaf + aac + aag = 1 \<Longrightarrow>
+       aal + acu + acs + act = 1 \<Longrightarrow>
+       aas + xb + aat = 1 \<Longrightarrow>
+       zk + zj + zi = 1 \<Longrightarrow>
+       ack + acj + xc + aci = 1 \<Longrightarrow>
+       aav + aah + xd = 1 \<Longrightarrow>
+       abt + xo + abu + xn = 1 \<Longrightarrow>
+       adc + abz + adc + abz = 1 \<Longrightarrow>
+       xt + zz + aab + aaa = 1 \<Longrightarrow>
+       adq + xl + adr + adb = 1 \<Longrightarrow>
+       zf + yh + yi = 1 \<Longrightarrow>
+       aao + aam + xe = 1 \<Longrightarrow>
+       abk + aby + abj + abx = 1 \<Longrightarrow>
+       yp + yp = 1 \<Longrightarrow>
+       yl + yj + ym = 1 \<Longrightarrow>
+       acw + acw + acw + acw = 1 \<Longrightarrow>
+       adk + adg + adj + adf = 1 \<Longrightarrow>
+       adv + xf + adu = 1 \<Longrightarrow>
+       yd = 0 \<or> yb = 0 \<Longrightarrow>
+       xh = 0 \<or> ye = 0 \<Longrightarrow>
+       yy = 0 \<or> za = 0 \<Longrightarrow>
+       acc = 0 \<or> xi = 0 \<Longrightarrow>
+       xv = 0 \<or> xr = 0 \<Longrightarrow>
+       yr = 0 \<or> abc = 0 \<Longrightarrow>
+       zn = 0 \<or> abw = 0 \<Longrightarrow>
+       xo = 0 \<or> abu = 0 \<Longrightarrow>
+       xl = 0 \<or> adr = 0 \<Longrightarrow>
+       (yr + abd < abl \<or>
+        yr + (abd + abb) < 1) \<or>
+       yr + abd = abl \<and>
+       yr + (abd + abb) = 1 \<Longrightarrow>
+       adb + adr < xn + abu \<or>
+       adb + adr = xn + abu \<Longrightarrow>
+       (abl < abt \<or> abl < abt + xo) \<or>
+       abl = abt \<and> abl = abt + xo \<Longrightarrow>
+       yd + yc < abc + abd \<or>
+       yd + yc = abc + abd \<Longrightarrow>
+       aca < abb + yr \<or> aca = abb + yr \<Longrightarrow>
+       acb + acc < xu + xv \<or>
+       acb + acc = xu + xv \<Longrightarrow>
+       (yq < xu + xr \<or>
+        yq + zp < xu + (xr + xw)) \<or>
+       yq = xu + xr \<and>
+       yq + zp = xu + (xr + xw) \<Longrightarrow>
+       (zw < xw \<or>
+        zw < xw + xv \<or>
+        zw + zy < xw + (xv + xu)) \<or>
+       zw = xw \<and>
+       zw = xw + xv \<and>
+       zw + zy = xw + (xv + xu) \<Longrightarrow>
+       xs + yw < zs + zt \<or>
+       xs + yw = zs + zt \<Longrightarrow>
+       aab + xt < ye + yf \<or>
+       aab + xt = ye + yf \<Longrightarrow>
+       (ya + yb < yg + ye \<or>
+        ya + (yb + yc) < yg + (ye + yf)) \<or>
+       ya + yb = yg + ye \<and>
+       ya + (yb + yc) = yg + (ye + yf) \<Longrightarrow>
+       (xu + xv < acb + acc \<or>
+        xu + (xv + xw) < acb + (acc + acd)) \<or>
+       xu + xv = acb + acc \<and>
+       xu + (xv + xw) = acb + (acc + acd) \<Longrightarrow>
+       (zs < xz + zy \<or>
+        zs + xy < xz + (zy + zw)) \<or>
+       zs = xz + zy \<and>
+       zs + xy = xz + (zy + zw) \<Longrightarrow>
+       (zs + zt < xz + zy \<or>
+        zs + (zt + xy) < xz + (zy + zw)) \<or>
+       zs + zt = xz + zy \<and>
+       zs + (zt + xy) = xz + (zy + zw) \<Longrightarrow>
+       yg + ye < ya + yb \<or>
+       yg + ye = ya + yb \<Longrightarrow>
+       (abd < yc \<or> abd + abc < yc + yd) \<or>
+       abd = yc \<and> abd + abc = yc + yd \<Longrightarrow>
+       (ye + yf < adr + adq \<or>
+        ye + (yf + yg) < adr + (adq + adb)) \<or>
+       ye + yf = adr + adq \<and>
+       ye + (yf + yg) = adr + (adq + adb) \<Longrightarrow>
+       yh + yi < ym + yj \<or>
+       yh + yi = ym + yj \<Longrightarrow>
+       (abq < yl \<or> abq + abo < yl + ym) \<or>
+       abq = yl \<and> abq + abo = yl + ym \<Longrightarrow>
+       (yp < zp \<or>
+        yp < zp + zo \<or> 1 < zp + (zo + yq)) \<or>
+       yp = zp \<and>
+       yp = zp + zo \<and> 1 = zp + (zo + yq) \<Longrightarrow>
+       (abb + yr < aca \<or>
+        abb + (yr + abd) < aca + acf) \<or>
+       abb + yr = aca \<and>
+       abb + (yr + abd) = aca + acf \<Longrightarrow>
+       adw + zg < abe + ys \<or>
+       adw + zg = abe + ys \<Longrightarrow>
+       zd + abq < ys + abg \<or>
+       zd + abq = ys + abg \<Longrightarrow>
+       yt + abs1 < aby + abk \<or>
+       yt + abs1 = aby + abk \<Longrightarrow>
+       (yu < abx \<or>
+        yu < abx + aby \<or>
+        yu + yw < abx + (aby + abk)) \<or>
+       yu = abx \<and>
+       yu = abx + aby \<and>
+       yu + yw = abx + (aby + abk) \<Longrightarrow>
+       aaf < adv \<or> aaf = adv \<Longrightarrow>
+       abj + abk < yy + zb \<or>
+       abj + abk = yy + zb \<Longrightarrow>
+       (abb < yz \<or>
+        abb + abc < yz + za \<or>
+        abb + (abc + abd) < yz + (za + zb)) \<or>
+       abb = yz \<and>
+       abb + abc = yz + za \<and>
+       abb + (abc + abd) = yz + (za + zb) \<Longrightarrow>
+       (acg + zc < zd + abq \<or>
+        acg + (zc + ach)
+        < zd + (abq + abo)) \<or>
+       acg + zc = zd + abq \<and>
+       acg + (zc + ach) =
+       zd + (abq + abo) \<Longrightarrow>
+       zf < acm \<or> zf = acm \<Longrightarrow>
+       (zg + ady < acn + acm \<or>
+        zg + (ady + adw)
+        < acn + (acm + aco)) \<or>
+       zg + ady = acn + acm \<and>
+       zg + (ady + adw) =
+       acn + (acm + aco) \<Longrightarrow>
+       aay + zh < zi + zj \<or>
+       aay + zh = zi + zj \<Longrightarrow>
+       zy < zk \<or> zy = zk \<Longrightarrow>
+       (adn < zm + zn \<or>
+        adn + adl < zm + (zn + abv)) \<or>
+       adn = zm + zn \<and>
+       adn + adl = zm + (zn + abv) \<Longrightarrow>
+       zo + zp < zs + zt \<or>
+       zo + zp = zs + zt \<Longrightarrow>
+       zq + zr < zs + zt \<or>
+       zq + zr = zs + zt \<Longrightarrow>
+       (aai < adi \<or> aai < adi + adh) \<or>
+       aai = adi \<and> aai = adi + adh \<Longrightarrow>
+       (abr < acj \<or>
+        abr + (abs1 + zu)
+        < acj + (aci + ack)) \<or>
+       abr = acj \<and>
+       abr + (abs1 + zu) =
+       acj + (aci + ack) \<Longrightarrow>
+       (abl < zw \<or> 1 < zw + zy) \<or>
+       abl = zw \<and> 1 = zw + zy \<Longrightarrow>
+       (zz + aaa < act + acu \<or>
+        zz + (aaa + aab)
+        < act + (acu + aal)) \<or>
+       zz + aaa = act + acu \<and>
+       zz + (aaa + aab) =
+       act + (acu + aal) \<Longrightarrow>
+       (aam < aac \<or> aam + aao < aac + aaf) \<or>
+       aam = aac \<and> aam + aao = aac + aaf \<Longrightarrow>
+       (aak < aaf \<or> aak + aad < aaf + aag) \<or>
+       aak = aaf \<and> aak + aad = aaf + aag \<Longrightarrow>
+       (aah < aai \<or> aah + aav < aai + aak) \<or>
+       aah = aai \<and> aah + aav = aai + aak \<Longrightarrow>
+       act + (acu + aal) < aam + aao \<or>
+       act + (acu + aal) = aam + aao \<Longrightarrow>
+       (ads < aat \<or> 1 < aat + aas) \<or>
+       ads = aat \<and> 1 = aat + aas \<Longrightarrow>
+       (aba < aas \<or> aba + aay < aas + aat) \<or>
+       aba = aas \<and> aba + aay = aas + aat \<Longrightarrow>
+       acm < aav \<or> acm = aav \<Longrightarrow>
+       (ada < aay \<or> 1 < aay + aba) \<or>
+       ada = aay \<and> 1 = aay + aba \<Longrightarrow>
+       abb + (abc + abd) < abe + abg \<or>
+       abb + (abc + abd) = abe + abg \<Longrightarrow>
+       (abh < abj \<or> abh < abj + abk) \<or>
+       abh = abj \<and> abh = abj + abk \<Longrightarrow>
+       1 < abo + abq \<or> 1 = abo + abq \<Longrightarrow>
+       (acj < abr \<or> acj + aci < abr + abs1) \<or>
+       acj = abr \<and> acj + aci = abr + abs1 \<Longrightarrow>
+       (abt < abv \<or> abt + abu < abv + abw) \<or>
+       abt = abv \<and> abt + abu = abv + abw \<Longrightarrow>
+       (abx < adc \<or> abx + aby < adc + abz) \<or>
+       abx = adc \<and> abx + aby = adc + abz \<Longrightarrow>
+       (acf < acd \<or>
+        acf < acd + acc \<or>
+        1 < acd + (acc + acb)) \<or>
+       acf = acd \<and>
+       acf = acd + acc \<and>
+       1 = acd + (acc + acb) \<Longrightarrow>
+       acc + acd < acf \<or> acc + acd = acf \<Longrightarrow>
+       (acg < acq \<or> acg + ach < acq + acr) \<or>
+       acg = acq \<and> acg + ach = acq + acr \<Longrightarrow>
+       aci + (acj + ack) < acr + acp \<or>
+       aci + (acj + ack) = acr + acp \<Longrightarrow>
+       (acm < acp \<or>
+        acm + acn < acp + acq \<or>
+        acm + (acn + aco)
+        < acp + (acq + acr)) \<or>
+       acm = acp \<and>
+       acm + acn = acp + acq \<and>
+       acm + (acn + aco) =
+       acp + (acq + acr) \<Longrightarrow>
+       (acs + act < acw + acw \<or>
+        acs + (act + acu)
+        < acw + (acw + acw)) \<or>
+       acs + act = acw + acw \<and>
+       acs + (act + acu) =
+       acw + (acw + acw) \<Longrightarrow>
+       (ada < adb + adr \<or>
+        1 < adb + (adr + adq)) \<or>
+       ada = adb + adr \<and>
+       1 = adb + (adr + adq) \<Longrightarrow>
+       (adc + adc < adf + adg \<or>
+        adc + (adc + abz)
+        < adf + (adg + adk)) \<or>
+       adc + adc = adf + adg \<and>
+       adc + (adc + abz) =
+       adf + (adg + adk) \<Longrightarrow>
+       adh + adi < adj + adk \<or>
+       adh + adi = adj + adk \<Longrightarrow>
+       (adl < aec \<or> 1 < aec + adp) \<or>
+       adl = aec \<and> 1 = aec + adp \<Longrightarrow>
+       (adq < ads \<or> adq + adr < ads) \<or>
+       adq = ads \<and> adq + adr = ads \<Longrightarrow>
+       adu + adv < aed + aea \<or>
+       adu + adv = aed + aea \<Longrightarrow>
+       (adw < aee \<or> adw + ady < aee + aea) \<or>
+       adw = aee \<and> adw + ady = aee + aea \<Longrightarrow>
+       (aeb < aed \<or> aeb + aec < aed + aee) \<or>
+       aeb = aed \<and> aeb + aec = aed + aee \<Longrightarrow>
+       False"
+       by argo
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_cc.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,654 @@
+(*  Title:      Tools/Argo/argo_cc.ML
+    Author:     Sascha Boehme
+    Author:     Dmitriy Traytel and Matthias Franze, TU Muenchen
+
+Equality reasoning based on congurence closure. It features:
+
+  * congruence closure for any term that participates in equalities
+  * support for predicates
+
+These features might be added:
+
+  * caching of explanations while building proofs to obtain shorter proofs
+    and faster proof checking
+  * propagating relevant merges of equivalence classes to all other theory solvers
+  * propagating new relevant negated equalities to all other theory solvers
+  * creating lemma "f ~= g | a ~= b | f a = g b" for asserted negated equalities
+    between "f a" and "g b" (dynamic ackermannization)
+
+The implementation is inspired by:
+
+  Robert Nieuwenhuis and Albert Oliveras. Fast Congruence Closure and
+  Extensions. In Information and Computation, volume 205(4),
+  pages 557-580, 2007.
+
+  Harald Ganzinger, George Hagen, Robert Nieuwenhuis, Albert Oliveras,
+  Cesare Tinelli. DPLL(T): Fast decision procedures. In Lecture Notes in
+  Computer Science, volume 3114, pages 175-188. Springer, 2004.
+*)
+
+signature ARGO_CC =
+sig
+  (* context *)
+  type context
+  val context: context
+
+  (* simplification *)
+  val simplify: Argo_Rewr.context -> Argo_Rewr.context
+  
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> Argo_Lit.literal option * context
+
+  (* main operations *)
+  val assume: Argo_Common.literal -> context -> Argo_Lit.literal Argo_Common.implied * context
+  val check: context -> Argo_Lit.literal Argo_Common.implied * context
+  val explain: Argo_Lit.literal -> context -> (Argo_Cls.clause * context) option
+  val add_level: context -> context
+  val backtrack: context -> context
+end
+
+structure Argo_Cc: ARGO_CC =
+struct
+
+(* tables indexed by pairs of terms *)
+
+val term2_ord = prod_ord Argo_Term.term_ord Argo_Term.term_ord
+
+structure Argo_Term2tab = Table(type key = Argo_Term.term * Argo_Term.term val ord = term2_ord)
+
+
+(* equality certificates *)
+
+(*
+  The solver keeps assumed equalities to produce explanations later on.
+
+  A flat equality (lp, (t1, t2)) consists of the assumed literal and its proof
+  as well as the terms t1 and t2 that are assumed to be equal. The literal expresses
+  the equality t1 = t2.
+
+  A congruence equality (t1, t2) is an equality t1 = t2 where both terms are
+  applications (f a) and (g b).
+
+  A symmetric equality eq is a marker for applying the symmetry rule to eq.
+*)
+
+datatype eq =
+  Flat of Argo_Common.literal * (Argo_Term.term * Argo_Term.term) |
+  Cong of Argo_Term.term * Argo_Term.term |
+  Symm of eq
+
+fun dest_eq (Flat (_, tp)) = tp
+  | dest_eq (Cong tp) = tp
+  | dest_eq (Symm eq) = swap (dest_eq eq)
+
+fun symm (Symm eq) = eq
+  | symm eq = Symm eq
+
+fun negate (Flat ((lit, p), tp)) = Flat ((Argo_Lit.negate lit, p), tp)
+  | negate (Cong tp) = Cong tp
+  | negate (Symm eq) = Symm (negate eq)
+
+fun dest_app (Argo_Term.T (_, Argo_Expr.App, [t1, t2])) = (t1, t2)
+  | dest_app _ = raise Fail "bad application"
+
+
+(* context *)
+
+(*
+  Each representative keeps track of the yet unimplied atoms in which this any member of
+  this representative's equivalence class occurs. An atom is either a list of equalities
+  between two terms, a list of predicates or a certificate. The certificate denotes that
+  this equivalence class contains already implied predicates, and the literal accompanying
+  the certificate specifies the polarity of these predicates.
+*)
+
+datatype atoms =
+  Eqs of (Argo_Term.term * Argo_Term.term) list |
+  Preds of Argo_Term.term list |
+  Cert of Argo_Common.literal
+
+(*
+  Each representative has an associated ritem that contains the members of the
+  equivalence class, the yet unimplied atoms and further information.
+*)
+
+type ritem = {
+  size: int, (* the size of the equivalence class *)
+  class: Argo_Term.term list, (* the equivalence class as a list of distinct terms *)
+  occs: Argo_Term.term list, (* a list of all application terms in which members of
+    the equivalence class occur either as function or as argument *)
+  neqs: (Argo_Term.term * eq) list, (* a list of terms from disjoint equivalence classes,
+    for each term of this list there is a certificate of a negated equality that is
+    required to explain why the equivalence classes are disjoint *)
+  atoms: atoms} (* the atoms of the representative *)
+
+type repr = Argo_Term.term Argo_Termtab.table
+type rdata = ritem Argo_Termtab.table
+type apps = Argo_Term.term Argo_Term2tab.table
+type trace = (Argo_Term.term * eq) Argo_Termtab.table
+
+type context = {
+  repr: repr, (* a table mapping terms to their representatives *)
+  rdata: rdata, (* a table mapping representatives to their ritems *)
+  apps: apps, (* a table mapping a function and an argument to their application *)
+  trace: trace, (* the proof forest used to trace assumed and implied equalities *)
+  prf: Argo_Proof.context, (* the proof context *)
+  back: (repr * rdata * apps * trace) list} (* backtracking information *)
+
+fun mk_context repr rdata apps trace prf back: context =
+  {repr=repr, rdata=rdata, apps=apps, trace=trace, prf=prf, back=back}
+
+val context =
+  mk_context Argo_Termtab.empty Argo_Termtab.empty Argo_Term2tab.empty Argo_Termtab.empty
+    Argo_Proof.cc_context []
+
+fun repr_of repr t = the_default t (Argo_Termtab.lookup repr t)
+fun repr_of' ({repr, ...}: context) = repr_of repr
+fun put_repr t r = Argo_Termtab.update (t, r)
+
+fun mk_ritem size class occs neqs atoms: ritem =
+  {size=size, class=class, occs=occs, neqs=neqs, atoms=atoms}
+
+fun as_ritem t = mk_ritem 1 [t] [] [] (Eqs [])
+fun as_pred_ritem t = mk_ritem 1 [t] [] [] (Preds [t])
+fun gen_ritem_of mk rdata r = the_default (mk r) (Argo_Termtab.lookup rdata r)
+fun ritem_of rdata = gen_ritem_of as_ritem rdata
+fun ritem_of_pred rdata = gen_ritem_of as_pred_ritem rdata
+fun ritem_of' ({rdata, ...}: context) = ritem_of rdata
+fun put_ritem r ri = Argo_Termtab.update (r, ri)
+
+fun add_occ r occ = Argo_Termtab.map_default (r, as_ritem r)
+  (fn {size, class, occs, neqs, atoms}: ritem => mk_ritem size class (occ :: occs) neqs atoms)
+
+fun put_atoms atoms ({size, class, occs, neqs, ...}: ritem) = mk_ritem size class occs neqs atoms
+
+fun add_eq_atom r atom = Argo_Termtab.map_default (r, as_ritem r)
+  (fn ri as {atoms=Eqs atoms, ...}: ritem => put_atoms (Eqs (atom :: atoms)) ri
+    | ri => put_atoms (Eqs [atom]) ri)
+
+fun lookup_app apps tp = Argo_Term2tab.lookup apps tp
+fun put_app tp app = Argo_Term2tab.update_new (tp, app)
+
+
+(* traces for explanations *)
+
+(*
+  Assumed and implied equalities are collected in a proof forest for being able to
+  produce explanations. For each equivalence class there is one proof tree. The
+  equality certificates are oriented towards a root term, that is not necessarily
+  the representative of the equivalence class.
+*)
+
+(*
+  Whenever two equivalence classes are merged due to an equality t1 = t2, the shorter
+  of the two paths, either from t1 to its root or t2 to its root, is re-oriented such
+  that the relevant ti becomes the new root of its tree. Then, a new edge between ti
+  and the other term of the equality t1 = t2 is added to connect the two proof trees.
+*)
+
+fun depth_of trace t =
+  (case Argo_Termtab.lookup trace t of
+    NONE => 0
+  | SOME (t', _) => 1 + depth_of trace t')
+
+fun reorient t trace =
+  (case Argo_Termtab.lookup trace t of
+    NONE => trace
+  | SOME (t', eq) => Argo_Termtab.update (t', (t, symm eq)) (reorient t' trace))
+
+fun new_edge from to eq trace = Argo_Termtab.update (from, (to, eq)) (reorient from trace)
+
+fun with_shortest f (t1, t2) eq trace =
+  (if depth_of trace t1 <= depth_of trace t2 then f t1 t2 eq else f t2 t1 (symm eq)) trace
+
+fun add_edge eq trace = with_shortest new_edge (dest_eq eq) eq trace
+
+(*
+  To produce an explanation that t1 and t2 are equal, the paths to their root are
+  extracted from the proof forest. Common ancestors in both paths are dropped.
+*)
+
+fun path_to_root trace path t =
+  (case Argo_Termtab.lookup trace t of
+    NONE => (t, path)
+  | SOME (t', _) => path_to_root trace (t :: path) t')
+
+fun drop_common root (t1 :: path1) (t2 :: path2) =
+      if Argo_Term.eq_term (t1, t2) then drop_common t1 path1 path2 else root
+  | drop_common root _ _ = root
+
+fun common_ancestor trace t1 t2 =
+  let val ((root, path1), (_, path2)) = apply2 (path_to_root trace []) (t1, t2)
+  in drop_common root path1 path2 end
+
+(*
+  The proof of an assumed literal is typically a hypothesis. If the assumed literal is
+  already known to be a unit literal, then there is already a proof for it.
+*)
+
+fun proof_of (lit, NONE) lits prf =
+      (insert Argo_Lit.eq_lit (Argo_Lit.negate lit) lits, Argo_Proof.mk_hyp lit prf)
+  | proof_of (_, SOME p) lits prf = (lits, (p, prf))
+
+(*
+  The explanation of equality between two terms t1 and t2 is computed based on the
+  paths from t1 and t2 to their common ancestor t in the proof forest. For each of
+  the two paths, a transitive proof of equality t1 = t and t = t2 is constructed,
+  such that t1 = t2 follows by transitivity.
+  
+  Each edge of the paths denotes an assumed or implied equality. Implied equalities
+  might be due to congruences (f a = g b) for which the equalities f = g and a = b
+  need to be explained recursively.
+*)
+
+fun mk_eq_proof trace t1 t2 lits prf =
+  if Argo_Term.eq_term (t1, t2) then (lits, Argo_Proof.mk_refl t1 prf)
+  else
+    let
+      val root = common_ancestor trace t1 t2
+      val (lits, (p1, prf)) = trans_proof I I trace t1 root lits prf
+      val (lits, (p2, prf)) = trans_proof swap symm trace t2 root lits prf
+    in (lits, Argo_Proof.mk_trans p1 p2 prf) end
+
+and trans_proof sw sy trace t root lits prf =
+  if Argo_Term.eq_term (t, root) then (lits, Argo_Proof.mk_refl t prf)
+  else
+    (case Argo_Termtab.lookup trace t of
+      NONE => raise Fail "bad trace"
+    | SOME (t', eq) => 
+        let
+          val (lits, (p1, prf)) = proof_step trace (sy eq) lits prf
+          val (lits, (p2, prf)) = trans_proof sw sy trace t' root lits prf
+        in (lits, uncurry Argo_Proof.mk_trans (sw (p1, p2)) prf) end)
+
+and proof_step _ (Flat (cert, _)) lits prf = proof_of cert lits prf
+  | proof_step trace (Cong tp) lits prf =
+      let
+        val ((t1, t2), (u1, u2)) = apply2 dest_app tp
+        val (lits, (p1, prf)) = mk_eq_proof trace t1 u1 lits prf
+        val (lits, (p2, prf)) = mk_eq_proof trace t2 u2 lits prf
+      in (lits, Argo_Proof.mk_cong p1 p2 prf) end
+  | proof_step trace (Symm eq) lits prf =
+      proof_step trace eq lits prf ||> uncurry Argo_Proof.mk_symm
+
+(*
+  All clauses produced by a theory solver are expected to be a lemma.
+  The lemma proof must hence be the last proof step.
+*)
+
+fun close_proof lit lits (p, prf) = (lit :: lits, Argo_Proof.mk_lemma [lit] p prf)
+
+(*
+  The explanation for the equality of t1 and t2 used the above algorithm.
+*)
+
+fun explain_eq lit t1 t2 ({repr, rdata, apps, trace, prf, back}: context) =
+  let val (lits, (p, prf)) = mk_eq_proof trace t1 t2 [] prf |-> close_proof lit
+  in ((lits, p), mk_context repr rdata apps trace prf back) end
+
+(*
+  The explanation that t1 and t2 are distinct uses the negated equality u1 ~= u2 that
+  explains why the equivalence class containing t1 and u1 and the equivalence class
+  containing t2 and u2 are disjoint. The explanations for t1 = u1 and u2 = t2 are
+  constructed using the above algorithm. By transitivity, it follows that t1 ~= t2.  
+*)
+
+fun finish_proof (Flat ((lit, _), _)) lits p prf = close_proof lit lits (p, prf)
+  | finish_proof (Cong _) _ _ _ = raise Fail "bad equality"
+  | finish_proof (Symm eq) lits p prf = Argo_Proof.mk_symm p prf |-> finish_proof eq lits
+
+fun explain_neq eq eq' ({repr, rdata, apps, trace, prf, back}: context) =
+  let
+    val (t1, t2) = dest_eq eq
+    val (u1, u2) = dest_eq eq'
+
+    val (lits, (p, prf)) = proof_step trace eq' [] prf
+    val (lits, (p1, prf)) = mk_eq_proof trace t1 u1 lits prf
+    val (lits, (p2, prf)) = mk_eq_proof trace u2 t2 lits prf
+    val (lits, (p, prf)) = 
+      Argo_Proof.mk_trans p p2 prf |-> Argo_Proof.mk_trans p1 |-> finish_proof eq lits
+  in ((lits, p), mk_context repr rdata apps trace prf back) end
+
+
+(* propagating new equalities *)
+
+exception CONFLICT of Argo_Cls.clause * context
+
+(*
+  comment missing
+*)
+
+fun same_repr repr r (t, _) = Argo_Term.eq_term (r, repr_of repr t)
+
+fun has_atom rdata r eq =
+  (case #atoms (ritem_of rdata r) of
+    Eqs eqs => member (Argo_Term.eq_term o snd) eqs eq
+  | _ => false)
+
+fun add_implied mk_lit repr rdata r neqs (atom as (t, eq)) (eqs, ls) =
+  let val r' = repr_of repr t
+  in
+    if Argo_Term.eq_term (r, r') then (eqs, insert Argo_Lit.eq_lit (mk_lit eq) ls)
+    else if exists (same_repr repr r') neqs andalso has_atom rdata r' eq then
+      (eqs, Argo_Lit.Neg eq :: ls)
+    else (atom :: eqs, ls)
+  end
+
+(*
+  comment missing
+*)
+
+fun copy_occ repr app (eqs, occs, apps) =
+  let val rp = apply2 (repr_of repr) (dest_app app)
+  in
+    (case lookup_app apps rp of
+      SOME app' => (Cong (app, app') :: eqs, occs, apps)
+    | NONE => (eqs, app :: occs, put_app rp app apps))
+  end
+
+(*
+  comment missing
+*)
+
+fun add_lits (Argo_Lit.Pos _, _) = fold (cons o Argo_Lit.Pos)
+  | add_lits (Argo_Lit.Neg _, _) = fold (cons o Argo_Lit.Neg)
+
+fun join_atoms f (Eqs eqs1) (Eqs eqs2) ls = f eqs1 eqs2 ls
+  | join_atoms _ (Preds ts1) (Preds ts2) ls = (Preds (union Argo_Term.eq_term ts1 ts2), ls)
+  | join_atoms _ (Preds ts) (Cert lp) ls = (Cert lp, add_lits lp ts ls)
+  | join_atoms _ (Cert lp) (Preds ts) ls = (Cert lp, add_lits lp ts ls)
+  | join_atoms _ (Cert lp) (Cert _) ls = (Cert lp, ls)
+  | join_atoms _ _ _ _ = raise Fail "bad atoms"
+
+(*
+  comment missing
+*)
+
+fun join r1 ri1 r2 ri2 eq (eqs, ls, {repr, rdata, apps, trace, prf, back}: context) =
+  let
+    val {size=size1, class=class1, occs=occs1, neqs=neqs1, atoms=atoms1}: ritem = ri1
+    val {size=size2, class=class2, occs=occs2, neqs=neqs2, atoms=atoms2}: ritem = ri2
+
+    val repr = fold (fn t => put_repr t r1) class2 repr
+    val class = fold cons class2 class1
+    val (eqs, occs, apps) = fold (copy_occ repr) occs2 (eqs, occs1, apps)
+    val trace = add_edge eq trace
+    val neqs = AList.merge Argo_Term.eq_term (K true) (neqs1, neqs2)
+    fun add r neqs = fold (add_implied Argo_Lit.Pos repr rdata r neqs)
+    fun adds eqs1 eqs2 ls = ([], ls) |> add r2 neqs2 eqs1 |> add r1 neqs1 eqs2 |>> Eqs
+    val (atoms, ls) = join_atoms adds atoms1 atoms2 ls
+    (* TODO: make sure that all implied literals are propagated *)
+    val rdata = put_ritem r1 (mk_ritem (size1 + size2) class occs neqs atoms) rdata
+  in (eqs, ls, mk_context repr rdata apps trace prf back) end
+
+(*
+  comment missing
+*)
+
+fun find_neq ({repr, ...}: context) ({neqs, ...}: ritem) r = find_first (same_repr repr r) neqs
+
+fun check_join (r1, r2) (ri1, ri2) eq (ecx as (_, _, cx)) =
+  (case find_neq cx ri2 r1 of
+    SOME (_, eq') => raise CONFLICT (explain_neq (negate (symm eq)) eq' cx)
+  | NONE =>
+      (case find_neq cx ri1 r2 of
+        SOME (_, eq') => raise CONFLICT (explain_neq (negate eq) eq' cx)
+      | NONE => join r1 ri1 r2 ri2 eq ecx))
+
+(*
+  comment missing
+*)
+
+fun with_max_class f (rp as (r1, r2)) (rip as (ri1: ritem, ri2: ritem)) eq =
+  if #size ri1 >= #size ri2 then f rp rip eq else f (r2, r1) (ri2, ri1) (symm eq)
+
+(*
+  comment missing
+*)
+
+fun propagate ([], ls, cx) = (rev ls, cx)
+  | propagate (eq :: eqs, ls, cx) =
+      let val rp = apply2 (repr_of' cx) (dest_eq eq)
+      in 
+        if Argo_Term.eq_term rp then propagate (eqs, ls, cx)
+        else propagate (with_max_class check_join rp (apply2 (ritem_of' cx) rp) eq (eqs, ls, cx))
+      end
+
+fun without lit (lits, cx) = (Argo_Common.Implied (remove Argo_Lit.eq_lit lit lits), cx)
+
+fun flat_merge (lp as (lit, _)) eq cx = without lit (propagate ([Flat (lp, eq)], [], cx))
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+(*
+  comment missing
+*)
+
+fun app_merge app tp (cx as {repr, rdata, apps, trace, prf, back}: context) =
+  let val rp as (r1, r2) = apply2 (repr_of repr) tp
+  in
+    (case lookup_app apps rp of
+      SOME app' =>
+        (case propagate ([Cong (app, app')], [], cx) of
+          ([], cx) => cx
+        | _ => raise Fail "bad application merge")
+    | NONE =>
+        let val rdata = add_occ r1 app (add_occ r2 app rdata)
+        in mk_context repr rdata (put_app rp app apps) trace prf back end)
+  end
+
+(*
+  A negated equality between t1 and t2 is only recorded if t1 and t2 are not already known
+  to belong to the same class. In that case, a conflict is raised with an explanation
+  why t1 and t2 are equal. Otherwise, the classes of t1 and t2 are marked as disjoint by
+  storing the negated equality in the ritems of t1's and t2's representative. All equalities
+  between terms of t1's and t2's class are implied as negated equalities. Those equalities
+  are found in the ritems of t1's and t2's representative.
+*)
+
+fun note_neq eq (r1, r2) (t1, t2) ({repr, rdata, apps, trace, prf, back}: context) =
+  let
+    val {size=size1, class=class1, occs=occs1, neqs=neqs1, atoms=atoms1}: ritem = ritem_of rdata r1
+    val {size=size2, class=class2, occs=occs2, neqs=neqs2, atoms=atoms2}: ritem = ritem_of rdata r2
+
+    fun add r (Eqs eqs) ls = fold (add_implied Argo_Lit.Neg repr rdata r []) eqs ([], ls) |>> Eqs
+      | add _ _ _ = raise Fail "bad negated equality between predicates"
+    val ((atoms1, atoms2), ls) = [] |> add r2 atoms1 ||>> add r1 atoms2
+    val ri1 = mk_ritem size1 class1 occs1 ((t2, eq) :: neqs1) atoms1
+    val ri2 = mk_ritem size2 class2 occs2 ((t1, symm eq) :: neqs2) atoms2
+  in (ls, mk_context repr (put_ritem r1 ri1 (put_ritem r2 ri2 rdata)) apps trace prf back) end
+
+fun flat_neq (lp as (lit, _)) (tp as (t1, t2)) cx =
+  let val rp = apply2 (repr_of' cx) tp
+  in
+    if Argo_Term.eq_term rp then
+      let val (cls, cx) = explain_eq (Argo_Lit.negate lit) t1 t2 cx
+      in (Argo_Common.Conflict cls, cx) end
+    else without lit (note_neq (Flat (lp, tp)) rp tp cx)
+  end
+
+
+(* simplification *)
+
+(*
+  Only equalities are subject to normalizations. An equality between two expressions e1 and e2
+  is normalized, if e1 is less than e2 based on the expression ordering. If e1 and e2 are
+  syntactically equal, the equality between these two expressions is normalized to the true
+  expression.
+*)
+
+fun norm_eq env =
+  let val e1 = Argo_Rewr.get env 1 and e2 = Argo_Rewr.get env 2
+  in
+    (case Argo_Expr.expr_ord (e1, e2) of
+      EQUAL => SOME (Argo_Proof.Rewr_Eq_Refl, Argo_Rewr.E Argo_Expr.true_expr)
+    | LESS => NONE
+    | GREATER => SOME (Argo_Proof.Rewr_Eq_Symm, Argo_Rewr.E (Argo_Expr.mk_eq e2 e1)))
+  end
+
+val simplify = Argo_Rewr.func "(eq (? 1) (? 2))" norm_eq
+
+
+(* declaring atoms *)
+
+(*
+  Only a genuinely new equality term t for the equality "t1 = t2" is added. If t1 and t2 belong
+  to the same equality class or if the classes of t1 and t2 are known to be disjoint, the
+  respective literal is returned together with an unmodified context.
+*)
+
+fun add_eq_term t t1 t2 (rp as (r1, r2)) (cx as {repr, rdata, apps, trace, prf, back}: context) =
+  if Argo_Term.eq_term rp then (SOME (Argo_Lit.Pos t), cx)
+  else if is_some (find_neq cx (ritem_of rdata r1) r2) then (SOME (Argo_Lit.Neg t), cx)
+  else
+    let val rdata = add_eq_atom r1 (t2, t) (add_eq_atom r2 (t1, t) rdata)
+    in (NONE, mk_context repr rdata apps trace prf back) end
+
+(*
+  Only a genuinely new predicate t, which is an application "t1 t2", is added.
+  If there is a predicate that is known to be congruent to the representatives of t1 and t2,
+  and that predicate or its negation has already been assummed, the respective literal of t
+  is returned together with an unmodified context.
+*)
+
+fun add_pred_term t rp (cx as {repr, rdata, apps, trace, prf, back}: context) =
+  (case lookup_app apps rp of
+    NONE => (NONE, mk_context repr (put_ritem t (as_pred_ritem t) rdata) apps trace prf back)
+  | SOME app =>
+      (case `(ritem_of_pred rdata) (repr_of repr app) of
+        ({atoms=Cert (Argo_Lit.Pos _, _), ...}: ritem, _) => (SOME (Argo_Lit.Pos t), cx)
+      | ({atoms=Cert (Argo_Lit.Neg _, _), ...}: ritem, _) => (SOME (Argo_Lit.Neg t), cx)
+      | (ri as {atoms=Preds ts, ...}: ritem, r) =>
+          let val rdata = put_ritem r (put_atoms (Preds (t :: ts)) ri) rdata
+          in (NONE, mk_context repr rdata apps trace prf back) end
+      | ({atoms=Eqs _, ...}: ritem, _) => raise Fail "bad predicate"))
+
+(*
+  For each term t that is an application "t1 t2", the reflexive equality t = t1 t2 is added
+  to the context. This is required for propagations of congruences.
+*)
+
+fun flatten (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])) cx =
+      flatten t1 (flatten t2 (app_merge t (t1, t2) cx))
+  | flatten _ cx = cx
+
+(*
+  Atoms to be added to the context must either be an equality "t1 = t2" or
+  an application "t1 t2" (a predicate). Besides adding the equality or the application,
+  reflexive equalities for for all applications in the terms t1 and t2 are added.
+*)
+
+fun add_atom (t as Argo_Term.T (_, Argo_Expr.Eq, [t1, t2])) cx =
+      add_eq_term t t1 t2 (apply2 (repr_of' cx) (t1, t2)) (flatten t1 (flatten t2 cx))
+  | add_atom (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])) cx =
+      let val cx = flatten t1 (flatten t2 (app_merge t (t1, t2) cx))
+      in add_pred_term t (apply2 (repr_of' cx) (t1, t2)) cx end
+  | add_atom _ cx = (NONE, cx)
+
+
+(* assuming external knowledge *)
+
+(*
+  Assuming a predicate r replaces all predicate atoms of r's ritem with the assumed certificate.
+  The predicate atoms are implied, either with positive or with negative polarity based on
+  the assumption.
+
+  There must not be a certificate for r since otherwise r would have been assumed before already.
+*)
+
+fun assume_pred lit mk_lit cert r ({repr, rdata, apps, trace, prf, back}: context) =
+  (case ritem_of_pred rdata r of
+    {size, class, occs, neqs, atoms=Preds ts}: ritem =>
+      let val rdata = put_ritem r (mk_ritem size class occs neqs cert) rdata
+      in without lit (map mk_lit ts, mk_context repr rdata apps trace prf back) end
+  | _ => raise Fail "bad predicate assumption")
+
+(*
+  Assumed equalities "t1 = t2" are treated as flat equalities between terms t1 and t2.
+  If t1 and t2 are applications, congruences are propagated as part of the merge between t1 and t2.
+  Negated equalities are handled likewise.
+
+  Assumed predicates do not trigger congruences. Only predicates of the same class are implied.
+*)
+
+fun assume (lp as (Argo_Lit.Pos (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2])), _)) cx =
+      flat_merge lp (t1, t2) cx
+  | assume (lp as (Argo_Lit.Neg (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2])), _)) cx =
+      flat_neq lp (t1, t2) cx
+  | assume (lp as (lit as Argo_Lit.Pos (t as Argo_Term.T (_, Argo_Expr.App, [_, _])), _)) cx =
+      assume_pred lit Argo_Lit.Pos (Cert lp) (repr_of' cx t) cx
+  | assume (lp as (lit as Argo_Lit.Neg (t as Argo_Term.T (_, Argo_Expr.App, [_, _])), _)) cx =
+      assume_pred lit Argo_Lit.Neg (Cert lp) (repr_of' cx t) cx
+  | assume _ cx = (Argo_Common.Implied [], cx)
+
+
+(* checking for consistency and pending implications *)
+
+(*
+  The internal model is always kept consistent. All implications are propagated as soon as
+  new information is assumed. Hence, there is nothing to be done here.
+*)
+
+fun check cx = (Argo_Common.Implied [], cx)
+
+
+(* explanations *)
+
+(*
+  The explanation for the predicate t, which is an application of t1 and t2, is constructed
+  from the explanation of the predicate application "u1 u2" as well as the equalities "u1 = t1"
+  and "u2 = t2" which both are constructed from the proof forest. The substitution rule is
+  the proof step that concludes "t1 t2" from "u1 u2" and the two equalities "u1 = t1"
+  and "u2 = t2".
+
+  The atoms part of the ritem of t's representative must be a certificate of an already
+  assumed predicate for otherwise there would be no explanation for t.
+*)
+
+fun explain_pred lit t t1 t2 ({repr, rdata, apps, trace, prf, back}: context) =
+  (case ritem_of_pred rdata (repr_of repr t) of
+    {atoms=Cert (cert as (lit', _)), ...}: ritem =>
+      let
+        val (u1, u2) = dest_app (Argo_Lit.term_of lit')
+        val (lits, (p, prf)) = proof_of cert [] prf
+        val (lits, (p1, prf)) = mk_eq_proof trace u1 t1 lits prf
+        val (lits, (p2, prf)) = mk_eq_proof trace u2 t2 lits prf
+        val (lits, (p, prf)) = Argo_Proof.mk_subst p p1 p2 prf |> close_proof lit lits
+      in ((lits, p), mk_context repr rdata apps trace prf back) end
+  | _ => raise Fail "no explanation for bad predicate")
+
+(*
+  Explanations are produced based on the proof forest that is constructed while assuming new
+  information and propagating this among the internal data structures.
+  
+  For predicates, no distinction between both polarities needs to be done here. The atoms
+  part of the relevant ritem knows the assumed polarity.
+*)
+
+fun explain (lit as Argo_Lit.Pos (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2]))) cx =
+      SOME (explain_eq lit t1 t2 cx)
+  | explain (lit as Argo_Lit.Neg (Argo_Term.T (_, Argo_Expr.Eq, [t1, t2]))) cx =
+      let val (_, eq) = the (find_neq cx (ritem_of' cx (repr_of' cx t1)) (repr_of' cx t2))
+      in SOME (explain_neq (Flat ((lit, NONE), (t1, t2))) eq cx) end
+  | explain (lit as (Argo_Lit.Pos (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])))) cx =
+      SOME (explain_pred lit t t1 t2 cx)
+  | explain (lit as (Argo_Lit.Neg (t as Argo_Term.T (_, Argo_Expr.App, [t1, t2])))) cx =
+      SOME (explain_pred lit t t1 t2 cx)
+  | explain _ _ = NONE
+
+
+(* backtracking *)
+
+(*
+  All information that needs to be reconstructed on backtracking is stored on the backtracking
+  stack. On backtracking any current information is replaced by what was stored before. No copying
+  nor subtle updates are required thanks to immutable data structures.
+*)
+
+fun add_level ({repr, rdata, apps, trace, prf, back}: context) =
+  mk_context repr rdata apps trace prf ((repr, rdata, apps, trace) :: back)
+
+fun backtrack ({back=[], ...}: context) = raise Empty
+  | backtrack ({prf, back=(repr, rdata, apps, trace) :: back, ...}: context) =
+      mk_context repr rdata apps trace prf back
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_cdcl.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,477 @@
+(*  Title:      Tools/Argo/argo_cdcl.ML
+    Author:     Sascha Boehme
+
+Propositional satisfiability solver in the style of conflict-driven
+clause-learning (CDCL). It features:
+
+ * conflict analysis and clause learning based on the first unique implication point
+ * nonchronological backtracking
+ * dynamic variable ordering (VSIDS)
+ * restarting
+ * polarity caching
+ * propagation via two watched literals
+ * special propagation of binary clauses 
+ * minimizing learned clauses
+ * support for external knowledge
+
+These features might be added:
+
+ * pruning of unnecessary learned clauses
+ * rebuilding the variable heap
+ * aligning the restart level with the decision heuristics: keep decisions that would
+   be recovered instead of backjumping to level 0
+
+The implementation is inspired by:
+
+  Niklas E'en and Niklas S"orensson. An Extensible SAT-solver. In Enrico
+  Giunchiglia and Armando Tacchella, editors, Theory and Applications of
+  Satisfiability Testing. Volume 2919 of Lecture Notes in Computer
+  Science, pages 502-518. Springer, 2003.
+
+  Niklas S"orensson and Armin Biere. Minimizing Learned Clauses. In
+  Oliver Kullmann, editor, Theory and Applications of Satisfiability
+  Testing. Volume 5584 of Lecture Notes in Computer Science,
+  pages 237-243. Springer, 2009.
+*)
+
+signature ARGO_CDCL =
+sig
+  (* types *)
+  type 'a explain = Argo_Lit.literal -> 'a -> Argo_Cls.clause * 'a
+
+  (* context *)
+  type context
+  val context: context
+  val assignment_of: context -> Argo_Lit.literal -> bool option
+
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> context
+  val add_axiom: Argo_Cls.clause -> context -> int * context
+
+  (* main operations *)
+  val assume: 'a explain -> Argo_Lit.literal -> context -> 'a ->
+    Argo_Cls.clause option * context * 'a
+  val propagate: context -> Argo_Common.literal Argo_Common.implied * context
+  val decide: context -> context option
+  val analyze: 'a explain -> Argo_Cls.clause -> context -> 'a -> int * context * 'a
+  val restart: context -> int * context
+end
+
+structure Argo_Cdcl: ARGO_CDCL =
+struct
+
+(* basic types and operations *)
+
+type 'a explain = Argo_Lit.literal -> 'a -> Argo_Cls.clause * 'a
+
+datatype reason =
+  Level0 of Argo_Proof.proof |
+  Decided of int * int * (bool * reason) Argo_Termtab.table |
+  Implied of int * int * (Argo_Lit.literal * reason) list * Argo_Proof.proof |
+  External of int
+
+fun level_of (Level0 _) = 0
+  | level_of (Decided (l, _, _)) = l
+  | level_of (Implied (l, _, _, _)) = l
+  | level_of (External l) = l
+
+type justified = Argo_Lit.literal * reason
+
+type watches = Argo_Cls.clause list * Argo_Cls.clause list
+
+fun get_watches wts t = Argo_Termtab.lookup wts t
+fun map_watches f t wts = Argo_Termtab.map_default (t, ([], [])) f wts
+
+fun map_lit_watches f (Argo_Lit.Pos t) = map_watches (apsnd f) t
+  | map_lit_watches f (Argo_Lit.Neg t) = map_watches (apfst f) t
+
+fun watches_of wts (Argo_Lit.Pos t) = (case get_watches wts t of SOME (ws, _) => ws | NONE => [])
+  | watches_of wts (Argo_Lit.Neg t) = (case get_watches wts t of SOME (_, ws) => ws | NONE => [])
+
+fun attach cls lit = map_lit_watches (cons cls) lit
+fun detach cls lit = map_lit_watches (remove Argo_Cls.eq_clause cls) lit
+
+
+(* literal values *)
+
+fun raw_val_of vals lit = Argo_Termtab.lookup vals (Argo_Lit.term_of lit)
+
+fun val_of vals (Argo_Lit.Pos t) = Argo_Termtab.lookup vals t
+  | val_of vals (Argo_Lit.Neg t) = Option.map (apfst not) (Argo_Termtab.lookup vals t)
+
+fun value_of vals (Argo_Lit.Pos t) = Option.map fst (Argo_Termtab.lookup vals t)
+  | value_of vals (Argo_Lit.Neg t) = Option.map (not o fst) (Argo_Termtab.lookup vals t)
+
+fun justified vals lit = Option.map (pair lit o snd) (raw_val_of vals lit)
+fun the_reason_of vals lit = snd (the (raw_val_of vals lit))
+
+fun assign (Argo_Lit.Pos t) r = Argo_Termtab.update (t, (true, r))
+  | assign (Argo_Lit.Neg t) r = Argo_Termtab.update (t, (false, r))
+
+
+(* context *)
+
+type trail = int * justified list (* the trail height and the sequence of assigned literals *)
+
+type context = {
+  units: Argo_Common.literal list, (* the literals that await propagation *)
+  level: int, (* the decision level *)
+  trail: int * justified list, (* the trail height and the sequence of assigned literals *)
+  vals: (bool * reason) Argo_Termtab.table, (* mapping of terms to polarity and reason *)
+  wts: watches Argo_Termtab.table, (* clauses watched by terms *)
+  heap: Argo_Heap.heap, (* max-priority heap for decision heuristics *)
+  clss: Argo_Cls.table, (* information about clauses *)
+  prf: Argo_Proof.context} (* the proof context *)
+
+fun mk_context units level trail vals wts heap clss prf: context =
+  {units=units, level=level, trail=trail, vals=vals, wts=wts, heap=heap, clss=clss, prf=prf}
+
+val context =
+  mk_context [] 0 (0, []) Argo_Termtab.empty Argo_Termtab.empty Argo_Heap.heap
+    Argo_Cls.table Argo_Proof.cdcl_context
+
+fun drop_levels n (Decided (l, h, vals)) trail heap =
+      if l = n + 1 then ((h, trail), vals, heap) else drop_literal n trail heap
+  | drop_levels n _ tr heap = drop_literal n tr heap
+
+and drop_literal n ((lit, r) :: trail) heap = drop_levels n r trail (Argo_Heap.insert lit heap)
+  | drop_literal _ [] _ = raise Fail "bad trail"
+
+fun backjump_to new_level (cx as {level, trail=(_, tr), wts, heap, clss, prf, ...}: context) =
+  if new_level >= level then (0, cx)
+  else
+    let val (trail, vals, heap) = drop_literal (Integer.max 0 new_level) tr heap
+    in (level - new_level, mk_context [] new_level trail vals wts heap clss prf) end
+
+
+(* proofs *)
+
+fun tag_clause (lits, p) prf = Argo_Proof.mk_clause lits p prf |>> pair lits
+
+fun level0_unit_proof (lit, Level0 p') (p, prf) = Argo_Proof.mk_unit_res lit p p' prf
+  | level0_unit_proof _ _ = raise Fail "bad reason"
+
+fun level0_unit_proofs lrs p prf = fold level0_unit_proof lrs (p, prf)
+
+fun unsat ({vals, prf, ...}: context) (lits, p) =
+  let val lrs = map (fn lit => (lit, the_reason_of vals lit)) lits
+  in Argo_Proof.unsat (fst (level0_unit_proofs lrs p prf)) end
+
+
+(* literal operations *)
+
+fun push lit p reason prf ({units, level, trail=(h, tr), vals, wts, heap, clss, ...}: context) =
+  let val vals = assign lit reason vals
+  in mk_context ((lit, p) :: units) level (h + 1, (lit, reason) :: tr) vals wts heap clss prf end
+
+fun push_level0 lit p lrs (cx as {prf, ...}: context) =
+  let val (p, prf) = level0_unit_proofs lrs p prf
+  in push lit (SOME p) (Level0 p) prf cx end
+
+fun push_implied lit p lrs (cx as {level, trail=(h, _), prf, ...}: context) =
+  if level > 0 then push lit NONE (Implied (level, h, lrs, p)) prf cx
+  else push_level0 lit p lrs cx
+
+fun push_decided lit (cx as {level, trail=(h, _), vals, prf, ...}: context) =
+  push lit NONE (Decided (level, h, vals)) prf cx
+
+fun assignment_of ({vals, ...}: context) = value_of vals
+
+fun replace_watches old new cls ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  mk_context units level trail vals (attach cls new (detach cls old wts)) heap clss prf
+
+
+(* clause operations *)
+
+fun as_clause cls ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  let val (cls, prf) = tag_clause cls prf
+  in (cls, mk_context units level trail vals wts heap clss prf) end
+
+fun note_watches ([_, _], _) _ clss = clss
+  | note_watches cls lp clss = Argo_Cls.put_watches cls lp clss
+
+fun attach_clause lit1 lit2 (cls as (lits, _)) cx =
+  let
+    val {units, level, trail, vals, wts, heap, clss, prf}: context = cx
+    val wts = attach cls lit1 (attach cls lit2 wts)
+    val clss = note_watches cls (lit1, lit2) clss
+  in mk_context units level trail vals wts (fold Argo_Heap.count lits heap) clss prf end
+
+fun change_watches _ (false, _, _) cx = cx
+  | change_watches cls (true, l1, l2) ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+      mk_context units level trail vals wts heap (Argo_Cls.put_watches cls (l1, l2) clss) prf
+
+fun add_asserting lit lit' (cls as (_, p)) lrs cx =
+  attach_clause lit lit' cls (push_implied lit p lrs cx)
+
+(*
+  When learning a non-unit clause, the context is backtracked to the highest decision level
+  of the assigned literals.
+*)
+
+fun learn_clause _ ([lit], p) cx = backjump_to 0 cx ||> push_level0 lit p []
+  | learn_clause lrs (cls as (lits, _)) cx =
+      let
+        fun max_level (l, r) (ll as (_, lvl)) = if level_of r > lvl then (l, level_of r) else ll
+        val (lit, lvl) = fold max_level lrs (hd lits, 0)
+      in backjump_to lvl cx ||> add_asserting (hd lits) lit cls lrs end
+
+(*
+  An axiom with one unassigned literal and all remaining literals being assigned to
+  false is asserting. An axiom with all literals assigned to false on level 0 makes the
+  context unsatisfiable. An axiom with all literals assigned to false on higher levels
+  causes backjumping before the highest level, and then the axiom might be asserting if
+  only one literal is unassigned on that level.
+*)
+
+fun min lit i NONE = SOME (lit, i)
+  | min lit i (SOME (lj as (_, j))) = SOME (if i < j then (lit, i) else lj)
+
+fun level_ord ((_, r1), (_, r2)) = int_ord (level_of r2, level_of r1)
+fun add_max lr lrs = Ord_List.insert level_ord lr lrs
+
+fun part [] [] t us fs = (t, us, fs)
+  | part (NONE :: vs) (l :: ls) t us fs = part vs ls t (l :: us) fs
+  | part (SOME (true, r) :: vs) (l :: ls) t us fs = part vs ls (min l (level_of r) t) us fs
+  | part (SOME (false, r) :: vs) (l :: ls) t us fs = part vs ls t us (add_max (l, r) fs)
+  | part _ _ _ _ _ = raise Fail "mismatch between values and literals"
+
+fun backjump_add (lit, r) (lit', r') cls lrs cx =
+  let
+    val add =
+      if level_of r = level_of r' then attach_clause lit lit' cls
+      else add_asserting lit lit' cls lrs
+  in backjump_to (level_of r - 1) cx ||> add end
+
+fun analyze_axiom vs (cls as (lits, p), cx) =
+  (case part vs lits NONE [] [] of
+    (SOME (lit, lvl), [], []) =>
+      if lvl > 0 then backjump_to 0 cx ||> push_implied lit p [] else (0, cx)
+  | (SOME (lit, lvl), [], (lit', _) :: _) => (0, cx |> (lvl > 0) ? attach_clause lit lit' cls)
+  | (SOME (lit, lvl), lit' :: _, _) => (0, cx |> (lvl > 0) ? attach_clause lit lit' cls)
+  | (NONE, [], (_, Level0 _) :: _) => unsat cx cls
+  | (NONE, [], [(lit, _)]) => backjump_to 0 cx ||> push_implied lit p []
+  | (NONE, [], lrs as (lr :: lr' :: _)) => backjump_add lr lr' cls lrs cx
+  | (NONE, [lit], []) => backjump_to 0 cx ||> push_implied lit p []
+  | (NONE, [lit], lrs as (lit', _) :: _) => (0, add_asserting lit lit' cls lrs cx)
+  | (NONE, lit1 :: lit2 :: _, _) => (0, attach_clause lit1 lit2 cls cx)
+  | _ => raise Fail "bad clause")
+
+
+(* enriching the context *)
+
+fun add_atom t ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  let val heap = Argo_Heap.insert (Argo_Lit.Pos t) heap
+  in mk_context units level trail vals wts heap clss prf end
+
+fun add_axiom ([], p) _ = Argo_Proof.unsat p
+  | add_axiom (cls as (lits, _)) (cx as {vals, ...}: context) =
+      if has_duplicates Argo_Lit.eq_lit lits then raise Fail "clause with duplicate literals"
+      else if has_duplicates Argo_Lit.dual_lit lits then (0, cx)
+      else analyze_axiom (map (val_of vals) lits) (as_clause cls cx)
+
+
+(* external knowledge *)
+
+fun assume explain lit (cx as {level, vals, prf, ...}: context) x =
+  (case value_of vals lit of
+    SOME true => (NONE, cx, x)
+  | SOME false => 
+      let val (cls, x) = explain lit x
+      in if level = 0 then unsat cx cls else (SOME cls, cx, x) end
+  | NONE =>
+      if level = 0 then
+        let val ((lits, p), x) = explain lit x
+        in (NONE, push_level0 lit p (map_filter (justified vals) lits) cx, x) end
+      else (NONE, push lit NONE (External level) prf cx, x))
+
+
+(* propagation *)
+
+exception CONFLICT of Argo_Cls.clause * context
+
+fun order_lits_by lit (l1, l2) =
+  if Argo_Lit.eq_id (l1, lit) then (true, l2, l1) else (false, l1, l2)
+
+fun prop_binary (_, implied_lit, other_lit) (cls as (_, p)) (cx as {level, vals, ...}: context) =
+  (case value_of vals implied_lit of
+    NONE => push_implied implied_lit p [(other_lit, the_reason_of vals other_lit)] cx
+  | SOME true => cx
+  | SOME false => if level = 0 then unsat cx cls else raise CONFLICT (cls, cx))
+
+datatype next = Lit of Argo_Lit.literal | None of justified list
+
+fun with_non_false f l (SOME (false, r)) lrs = f ((l, r) :: lrs)
+  | with_non_false _ l _ _ = Lit l
+
+fun first_non_false _ _ [] lrs = None lrs
+  | first_non_false vals lit (l :: ls) lrs =
+      if Argo_Lit.eq_lit (l, lit) then first_non_false vals lit ls lrs
+      else with_non_false (first_non_false vals lit ls) l (val_of vals l) lrs
+
+fun prop_nary (lp as (_, lit1, lit2)) (cls as (lits, p)) (cx as {level, vals, ...}: context) =
+  let val v = value_of vals lit1
+  in
+    if v = SOME true then change_watches cls lp cx
+    else
+      (case first_non_false vals lit1 lits [] of
+        Lit lit2' => change_watches cls (true, lit1, lit2') (replace_watches lit2 lit2' cls cx)
+      | None lrs =>
+          if v = NONE then push_implied lit1 p lrs (change_watches cls lp cx)
+          else if level = 0 then unsat cx cls
+          else raise CONFLICT (cls, change_watches cls lp cx))
+  end
+
+fun prop_cls lit (cls as ([l1, l2], _)) cx = prop_binary (order_lits_by lit (l1, l2)) cls cx
+  | prop_cls lit cls (cx as {clss, ...}: context) =
+      prop_nary (order_lits_by lit (Argo_Cls.get_watches clss cls)) cls cx
+
+fun prop_lit (lp as (lit, _)) (lps, cx as {wts, ...}: context) =
+  (lp :: lps, fold (prop_cls lit) (watches_of wts lit) cx)
+
+fun prop lps (cx as {units=[], ...}: context) = (Argo_Common.Implied (rev lps), cx)
+  | prop lps ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+      fold_rev prop_lit units (lps, mk_context [] level trail vals wts heap clss prf) |-> prop
+
+fun propagate cx = prop [] cx
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+
+(* decisions *)
+
+(*
+  Decisions are based on an activity heuristics. The most active variable that is
+  still unassigned is chosen.
+*)
+
+fun decide ({units, level, trail, vals, wts, heap, clss, prf}: context) =
+  let
+    fun check NONE = NONE
+      | check (SOME (lit, heap)) =
+          if Argo_Termtab.defined vals (Argo_Lit.term_of lit) then check (Argo_Heap.extract heap)
+          else SOME (push_decided lit (mk_context units (level + 1) trail vals wts heap clss prf))
+  in check (Argo_Heap.extract heap) end
+
+
+(* conflict analysis and clause learning *)
+
+(*
+  Learned clauses often contain literals that are redundant, because they are
+  subsumed by other literals of the clause. By analyzing the implication graph beyond
+  the unique implication point, such redundant literals can be identified and hence
+  removed from the learned clause. Only literals occurring in the learned clause and
+  their reasons need to be analyzed.
+*)
+
+exception ESSENTIAL of unit
+
+fun history_ord ((h1, lit1, _), (h2, lit2, _)) =
+  if h1 < 0 andalso h2 < 0 then int_ord (apply2 Argo_Lit.signed_id_of (lit1, lit2))
+  else int_ord (h2, h1)
+
+fun rec_redundant stop (lit, Implied (lvl, h, lrs, p)) lps =
+      if stop lit lvl then lps
+      else fold (rec_redundant stop) lrs ((h, lit, p) :: lps)
+  | rec_redundant stop (lit, Decided (lvl, _, _)) lps =
+      if stop lit lvl then lps
+      else raise ESSENTIAL ()
+  | rec_redundant _ (lit, Level0 p) lps = ((~1, lit, p) :: lps)
+  | rec_redundant _ _ _ = raise ESSENTIAL ()
+
+fun redundant stop (lr as (lit, Implied (_, h, lrs, p))) (lps, essential_lrs) = (
+      (fold (rec_redundant stop) lrs ((h, lit, p) :: lps), essential_lrs)
+      handle ESSENTIAL () => (lps, lr :: essential_lrs))
+  | redundant _ lr (lps, essential_lrs) = (lps, lr :: essential_lrs)
+
+fun resolve_step (_, l, p') (p, prf) = Argo_Proof.mk_unit_res l p p' prf
+
+fun reduce lrs p prf =
+  let
+    val lits = map fst lrs
+    val levels = fold (insert (op =) o level_of o snd) lrs []
+    fun stop lit level =
+      if member Argo_Lit.eq_lit lits lit then true
+      else if member (op =) levels level then false
+      else raise ESSENTIAL ()
+
+    val (lps, lrs) = fold (redundant stop) lrs ([], [])
+  in (lrs, fold resolve_step (sort_distinct history_ord lps) (p, prf)) end
+
+(*
+  Literals that are candidates for the learned lemma are marked and unmarked while
+  traversing backwards through the trail. The last remaining marked literal is the first
+  unique implication point.
+*)
+
+fun unmark lit ms = remove Argo_Lit.eq_id lit ms
+fun marked ms lit = member Argo_Lit.eq_id ms lit
+
+(*
+  Whenever an implication is recorded, the reason for the false literals of the
+  asserting clause are known. It is reasonable to store this justification list as part
+  of the implication reason. Consequently, the implementation of conflict analysis can
+  benefit from this information, which does not need to be re-computed here.
+*)
+
+fun justification_for _ _ _ (Implied (_, _, lrs, p)) x = (lrs, p, x)
+  | justification_for explain vals lit (External _) x =
+      let val ((lits, p), x) = explain lit x
+      in (map_filter (justified vals) lits, p, x) end
+  | justification_for _ _ _ _ _ = raise Fail "bad reason"
+
+fun first_lit pred ((lr as (lit, _)) :: lrs) = if pred lit then (lr, lrs) else first_lit pred lrs
+  | first_lit _ _ = raise Empty
+
+(*
+  Beginning from the conflicting clause, the implication graph is traversed to the first
+  unique implication point. This breadth-first search is controlled by the topological order of
+  the trail, which is traversed backwards. While traversing through the trail, the conflict
+  literals of lower levels are collected to form the conflict lemma together with the unique
+  implication point. Conflict literals assigned on level 0 are excluded from the conflict lemma.
+  Conflict literals assigned on the current level are candidates for the first unique
+  implication point.
+*)
+
+fun analyze explain cls (cx as {level, trail, vals, wts, heap, clss, prf, ...}: context) x =
+  let
+    fun from_clause [] trail ms lrs h p prf x =
+          from_trail (first_lit (marked ms) trail) ms lrs h p prf x
+      | from_clause ((lit, r) :: clause_lrs) trail ms lrs h p prf x =
+          from_reason r lit clause_lrs trail ms lrs h p prf x
+ 
+    and from_reason (Level0 p') lit clause_lrs trail ms lrs h p prf x =
+          let val (p, prf) = Argo_Proof.mk_unit_res lit p p' prf
+          in from_clause clause_lrs trail ms lrs h p prf x end
+      | from_reason r lit clause_lrs trail ms lrs h p prf x =
+          if level_of r = level then
+            if marked ms lit then from_clause clause_lrs trail ms lrs h p prf x
+            else from_clause clause_lrs trail (lit :: ms) lrs (Argo_Heap.increase lit h) p prf x
+          else
+            let
+              val (lrs, h) =
+                if AList.defined Argo_Lit.eq_id lrs lit then (lrs, h)
+                else ((lit, r) :: lrs, Argo_Heap.increase lit h)
+            in from_clause clause_lrs trail ms lrs h p prf x end
+
+    and from_trail ((lit, _), _) [_] lrs h p prf x =
+          let val (lrs, (p, prf)) = reduce lrs p prf
+          in (Argo_Lit.negate lit :: map fst lrs, lrs, h, p, prf, x) end
+      | from_trail ((lit, r), trail) ms lrs h p prf x =
+          let
+            val (clause_lrs, p', x) = justification_for explain vals lit r x
+            val (p, prf) = Argo_Proof.mk_unit_res lit p' p prf
+          in from_clause clause_lrs trail (unmark lit ms) lrs h p prf x end
+
+    val (ls, p) = cls
+    val lrs = if level = 0 then unsat cx cls else map (fn l => (l, the_reason_of vals l)) ls
+    val (lits, lrs, heap, p, prf, x) = from_clause lrs (snd trail) [] [] heap p prf x
+    val heap = Argo_Heap.decay heap
+    val (levels, cx) = learn_clause lrs (lits, p) (mk_context [] level trail vals wts heap clss prf)
+  in (levels, cx, x) end
+
+
+(* restarting *)
+
+fun restart cx = backjump_to 0 cx
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_clausify.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,168 @@
+(*  Title:      Tools/Argo/argo_clausify.ML
+    Author:     Sascha Boehme
+
+Conversion of propositional formulas to definitional CNF.
+
+The clausification implementation is based on:
+
+  G. S. Tseitin. On the complexity of derivation in propositional
+  calculus.  In A. O. Slisenko (editor) Structures in Constructive
+  Mathematics and Mathematical Logic, Part II, Seminars in Mathematics,
+  pages 115-125. Steklov Mathematic Institute, 1968.
+
+  D. A. Plaisted and S. Greenbaum. A Structure-Preserving Clause Form
+  Translation. Journal of Symbolic Computation, 1986.
+
+  L. de Moura and N. Bj\orner. Proofs and Refutations, and Z3. In
+  P. Rudnicki and G. Sutcliffe and B. Konev and R. A. Schmidt and
+  S. Schulz (editors) International Workshop on the Implementation of
+  Logics. CEUR Workshop Proceedings, 2008.
+*)
+
+signature ARGO_CLAUSIFY =
+sig
+  val clausify: Argo_Rewr.context -> Argo_Expr.expr * Argo_Proof.proof ->
+    Argo_Proof.context * Argo_Core.context -> Argo_Proof.context * Argo_Core.context
+end
+
+structure Argo_Clausify: ARGO_CLAUSIFY =
+struct
+
+(* lifting of if-then-else expressions *)
+
+(*
+  It is assumed that expressions are free of if-then-else expressions whose then- and else-branch
+  have boolean type. Such if-then-else expressions can be rewritten to expressions using only
+  negation, conjunction and disjunction.
+
+  All other modules treat if-then-else expressions as constant expressions. They do not analyze or
+  decend into sub-expressions of an if-then-else expression.
+
+  Lifting an if-then-else expression (ite P t u) introduces two new clauses
+    (or (not P) (= (ite P t u) t)) and
+    (or P (= (ite P t u) u))
+*)
+
+fun ite_clause simp k es (eps, (prf, core)) =
+  let
+    val e = Argo_Expr.mk_or es
+    val (p, prf) = Argo_Proof.mk_taut k e prf 
+    val (ep, prf) = Argo_Rewr.with_proof (Argo_Rewr.args (Argo_Rewr.rewrite_top simp)) (e, p) prf
+  in (ep :: eps, (prf, core)) end
+
+fun check_ite simp t (e as Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) (eps, (prf, core)) =
+      (case Argo_Core.identify (Argo_Term.Term t) core of
+        (Argo_Term.Known _, core) => (eps, (prf, core))
+      | (Argo_Term.New _, core) =>
+          (eps, (prf, core))
+          |> ite_clause simp Argo_Proof.Taut_Ite_Then [Argo_Expr.mk_not e1, Argo_Expr.mk_eq e e2]
+          |> ite_clause simp Argo_Proof.Taut_Ite_Else [e1, Argo_Expr.mk_eq e e3])
+  | check_ite _ _ _ cx = cx
+
+fun lift_ites simp (t as Argo_Term.T (_, _, ts)) =
+  check_ite simp t (Argo_Term.expr_of t) #>
+  fold (lift_ites simp) ts
+
+
+(* tagged expressions and terms *)
+
+fun pos x = (true, x)
+fun neg x = (false, x)
+
+fun mk_lit true t = Argo_Lit.Pos t
+  | mk_lit false t = Argo_Lit.Neg t
+
+fun expr_of (true, t) = Argo_Term.expr_of t
+  | expr_of (false, t) = Argo_Expr.mk_not (Argo_Term.expr_of t)
+
+
+(* adding literals *)
+
+fun lit_for (polarity, x) (new_atoms, core) =
+  (case Argo_Core.add_atom x core of
+    (Argo_Term.Known t, core) => (mk_lit polarity t, (new_atoms, core))
+  | (Argo_Term.New t, core) => (mk_lit polarity t, (t :: new_atoms, core)))
+
+fun lit_of (Argo_Expr.E (Argo_Expr.Not, [e])) = lit_for (neg (Argo_Term.Expr e))
+  | lit_of e = lit_for (pos (Argo_Term.Expr e))
+
+fun lit_of' (pol, Argo_Term.T (_, Argo_Expr.Not, [t])) = lit_for (not pol, Argo_Term.Term t)
+  | lit_of' (pol, t) = lit_for (pol, Argo_Term.Term t)
+
+
+(* adding clauses *)
+
+fun add_clause f xs p (new_atoms, (prf, core)) =
+  let val (lits, (new_atoms, core)) = fold_map f xs (new_atoms, core)
+  in (new_atoms, (prf, Argo_Core.add_axiom (lits, p) core)) end
+
+fun simp_lit (e as Argo_Expr.E (Argo_Expr.Not, [Argo_Expr.E (Argo_Expr.Not, [e'])])) =
+      Argo_Rewr.rewr Argo_Proof.Rewr_Not_Not e' e
+  | simp_lit e = Argo_Rewr.keep e
+
+fun simp_clause (e as Argo_Expr.E (Argo_Expr.Or, _)) = Argo_Rewr.args simp_lit e
+  | simp_clause e = Argo_Rewr.keep e
+
+fun new_clause k ls (new_atoms, (prf, core)) =
+  let
+    val e = Argo_Expr.mk_or (map expr_of ls)
+    val (p, prf) = Argo_Proof.mk_taut k e prf
+    val ((_, p), prf) = Argo_Rewr.with_proof simp_clause (e, p) prf
+  in add_clause lit_of' ls p (new_atoms, (prf, core)) end
+
+
+(* clausifying propositions *)
+
+fun clausify_and t ts cx =
+  let
+    val n = length ts
+    val k1 = Argo_Proof.Taut_And_1 n and k2 = Argo_Proof.Taut_And_2 o rpair n
+  in
+    cx
+    |> new_clause k1 (pos t :: map neg ts)
+    |> fold_index (fn (i, t') => new_clause (k2 i) [neg t, pos t']) ts
+  end
+
+fun clausify_or t ts cx =
+  let
+    val n = length ts
+    val k1 = Argo_Proof.Taut_Or_1 o rpair n and k2 = Argo_Proof.Taut_Or_2 n
+  in
+    cx
+    |> fold_index (fn (i, t') => new_clause (k1 i) [pos t, neg t']) ts
+    |> new_clause k2 (neg t :: map pos ts)
+  end
+
+fun clausify_iff t t1 t2 cx =
+  cx
+  |> new_clause Argo_Proof.Taut_Iff_1 [pos t, pos t1, pos t2]
+  |> new_clause Argo_Proof.Taut_Iff_2 [pos t, neg t1, neg t2]
+  |> new_clause Argo_Proof.Taut_Iff_3 [neg t, neg t1, pos t2]
+  |> new_clause Argo_Proof.Taut_Iff_4 [neg t, pos t1, neg t2]
+
+fun clausify_lit (t as Argo_Term.T (_, Argo_Expr.And, ts)) = clausify_and t ts
+  | clausify_lit (t as Argo_Term.T (_, Argo_Expr.Or, ts)) = clausify_or t ts
+  | clausify_lit (t as Argo_Term.T (_, Argo_Expr.Iff, [t1, t2])) = clausify_iff t t1 t2
+  | clausify_lit _ = I
+
+fun exhaust_new_atoms ([], cx) = cx
+  | exhaust_new_atoms (t :: new_atoms, cx) = exhaust_new_atoms (clausify_lit t (new_atoms, cx))
+
+fun clausify_expr _ (Argo_Expr.E (Argo_Expr.True, _), _) cx = cx
+  | clausify_expr _ (Argo_Expr.E (Argo_Expr.False, _), p) _ = Argo_Proof.unsat p
+  | clausify_expr f (Argo_Expr.E (Argo_Expr.And, es), p) cx =
+      fold_index (clausify_conj f (length es) p) es cx
+  | clausify_expr f (Argo_Expr.E (Argo_Expr.Or, es), p) cx = add_clausify f es p cx
+  | clausify_expr f (e, p) cx = add_clausify f [e] p cx
+
+and clausify_conj f n p (i, e) (prf, core) =
+  let val (p, prf) = Argo_Proof.mk_conj i n p prf
+  in clausify_expr f (e, p) (prf, core) end
+
+and add_clausify f es p cx =
+  let val ecx as (new_atoms, _) = add_clause lit_of es p ([], cx)
+  in fold f new_atoms ([], exhaust_new_atoms ecx) |-> fold (clausify_expr (K I)) end
+
+fun clausify simp ep cx = clausify_expr (lift_ites simp) ep cx
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_cls.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,45 @@
+(*  Title:      Tools/Argo/argo_cls.ML
+    Author:     Sascha Boehme
+
+Representation of clauses. Clauses are disjunctions of literals with a proof that explains
+why the disjunction holds.
+*)
+
+signature ARGO_CLS =
+sig
+  type clause = Argo_Lit.literal list * Argo_Proof.proof
+  val eq_clause: clause * clause -> bool
+
+  (* two-literal watches for clauses *)
+  type table
+  val table: table
+  val put_watches: clause -> Argo_Lit.literal * Argo_Lit.literal -> table -> table
+  val get_watches: table -> clause -> Argo_Lit.literal * Argo_Lit.literal
+end
+
+structure Argo_Cls: ARGO_CLS =
+struct
+
+type clause = Argo_Lit.literal list * Argo_Proof.proof
+
+fun eq_clause ((_, p1), (_, p2)) = Argo_Proof.eq_proof_id (apply2 Argo_Proof.id_of (p1, p2))
+fun clause_ord ((_, p1), (_, p2)) = Argo_Proof.proof_id_ord (apply2 Argo_Proof.id_of (p1, p2))
+
+
+(* two-literal watches for clauses *)
+
+(*
+  The CDCL solver keeps a mapping of some literals to clauses. Exactly two literals
+  of a clause are used to index the clause.
+*)
+
+structure Clstab = Table(type key = clause val ord = clause_ord)
+
+type table = (Argo_Lit.literal * Argo_Lit.literal) Clstab.table
+
+val table: table = Clstab.empty
+
+fun put_watches cls lp table = Clstab.update (cls, lp) table
+fun get_watches table cls = the (Clstab.lookup table cls)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_common.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,25 @@
+(*  Title:      Tools/Argo/argo_common.ML
+    Author:     Sascha Boehme
+
+Common infrastructure for the decision procedures of Argo.
+*)
+
+signature ARGO_COMMON =
+sig
+  type literal = Argo_Lit.literal * Argo_Proof.proof option
+  datatype 'a implied = Implied of 'a list | Conflict of Argo_Cls.clause
+end
+
+structure Argo_Common: ARGO_COMMON =
+struct
+
+type literal = Argo_Lit.literal * Argo_Proof.proof option
+  (* Implied new knowledge accompanied with an optional proof. If there is no proof,
+     the literal is to be treated hypothetically. If there is a proof, the literal is
+     to be treated as uni clause. *)
+
+datatype 'a implied = Implied of 'a list | Conflict of Argo_Cls.clause
+  (* A result of a step of a decision procedure, either an implication of new knowledge
+     or clause whose literals are known to be false. *)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_core.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,159 @@
+(*  Title:      Tools/Argo/argo_core.ML
+    Author:     Sascha Boehme
+
+Core of the Argo theorem prover implementing the DPLL(T) loop.
+
+The implementation is based on:
+
+  Harald Ganzinger, George Hagen, Robert Nieuwenhuis, Albert Oliveras,
+  Cesare Tinelli. DPLL(T): Fast decision procedures. In Lecture Notes in
+  Computer Science, volume 3114, pages 175-188. Springer, 2004.
+
+  Robert Nieuwenhuis, Albert Oliveras, Cesare Tinelli. Solving SAT and
+  SAT modulo theories: From an abstract Davis-Putnam-Logemann-Loveland
+  procedure to DPLL(T). In Journal of the ACM, volume 53(6), pages
+  937-977.  ACM, 2006.
+*)
+
+signature ARGO_CORE =
+sig
+  (* context *)
+  type context
+  val context: context
+
+  (* enriching the context *)
+  val identify: Argo_Term.item -> context -> Argo_Term.identified * context
+  val add_atom: Argo_Term.item -> context -> Argo_Term.identified * context
+  val add_axiom: Argo_Cls.clause -> context -> context
+
+  (* DPLL(T) loop *)
+  val run: context -> context (* raises Argo_Proof.UNSAT *)
+
+  (* model *)
+  val model_of: context -> string * Argo_Expr.typ -> bool option
+end
+
+structure Argo_Core: ARGO_CORE =
+struct
+
+(* context *)
+
+type context = {
+  terms: Argo_Term.context, (* the term context to identify equal expressions *)
+  iter: int, (* the current iteration of the search *)
+  cdcl: Argo_Cdcl.context, (* the context of the propositional solver *)
+  thy: Argo_Thy.context} (* the context of the theory solver *)
+
+fun mk_context terms iter cdcl thy: context = {terms=terms, iter=iter, cdcl=cdcl, thy=thy}
+
+val context = mk_context Argo_Term.context 1 Argo_Cdcl.context Argo_Thy.context
+
+fun backjump levels = funpow levels Argo_Thy.backtrack
+
+
+(* enriching the context *)
+
+fun identify i ({terms, iter, cdcl, thy}: context) =
+  let val (identified, terms) = Argo_Term.identify_item i terms
+  in (identified, mk_context terms iter cdcl thy) end
+
+fun add_atom i cx =
+  (case identify i cx of
+    known as (Argo_Term.Known _, _) => known
+  | (atom as Argo_Term.New t, {terms, iter, cdcl, thy}: context) =>
+      (case (Argo_Cdcl.add_atom t cdcl, Argo_Thy.add_atom t thy) of
+        (cdcl, (NONE, thy)) => (atom, mk_context terms iter cdcl thy)
+      | (cdcl, (SOME lit, thy)) =>
+          (case Argo_Cdcl.assume Argo_Thy.explain lit cdcl thy of
+            (NONE, cdcl, thy) => (atom, mk_context terms iter cdcl thy)
+          | (SOME _, _, _) => raise Fail "bad conflict with new atom")))
+
+fun add_axiom cls ({terms, iter, cdcl, thy}: context) =
+  let val (levels, cdcl) = Argo_Cdcl.add_axiom cls cdcl
+  in mk_context terms iter cdcl (backjump levels thy) end
+
+
+(* DPLL(T) loop: CDCL with theories *)
+
+datatype implications = None | Implications | Conflict of Argo_Cls.clause
+
+fun cdcl_assume [] cdcl thy = (NONE, cdcl, thy)
+  | cdcl_assume (lit :: lits) cdcl thy =
+      (* assume an assignment deduced by the theory solver *)
+      (case Argo_Cdcl.assume Argo_Thy.explain lit cdcl thy of
+        (NONE, cdcl, thy) => cdcl_assume lits cdcl thy
+      | (SOME cls, cdcl, thy) => (SOME cls, cdcl, thy))
+
+fun theory_deduce _ (conflict as (Conflict _, _, _)) = conflict
+  | theory_deduce f (result, cdcl, thy) =
+      (case f thy of
+        (Argo_Common.Implied [], thy) => (result, cdcl, thy)
+      | (Argo_Common.Implied lits, thy) => 
+          (* turn all implications of the theory solver into propositional assignments *)
+          (case cdcl_assume lits cdcl thy of
+            (NONE, cdcl, thy) => (Implications, cdcl, thy)
+          | (SOME cls, cdcl, thy) => (Conflict cls, cdcl, thy))
+      | (Argo_Common.Conflict cls, thy) => (Conflict cls, cdcl, thy))
+
+fun theory_assume [] cdcl thy = (None, cdcl, thy)
+  | theory_assume lps cdcl thy =
+      (None, cdcl, thy)
+      (* propagate all propositional implications to the theory solver *)
+      |> fold (theory_deduce o Argo_Thy.assume) lps
+      (* check the consistency of the theory model *)
+      |> theory_deduce Argo_Thy.check
+
+fun search limit cdcl thy =
+  (* collect all propositional implications of the last assignments *)
+  (case Argo_Cdcl.propagate cdcl of
+    (Argo_Common.Implied lps, cdcl) =>
+      (* propagate all propositional implications to the theory solver *)
+      (case theory_assume lps cdcl thy of
+        (None, cdcl, thy) =>
+          (* stop searching if the conflict limit has been exceeded *)
+          if limit <= 0 then (false, cdcl, thy)
+          else
+            (* no further propositional assignments, choose a value for the next unassigned atom *)
+            (case Argo_Cdcl.decide cdcl of
+              NONE => (true, cdcl, thy) (* the context is satisfiable *)
+            | SOME cdcl => search limit cdcl (Argo_Thy.add_level thy))
+      | (Implications, cdcl, thy) => search limit cdcl thy
+      | (Conflict ([], p), _, _) => Argo_Proof.unsat p
+      | (Conflict cls, cdcl, thy) => analyze cls limit cdcl thy)
+  | (Argo_Common.Conflict cls, cdcl) => analyze cls limit cdcl thy)
+
+and analyze cls limit cdcl thy =
+  (* analyze the conflict, probably using lazy explanations from the theory solver *)
+  let val (levels, cdcl, thy) = Argo_Cdcl.analyze Argo_Thy.explain cls cdcl thy
+  in search (limit - 1) cdcl (backjump levels thy) end
+
+fun luby_number i =
+  let
+    fun mult p = if p < i + 1 then mult (2 * p) else p
+    val p = mult 2
+  in if i = p - 1 then p div 2 else luby_number (i - (p div 2) + 1) end
+
+fun next_restart_limit iter = 100 * luby_number iter
+
+fun loop iter cdcl thy =
+  (* perform a limited search that is stopped after a certain number of conflicts *)
+  (case search (next_restart_limit iter) cdcl thy of
+    (true, cdcl, thy) => (iter + 1, cdcl, thy)
+  | (false, cdcl, thy) =>
+      (* restart the solvers to avoid that they get stuck in a fruitless search *)
+      let val (levels, cdcl) = Argo_Cdcl.restart cdcl
+      in loop (iter + 1) cdcl (backjump levels thy) end)
+
+fun run ({terms, iter, cdcl, thy}: context) =
+  let val (iter, cdcl, thy) = loop iter cdcl (Argo_Thy.prepare thy)
+  in mk_context terms iter cdcl thy end
+
+
+(* model *)
+
+fun model_of ({terms, cdcl, ...}: context) c =
+  (case Argo_Term.identify_item (Argo_Term.Expr (Argo_Expr.E (Argo_Expr.Con c, []))) terms of
+    (Argo_Term.Known t, _) => Argo_Cdcl.assignment_of cdcl (Argo_Lit.Pos t)
+  | (Argo_Term.New _, _) => NONE)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_expr.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,244 @@
+(*  Title:      Tools/Argo/sid_expr.ML
+    Author:     Sascha Boehme
+
+The input language of the Argo solver.
+*)
+
+signature ARGO_EXPR =
+sig
+  (* data types *)
+  datatype typ = Bool | Real | Func of typ * typ | Type of string
+  datatype kind =
+    True | False | Not | And | Or | Imp | Iff | Ite | Eq | App | Con of string * typ |
+    Le | Lt | Num of Rat.rat | Neg | Add | Sub | Mul | Div | Min | Max | Abs
+  datatype expr = E of kind * expr list
+
+  (* indices, equalities, orders *)
+  val int_of_kind: kind -> int
+  val con_ord: (string * typ) * (string * typ) -> order
+  val eq_kind: kind * kind -> bool
+  val kind_ord: kind * kind -> order
+  val eq_expr: expr * expr -> bool
+  val expr_ord: expr * expr -> order
+  val dual_expr: expr -> expr -> bool
+
+  (* constructors *)
+  val kind_of_string: string -> kind
+  val true_expr: expr
+  val false_expr: expr
+  val mk_not: expr -> expr
+  val mk_and: expr list -> expr
+  val mk_and2: expr -> expr -> expr
+  val mk_or: expr list -> expr
+  val mk_or2: expr -> expr -> expr
+  val mk_imp: expr -> expr -> expr
+  val mk_iff: expr -> expr -> expr
+  val mk_ite: expr -> expr -> expr -> expr
+  val mk_eq: expr -> expr -> expr
+  val mk_app: expr -> expr -> expr
+  val mk_con: string * typ -> expr
+  val mk_le: expr -> expr -> expr
+  val mk_lt: expr -> expr -> expr
+  val mk_num: Rat.rat -> expr
+  val mk_neg: expr -> expr
+  val mk_add: expr list -> expr
+  val mk_add2: expr -> expr -> expr
+  val mk_sub: expr -> expr -> expr
+  val mk_mul: expr -> expr -> expr
+  val mk_div: expr -> expr -> expr
+  val mk_min: expr -> expr -> expr
+  val mk_max: expr -> expr -> expr
+  val mk_abs: expr -> expr
+
+  (* type checking *)
+  exception TYPE of expr
+  exception EXPR of expr
+  val type_of: expr -> typ (* raises EXPR *)
+  val check: expr -> bool (* raises TYPE and EXPR *)
+end
+
+structure Argo_Expr: ARGO_EXPR =
+struct
+
+(* data types *)
+
+datatype typ = Bool | Real | Func of typ * typ | Type of string
+
+datatype kind =
+  True | False | Not | And | Or | Imp | Iff | Ite | Eq | App | Con of string * typ |
+  Le | Lt | Num of Rat.rat | Neg | Add | Sub | Mul | Div | Min | Max | Abs
+
+datatype expr = E of kind * expr list
+
+
+(* indices, equalities, orders *)
+
+fun int_of_type Bool = 0
+  | int_of_type Real = 1
+  | int_of_type (Func _) = 2
+  | int_of_type (Type _) = 3
+
+fun int_of_kind True = 0
+  | int_of_kind False = 1
+  | int_of_kind Not = 2
+  | int_of_kind And = 3
+  | int_of_kind Or = 4
+  | int_of_kind Imp = 5
+  | int_of_kind Iff = 6
+  | int_of_kind Ite = 7
+  | int_of_kind Eq = 8
+  | int_of_kind App = 9
+  | int_of_kind (Con _) = 10
+  | int_of_kind Le = 11
+  | int_of_kind Lt = 12
+  | int_of_kind (Num _) = 13
+  | int_of_kind Neg = 14
+  | int_of_kind Add = 15
+  | int_of_kind Sub = 16
+  | int_of_kind Mul = 17
+  | int_of_kind Div = 18
+  | int_of_kind Min = 19
+  | int_of_kind Max = 20
+  | int_of_kind Abs = 21
+
+fun eq_type (Bool, Bool) = true
+  | eq_type (Real, Real) = true
+  | eq_type (Func tys1, Func tys2) = eq_pair eq_type eq_type (tys1, tys2)
+  | eq_type (Type n1, Type n2) = (n1 = n2)
+  | eq_type _ = false
+
+fun type_ord (Bool, Bool) = EQUAL
+  | type_ord (Real, Real) = EQUAL
+  | type_ord (Type n1, Type n2) = fast_string_ord (n1, n2)
+  | type_ord (Func tys1, Func tys2) = prod_ord type_ord type_ord (tys1, tys2)
+  | type_ord (ty1, ty2) = int_ord (int_of_type ty1, int_of_type ty2)
+
+fun eq_con cp = eq_pair (op =) eq_type cp
+fun con_ord cp = prod_ord fast_string_ord type_ord cp
+
+fun eq_kind (Con c1, Con c2) = eq_con (c1, c2)
+  | eq_kind (Num n1, Num n2) = n1 = n2
+  | eq_kind (k1, k2) = (k1 = k2)
+
+fun kind_ord (Con c1, Con c2) = con_ord (c1, c2)
+  | kind_ord (Num n1, Num n2) = Rat.ord (n1, n2)
+  | kind_ord (k1, k2) = int_ord (int_of_kind k1, int_of_kind k2)
+
+fun eq_expr (E e1, E e2) = eq_pair eq_kind (eq_list eq_expr) (e1, e2)
+fun expr_ord (E e1, E e2) = prod_ord kind_ord (list_ord expr_ord) (e1, e2)
+
+fun dual_expr (E (Not, [e1])) e2 = eq_expr (e1, e2)
+  | dual_expr e1 (E (Not, [e2])) = eq_expr (e1, e2)
+  | dual_expr _ _ = false
+
+
+(* constructors *)
+
+val kind_of_string = the o Symtab.lookup (Symtab.make [
+  ("true", True),("false", False), ("not", Not), ("and", And), ("or", Or), ("imp", Imp),
+  ("iff", Iff), ("ite", Ite), ("eq", Eq), ("app", App), ("le", Le), ("lt", Lt), ("neg", Neg),
+  ("add", Add), ("sub", Sub), ("mul", Mul), ("div", Div), ("min", Min), ("max", Max), ("abs", Abs)])
+
+val true_expr = E (True, [])
+val false_expr = E (False, [])
+fun mk_not e = E (Not, [e])
+fun mk_and es = E (And, es)
+fun mk_and2 e1 e2 = mk_and [e1, e2]
+fun mk_or es = E (Or, es)
+fun mk_or2 e1 e2 = mk_or [e1, e2]
+fun mk_imp e1 e2 = E (Imp, [e1, e2])
+fun mk_iff e1 e2 = E (Iff, [e1, e2])
+fun mk_ite e1 e2 e3 = E (Ite, [e1, e2, e3])
+fun mk_eq e1 e2 = E (Eq, [e1, e2])
+fun mk_app e1 e2 = E (App, [e1, e2])
+fun mk_con n = E (Con n, [])
+fun mk_le e1 e2 = E (Le, [e1, e2])
+fun mk_lt e1 e2 = E (Lt, [e1, e2])
+fun mk_num r = E (Num r, [])
+fun mk_neg e = E (Neg, [e])
+fun mk_add es = E (Add, es)
+fun mk_add2 e1 e2 = mk_add [e1, e2]
+fun mk_sub e1 e2 = E (Sub, [e1, e2])
+fun mk_mul e1 e2 = E (Mul, [e1, e2])
+fun mk_div e1 e2 = E (Div, [e1, e2])
+fun mk_min e1 e2 = E (Min, [e1, e2])
+fun mk_max e1 e2 = E (Max, [e1, e2])
+fun mk_abs e = E (Abs, [e])
+
+
+(* type checking *)
+
+exception TYPE of expr
+exception EXPR of expr
+
+fun dest_func_type _ (Func tys) = tys
+  | dest_func_type e _ = raise TYPE e
+
+fun type_of (E (True, _)) = Bool
+  | type_of (E (False, _)) = Bool
+  | type_of (E (Not, _)) = Bool
+  | type_of (E (And, _)) = Bool
+  | type_of (E (Or, _)) = Bool
+  | type_of (E (Imp, _)) = Bool
+  | type_of (E (Iff, _)) = Bool
+  | type_of (E (Ite, [_, e, _])) = type_of e
+  | type_of (E (Eq, _)) = Bool
+  | type_of (E (App, [e, _])) = snd (dest_func_type e (type_of e))
+  | type_of (E (Con (_, ty), _)) = ty
+  | type_of (E (Le, _)) = Bool
+  | type_of (E (Lt, _)) = Bool
+  | type_of (E (Num _, _)) = Real
+  | type_of (E (Neg, _)) = Real
+  | type_of (E (Add, _)) = Real
+  | type_of (E (Sub, _)) = Real
+  | type_of (E (Mul, _)) = Real
+  | type_of (E (Div, _)) = Real
+  | type_of (E (Min, _)) = Real
+  | type_of (E (Max, _)) = Real
+  | type_of (E (Abs, _)) = Real
+  | type_of e = raise EXPR e
+
+fun all_type ty (E (_, es)) = forall (curry eq_type ty o type_of) es
+val all_bool = all_type Bool
+val all_real = all_type Real
+
+(*
+  Types as well as proper arities are checked.
+  Exception TYPE is raised for invalid types.
+  Exception EXPR is raised for invalid expressions and invalid arities.
+*)
+
+fun check (e as E (_, es)) = (forall check es andalso raw_check e) orelse raise TYPE e
+
+and raw_check (E (True, [])) = true
+  | raw_check (E (False, [])) = true
+  | raw_check (e as E (Not, [_])) = all_bool e
+  | raw_check (e as E (And, _ :: _)) = all_bool e
+  | raw_check (e as E (Or, _ :: _)) = all_bool e
+  | raw_check (e as E (Imp, [_, _])) = all_bool e
+  | raw_check (e as E (Iff, [_, _])) = all_bool e
+  | raw_check (E (Ite, [e1, e2, e3])) =
+      let val ty1 = type_of e1 and ty2 = type_of e2 and ty3 = type_of e3
+      in eq_type (ty1, Bool) andalso eq_type (ty2, ty3) end
+  | raw_check (E (Eq, [e1, e2])) =
+      let val ty1 = type_of e1 and ty2 = type_of e2
+      in eq_type (ty1, ty2) andalso not (eq_type (ty1, Bool)) end
+  | raw_check (E (App, [e1, e2])) =
+      eq_type (fst (dest_func_type e1 (type_of e1)), type_of e2)
+  | raw_check (E (Con _, [])) = true
+  | raw_check (E (Num _, [])) = true
+  | raw_check (e as E (Le, [_, _])) = all_real e
+  | raw_check (e as E (Lt, [_, _])) = all_real e
+  | raw_check (e as E (Neg, [_])) = all_real e
+  | raw_check (e as E (Add, _)) = all_real e
+  | raw_check (e as E (Sub, [_, _])) = all_real e
+  | raw_check (e as E (Mul, [_, _])) = all_real e
+  | raw_check (e as E (Div, [_, _])) = all_real e
+  | raw_check (e as E (Min, [_, _])) = all_real e
+  | raw_check (e as E (Max, [_, _])) = all_real e
+  | raw_check (e as E (Abs, [_])) = all_real e
+  | raw_check e = raise EXPR e
+
+end
+
+structure Argo_Exprtab = Table(type key = Argo_Expr.expr val ord = Argo_Expr.expr_ord)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_heap.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,186 @@
+(*  Title:      Tools/Argo/argo_heap.ML
+    Author:     Sascha Boehme
+
+A maximum-priority heap for literals with integer priorities and with inverse indices.
+The heap is intended to be used as VSIDS-like decision heuristics. This implementation
+is based on pairing heaps described in:
+
+  Chris Okasaki. Purely Functional Data Structures. Chapter 5.
+  Cambridge University Press, 1998.
+*)
+
+signature ARGO_HEAP =
+sig
+  type heap
+  val heap: heap
+  val insert: Argo_Lit.literal -> heap -> heap
+  val extract: heap -> (Argo_Lit.literal * heap) option
+  val increase: Argo_Lit.literal -> heap -> heap
+  val count: Argo_Lit.literal -> heap -> heap
+  val decay: heap -> heap
+  val rebuild: (Argo_Term.term -> bool) -> heap -> heap
+end
+
+structure Argo_Heap: ARGO_HEAP =
+struct
+
+(* heuristic activity constants *)
+
+val min_incr = 128
+fun decay_incr i = (i * 11) div 10
+val max_activity = Integer.pow 24 2
+val activity_rescale = Integer.pow 14 2
+
+
+(* data structures and basic operations *)
+
+datatype tree = E | T of Argo_Term.term * bool * tree list
+
+datatype parent = None | Root | Some of Argo_Term.term
+
+type heap = {
+  incr: int, (* the increment to apply in an increase operation *)
+  vals: ((int * int) * parent) Argo_Termtab.table, (* weights and parents of the stored terms *)
+  tree: tree} (* the pairing heap of literals; note: the tree caches literal polarities *)
+
+fun mk_heap incr vals tree: heap = {incr=incr, vals=vals, tree=tree}
+fun mk_heap' incr (tree, vals) = mk_heap incr vals tree
+
+val heap = mk_heap min_incr Argo_Termtab.empty E
+
+val empty_value = ((0, 0), None)
+fun value_of vals t = the_default empty_value (Argo_Termtab.lookup vals t)
+fun map_value t = Argo_Termtab.map_default (t, empty_value)
+
+
+(* weight operations *)
+
+(*
+  The weight of a term is a pair of activity and count. The activity describes how
+  often a term participated in conflicts. The count describes how often a term occurs
+  in clauses.
+*)
+
+val weight_ord = prod_ord int_ord int_ord
+
+fun weight_of vals t = fst (value_of vals t)
+
+fun less_than vals t1 t2 = weight_ord (weight_of vals t1, weight_of vals t2) = LESS
+
+fun rescale activity = activity div activity_rescale
+
+fun incr_activity incr t = map_value t (apfst (apfst (Integer.add incr)))
+fun incr_count t = map_value t (apfst (apsnd (Integer.add 1)))
+
+fun rescale_activities a incr vals =
+  if a <= max_activity then (incr, vals)
+  else (rescale incr, Argo_Termtab.map (fn _ => apfst (apfst rescale)) vals)
+
+
+(* reverse index operations *)
+
+(*
+  The reverse index is required to retrieve elements when increasing their priorities.
+*)
+
+fun contains vals t =
+  (case value_of vals t of
+    (_, None) => false
+  | _ => true)
+
+fun path_to vals t parents =
+  (case value_of vals t of
+    (_, Root) => parents
+  | (_, Some parent) => path_to vals parent (t :: parents)
+  | _ => raise Fail "bad heap")
+
+fun put_parent t parent = map_value t (apsnd (K parent))
+fun delete t = put_parent t None
+fun root t = put_parent t Root
+
+fun as_root (tree as T (t, _, _), vals) = (tree, root t vals)
+  | as_root x = x
+
+
+(* pairing heap operations *)
+
+fun merge E tree vals = (tree, vals)
+  | merge tree E vals = (tree, vals)
+  | merge (tree1 as T (t1, p1, trees1)) (tree2 as T (t2, p2, trees2)) vals =
+      if less_than vals t1 t2 then (T (t2, p2, tree1 :: trees2), put_parent t1 (Some t2) vals)
+      else (T (t1, p1, tree2 :: trees1), put_parent t2 (Some t1) vals)
+
+fun merge_pairs [] vals = (E, vals)
+  | merge_pairs [tree] vals = (tree, vals)
+  | merge_pairs (tree1 :: tree2 :: trees) vals =
+      vals |> merge tree1 tree2 ||>> merge_pairs trees |-> uncurry merge
+
+
+(* cutting subtrees specified by a path *)
+
+(*
+  The extractions are performed in such a way that the heap is changed in as few positions
+  as possible.
+*)
+
+fun with_index f u ((tree as T (t, _, _)) :: trees) =
+      if Argo_Term.eq_term (t, u) then f tree ||> (fn E => trees | tree => tree :: trees)
+      else with_index f u trees ||> cons tree
+  | with_index _ _ _ = raise Fail "bad heap"
+
+fun lift_index f u (T (t, p, trees)) = with_index f u trees ||> (fn trees => T (t, p, trees))
+  | lift_index _ _ E = raise Fail "bad heap"
+
+fun cut t [] tree = lift_index (fn tree => (tree, E)) t tree
+  | cut t (parent :: ts) tree = lift_index (cut t ts) parent tree
+
+
+(* filtering the heap *)
+
+val proper_trees = filter (fn E => false | T _ => true)
+
+fun filter_tree _ E vals = (E, vals)
+  | filter_tree pred (T (t, p, ts)) vals =
+      let val (ts, vals) = fold_map (filter_tree pred) ts vals |>> proper_trees
+      in if pred t then (T (t, p, ts), vals) else merge_pairs ts (delete t vals) end
+
+
+(* exported heap operations *)
+
+fun insert lit (h as {incr, vals, tree}: heap) = 
+  let val (t, p) = Argo_Lit.dest lit
+  in if contains vals t then h else mk_heap' incr (merge tree (T (t, p, [])) (root t vals)) end
+
+fun extract ({tree=E, ...}: heap) = NONE
+  | extract ({incr, vals, tree=T (t, p, ts)}: heap) =
+      SOME (Argo_Lit.literal t p, mk_heap' incr (as_root (merge_pairs ts (delete t vals))))
+
+fun with_term lit f = f (Argo_Lit.term_of lit)
+
+(*
+  If the changed weight violates the heap property, the corresponding tree
+  is extracted and merged with the root.
+*)
+
+fun fix t (w, Some parent) (incr, vals) tree =
+      if weight_ord (weight_of vals parent, w) = LESS then
+        let val (tree1, tree2) = cut t (path_to vals parent []) tree
+        in mk_heap' incr (merge tree1 tree2 (root t vals)) end
+      else mk_heap incr vals tree
+  | fix _ _ (incr, vals) tree = mk_heap incr vals tree
+
+fun increase lit ({incr, vals, tree}: heap) = with_term lit (fn t =>
+  let
+    val vals = incr_activity incr t vals
+    val value as ((a, _), _) = value_of vals t
+  in fix t value (rescale_activities a incr vals) tree end)
+
+fun count lit ({incr, vals, tree}: heap) = with_term lit (fn t =>
+  let val vals = incr_count t vals
+  in fix t (value_of vals t) (incr, vals) tree end)
+
+fun decay ({incr, vals, tree}: heap) = mk_heap (decay_incr incr) vals tree
+
+fun rebuild pred ({incr, vals, tree}: heap) = mk_heap' incr (filter_tree pred tree vals)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_lit.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,64 @@
+(*  Title:      Tools/Argo/argo_lit.ML
+    Author:     Sascha Boehme
+
+Representation of literals. Literals are terms with a polarity, either positive or negative.
+A literal for term t with positive polarity is equivalent to t.
+A literal for term t with negative polarity is equivalent to the propositional negation of t.
+*)
+
+signature ARGO_LIT =
+sig
+  datatype literal = Pos of Argo_Term.term | Neg of Argo_Term.term
+  val literal: Argo_Term.term -> bool -> literal
+  val dest: literal -> Argo_Term.term * bool
+  val term_of: literal -> Argo_Term.term
+  val signed_id_of: literal -> int
+  val signed_expr_of: literal -> Argo_Expr.expr
+  val negate: literal -> literal
+  val eq_id: literal * literal -> bool
+  val eq_lit: literal * literal -> bool
+  val dual_lit: literal * literal -> bool
+end
+
+structure Argo_Lit: ARGO_LIT =
+struct
+
+(* data type *)
+
+datatype literal = Pos of Argo_Term.term | Neg of Argo_Term.term
+
+
+(* operations *)
+
+fun literal t true = Pos t
+  | literal t false = Neg t
+
+fun dest (Pos t) = (t, true)
+  | dest (Neg t) = (t, false)
+
+fun term_of (Pos t) = t
+  | term_of (Neg t) = t
+
+fun signed_id_of (Pos t) = Argo_Term.id_of t
+  | signed_id_of (Neg t) = ~(Argo_Term.id_of t)
+
+fun signed_expr_of (Pos t) = Argo_Term.expr_of t
+  | signed_expr_of (Neg t) = Argo_Expr.mk_not (Argo_Term.expr_of t)
+
+fun id_of (Pos t) = Argo_Term.id_of t
+  | id_of (Neg t) = Argo_Term.id_of t
+
+fun negate (Pos t) = Neg t
+  | negate (Neg t) = Pos t
+
+fun eq_id (lit1, lit2) = (id_of lit1 = id_of lit2)
+
+fun eq_lit (Pos t1, Pos t2) = Argo_Term.eq_term (t1, t2)
+  | eq_lit (Neg t1, Neg t2) = Argo_Term.eq_term (t1, t2)
+  | eq_lit _ = false
+
+fun dual_lit (Pos t1, Neg t2) = Argo_Term.eq_term (t1, t2)
+  | dual_lit (Neg t1, Pos t2) = Argo_Term.eq_term (t1, t2)
+  | dual_lit _ = false
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_proof.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,387 @@
+(*  Title:      Tools/Argo/argo_proof.ML
+    Author:     Sascha Boehme
+
+The proof language of the Argo solver.
+
+Proofs trace the inferences of the solver. They can be used to check unsatisfiability results.
+
+The proof language is inspired by:
+
+  Leonardo  de  Moura  and  Nikolaj  Bj/orner. Proofs and Refutations, and Z3. In
+  Proceedings of the LPAR 2008 Workshops, Knowledge Exchange: Automated Provers and Proof
+  Assistants, and the 7th International Workshop on the Implementation of Logics,
+  volume 418 of CEUR Workshop Proceedings. CEUR-WS.org, 2008.
+*)
+
+signature ARGO_PROOF =
+sig
+  (* types *)
+  type proof_id
+  datatype tautology =
+    Taut_And_1 of int | Taut_And_2 of int * int | Taut_Or_1 of int * int | Taut_Or_2 of int |
+    Taut_Iff_1 | Taut_Iff_2 | Taut_Iff_3 | Taut_Iff_4 | Taut_Ite_Then | Taut_Ite_Else
+  datatype inequality = Le | Lt
+  datatype rewrite =
+    Rewr_Not_True | Rewr_Not_False | Rewr_Not_Not | Rewr_Not_And of int | Rewr_Not_Or of int |
+    Rewr_Not_Iff |
+    Rewr_And_False of int | Rewr_And_Dual of int * int | Rewr_And_Sort of int * int list list |
+    Rewr_Or_True of int | Rewr_Or_Dual of int * int | Rewr_Or_Sort of int * int list list |
+    Rewr_Iff_True | Rewr_Iff_False | Rewr_Iff_Not_Not | Rewr_Iff_Refl | Rewr_Iff_Symm |
+    Rewr_Iff_Dual |
+    Rewr_Imp | Rewr_Ite_Prop | Rewr_Ite_True | Rewr_Ite_False | Rewr_Ite_Eq |
+    Rewr_Eq_Refl | Rewr_Eq_Symm |
+    Rewr_Neg | Rewr_Add of (Rat.rat * int option) list * (Rat.rat * int option) list | Rewr_Sub |
+    Rewr_Mul_Nums of Rat.rat * Rat.rat | Rewr_Mul_Zero | Rewr_Mul_One | Rewr_Mul_Comm |
+    Rewr_Mul_Assoc | Rewr_Mul_Sum | Rewr_Div_Nums of Rat.rat * Rat.rat | Rewr_Div_Zero |
+    Rewr_Div_One | Rewr_Div_Mul | Rewr_Div_Inv | Rewr_Div_Left | Rewr_Div_Right | Rewr_Min |
+    Rewr_Max | Rewr_Abs | Rewr_Eq_Le | Rewr_Ineq_Nums of inequality * bool |
+    Rewr_Ineq_Add of inequality * Rat.rat | Rewr_Ineq_Sub of inequality |
+    Rewr_Ineq_Mul of inequality * Rat.rat | Rewr_Not_Ineq of inequality
+  datatype conv =
+    Keep_Conv | Then_Conv of conv * conv | Args_Conv of conv list | Rewr_Conv of rewrite
+  datatype rule =
+    Axiom of int | Taut of tautology * Argo_Expr.expr | Conjunct of int * int | Rewrite of conv |
+    Hyp of int * Argo_Expr.expr | Clause of int list | Lemma of int list | Unit_Res of int |
+    Refl of Argo_Expr.expr | Symm | Trans | Cong | Subst | Linear_Comb
+  type proof
+
+  (* equalities and orders *)
+  val eq_proof_id: proof_id * proof_id -> bool
+  val proof_id_ord: proof_id * proof_id -> order
+
+  (* conversion constructors *)
+  val keep_conv: conv
+  val mk_then_conv: conv -> conv -> conv
+  val mk_args_conv: conv list -> conv
+  val mk_rewr_conv: rewrite -> conv
+
+  (* context *)
+  type context
+  val cdcl_context: context
+  val cc_context: context
+  val simplex_context: context
+  val solver_context: context
+
+  (* proof constructors *)
+  val mk_axiom: int -> context -> proof * context
+  val mk_taut: tautology -> Argo_Expr.expr -> context -> proof * context
+  val mk_conj: int -> int -> proof -> context -> proof * context
+  val mk_rewrite: conv -> proof -> context -> proof * context
+  val mk_hyp: Argo_Lit.literal -> context -> proof * context
+  val mk_clause: Argo_Lit.literal list -> proof -> context -> proof * context
+  val mk_lemma: Argo_Lit.literal list -> proof -> context -> proof * context
+  val mk_unit_res: Argo_Lit.literal -> proof -> proof -> context -> proof * context
+  val mk_refl: Argo_Term.term -> context -> proof * context
+  val mk_symm: proof -> context -> proof * context
+  val mk_trans: proof -> proof -> context -> proof * context
+  val mk_cong: proof -> proof -> context -> proof * context
+  val mk_subst: proof -> proof -> proof -> context -> proof * context
+  val mk_linear_comb: proof list -> context -> proof * context
+
+  (* proof destructors *)
+  val id_of: proof -> proof_id
+  val dest: proof -> proof_id * rule * proof list
+
+  (* string representations *)
+  val string_of_proof_id: proof_id -> string
+  val string_of_taut: tautology -> string
+  val string_of_rule: rule -> string
+
+  (* unsatisfiability *)
+  exception UNSAT of proof
+  val unsat: proof -> 'a (* raises UNSAT *)
+end
+
+structure Argo_Proof: ARGO_PROOF =
+struct
+
+(* types *)
+
+datatype tautology =
+  Taut_And_1 of int | Taut_And_2 of int * int | Taut_Or_1 of int * int | Taut_Or_2 of int |
+  Taut_Iff_1 | Taut_Iff_2 | Taut_Iff_3 | Taut_Iff_4 | Taut_Ite_Then | Taut_Ite_Else
+
+datatype inequality = Le | Lt
+
+datatype rewrite =
+  Rewr_Not_True | Rewr_Not_False | Rewr_Not_Not | Rewr_Not_And of int | Rewr_Not_Or of int |
+  Rewr_Not_Iff |
+  Rewr_And_False of int | Rewr_And_Dual of int * int | Rewr_And_Sort of int * int list list |
+  Rewr_Or_True of int | Rewr_Or_Dual of int * int | Rewr_Or_Sort of int * int list list |
+  Rewr_Iff_True | Rewr_Iff_False | Rewr_Iff_Not_Not | Rewr_Iff_Refl | Rewr_Iff_Symm |
+  Rewr_Iff_Dual |
+  Rewr_Imp | Rewr_Ite_Prop | Rewr_Ite_True | Rewr_Ite_False | Rewr_Ite_Eq |
+  Rewr_Eq_Refl | Rewr_Eq_Symm |
+  Rewr_Neg | Rewr_Add of (Rat.rat * int option) list * (Rat.rat * int option) list | Rewr_Sub |
+  Rewr_Mul_Nums of Rat.rat * Rat.rat | Rewr_Mul_Zero | Rewr_Mul_One | Rewr_Mul_Comm |
+  Rewr_Mul_Assoc | Rewr_Mul_Sum | Rewr_Div_Nums of Rat.rat * Rat.rat | Rewr_Div_Zero |
+  Rewr_Div_One | Rewr_Div_Mul | Rewr_Div_Inv | Rewr_Div_Left | Rewr_Div_Right | Rewr_Min |
+  Rewr_Max | Rewr_Abs | Rewr_Eq_Le | Rewr_Ineq_Nums of inequality * bool |
+  Rewr_Ineq_Add of inequality * Rat.rat | Rewr_Ineq_Sub of inequality |
+  Rewr_Ineq_Mul of inequality * Rat.rat | Rewr_Not_Ineq of inequality
+
+datatype conv =
+  Keep_Conv | Then_Conv of conv * conv | Args_Conv of conv list | Rewr_Conv of rewrite
+
+datatype rule =
+  Axiom of int | Taut of tautology * Argo_Expr.expr | Conjunct of int * int | Rewrite of conv |
+  Hyp of int * Argo_Expr.expr | Clause of int list | Lemma of int list | Unit_Res of int |
+  Refl of Argo_Expr.expr | Symm | Trans | Cong | Subst | Linear_Comb
+
+(*
+  Proof identifiers are intentially hidden to prevent that functions outside of this structure
+  are able to build proofs. Proof can hence only be built by the functions provided by
+  this structure.
+*)
+
+datatype proof_id = Cdcl of int | Cc of int | Simplex of int | Solver of int
+
+datatype proof = Proof of proof_id * rule * proof list
+
+
+(* internal functions *)
+
+val proof_id_card = 4
+
+fun raw_proof_id (Cdcl i) = i
+  | raw_proof_id (Cc i) = i
+  | raw_proof_id (Simplex i) = i
+  | raw_proof_id (Solver i) = i
+
+
+(* equalities and orders *)
+
+fun int_of_proof_id (Cdcl _) = 0
+  | int_of_proof_id (Cc _) = 1
+  | int_of_proof_id (Simplex _) = 2
+  | int_of_proof_id (Solver _) = 3
+
+fun eq_proof_id (Cdcl i1, Cdcl i2) = (i1 = i2)
+  | eq_proof_id (Cc i1, Cc i2) = (i1 = i2)
+  | eq_proof_id (Simplex i1, Simplex i2) = (i1 = i2)
+  | eq_proof_id (Solver i1, Solver i2) = (i1 = i2)
+  | eq_proof_id _ = false
+
+fun proof_id_ord (Cdcl i1, Cdcl i2) = int_ord (i1, i2)
+  | proof_id_ord (Cc i1, Cc i2) = int_ord (i1, i2)
+  | proof_id_ord (Simplex i1, Simplex i2) = int_ord (i1, i2)
+  | proof_id_ord (Solver i1, Solver i2) = int_ord (i1, i2)
+  | proof_id_ord (id1, id2) = int_ord (int_of_proof_id id1, int_of_proof_id id2)
+
+
+(* conversion constructors *)
+
+val keep_conv = Keep_Conv
+
+fun mk_then_conv Keep_Conv c = c
+  | mk_then_conv c Keep_Conv = c
+  | mk_then_conv c1 c2 = Then_Conv (c1, c2)
+
+fun mk_args_conv cs =
+  if forall (fn Keep_Conv => true | _ => false) cs then Keep_Conv
+  else Args_Conv cs
+
+fun mk_rewr_conv r = Rewr_Conv r
+
+
+(* context *)
+
+(*
+  The proof context stores the next unused identifier. Incidentally, the same type as
+  for the proof identifier can be used as context. Every proof-producing module of the
+  solver has its own proof identifier domain to ensure globally unique identifiers
+  without sharing a single proof context.
+*)
+
+type context = proof_id
+
+val cdcl_context = Cdcl 0
+val cc_context = Cc 0
+val simplex_context = Simplex 0
+val solver_context = Solver 0
+
+fun next_id (id as Cdcl i) = (id, Cdcl (i + 1))
+  | next_id (id as Cc i) = (id, Cc (i + 1))
+  | next_id (id as Simplex i) = (id, Simplex (i + 1))
+  | next_id (id as Solver i) = (id, Solver (i + 1))
+
+
+(* proof destructors *)
+
+fun id_of (Proof (id, _, _)) = id
+
+fun dest (Proof p) = p
+
+
+(* proof constructors *)
+
+fun mk_proof r ps cx =
+  let val (id, cx) = next_id cx
+  in (Proof (id, r, ps), cx) end
+
+fun mk_axiom i = mk_proof (Axiom i) []
+fun mk_taut t e = mk_proof (Taut (t, e)) []
+fun mk_conj i n p = mk_proof (Conjunct (i, n)) [p]
+
+fun mk_rewrite Keep_Conv p cx = (p, cx)
+  | mk_rewrite c p cx = mk_proof (Rewrite c) [p] cx
+
+fun mk_hyp lit = mk_proof (Hyp (Argo_Lit.signed_id_of lit, Argo_Lit.signed_expr_of lit)) []
+fun mk_clause lits p cx = mk_proof (Clause (map Argo_Lit.signed_id_of lits)) [p] cx
+fun mk_lemma lits p = mk_proof (Lemma (map Argo_Lit.signed_id_of lits)) [p]
+
+(*
+  Replay of unit-resolution steps can be optimized if all premises follow a specific form.
+  Therefore, each premise is checked if it is in clausal form.
+*)
+
+fun check_clause (p as Proof (_, Clause _, _)) = p
+  | check_clause (p as Proof (_, Lemma _, _)) = p
+  | check_clause (p as Proof (_, Unit_Res _, _)) = p
+  | check_clause _ = raise Fail "bad clause proof"
+
+fun mk_unit t p1 p2 = mk_proof (Unit_Res (Argo_Term.id_of t)) (map check_clause [p1, p2])
+
+fun mk_unit_res (Argo_Lit.Pos t) p1 p2 = mk_unit t p1 p2
+  | mk_unit_res (Argo_Lit.Neg t) p1 p2 = mk_unit t p2 p1
+
+fun mk_refl t = mk_proof (Refl (Argo_Term.expr_of t)) []
+fun mk_symm p = mk_proof Symm [p]
+
+fun mk_trans (Proof (_, Refl _, _)) p2 = pair p2
+  | mk_trans p1 (Proof (_, Refl _, _)) = pair p1
+  | mk_trans p1 p2 = mk_proof Trans [p1, p2]
+
+fun mk_cong p1 p2 = mk_proof Cong [p1, p2]
+
+fun mk_subst p1 (Proof (_, Refl _, _)) (Proof (_, Refl _, _)) = pair p1
+  | mk_subst p1 p2 p3 = mk_proof Subst [p1, p2, p3]
+
+fun mk_linear_comb ps = mk_proof Linear_Comb ps
+
+
+(* string representations *)
+
+fun string_of_proof_id id = string_of_int (proof_id_card * raw_proof_id id + int_of_proof_id id)
+
+fun string_of_list l r f xs = enclose l r (space_implode ", " (map f xs))
+fun parens f xs = string_of_list "(" ")" f xs
+fun brackets f xs = string_of_list "[" "]" f xs
+
+fun string_of_taut (Taut_And_1 n) = "and " ^ string_of_int n
+  | string_of_taut (Taut_And_2 (i, n)) = "and " ^ parens string_of_int [i, n]
+  | string_of_taut (Taut_Or_1 (i, n)) = "or " ^ parens string_of_int [i, n]
+  | string_of_taut (Taut_Or_2 n) = "or " ^ string_of_int n
+  | string_of_taut Taut_Iff_1 = "(p1 == p2) | p1 | p2"
+  | string_of_taut Taut_Iff_2 = "(p1 == p2) | ~p1 | ~p2"
+  | string_of_taut Taut_Iff_3 = "~(p1 == p2) | ~p1 | p2"
+  | string_of_taut Taut_Iff_4 = "~(p1 == p2) | p1 | ~p2"
+  | string_of_taut Taut_Ite_Then = "~p | (ite p t1 t2) = t1"
+  | string_of_taut Taut_Ite_Else = "p | (ite p t1 t2) = t2"
+
+fun string_of_rewr Rewr_Not_True = "~T = F"
+  | string_of_rewr Rewr_Not_False = "~F = T"
+  | string_of_rewr Rewr_Not_Not = "~~p = p"
+  | string_of_rewr (Rewr_Not_And n) =
+      "~(and [" ^ string_of_int n ^ "]) = (or [" ^ string_of_int n ^ "])" 
+  | string_of_rewr (Rewr_Not_Or n) =
+      "~(or [" ^ string_of_int n ^ "]) = (and [" ^ string_of_int n ^ "])"
+  | string_of_rewr Rewr_Not_Iff = "~(p1 == p2) = (~p1 == ~p2)"
+  | string_of_rewr (Rewr_And_False i) = "(and ... F(" ^ string_of_int i ^ ") ...) = F"
+  | string_of_rewr (Rewr_And_Dual (i1, i2)) =
+      "(and ... p(" ^ string_of_int i1 ^ ") ... ~p(" ^ string_of_int i2 ^ ") ...) = F"
+  | string_of_rewr (Rewr_And_Sort (n, iss)) =
+      "(and [" ^ string_of_int n ^ "]) = " ^
+      "(and " ^ brackets (brackets string_of_int) iss ^ ")" 
+  | string_of_rewr (Rewr_Or_True i) = "(or ... T(" ^ string_of_int i ^ ") ...) = T"
+  | string_of_rewr (Rewr_Or_Dual (i1, i2)) =
+      "(or ... p(" ^ string_of_int i1 ^ ") ... ~p(" ^ string_of_int i2 ^ ") ...) = T"
+  | string_of_rewr (Rewr_Or_Sort (n, iss)) =
+      "(or [" ^ string_of_int n ^ "]) = " ^
+      "(or " ^ brackets (brackets string_of_int) iss ^ ")" 
+  | string_of_rewr Rewr_Iff_True = "(p == T) = p"
+  | string_of_rewr Rewr_Iff_False = "(p == F) = ~p"
+  | string_of_rewr Rewr_Iff_Not_Not = "(~p1 == ~p2) = (p1 == p2)"
+  | string_of_rewr Rewr_Iff_Refl = "(p == p) = T"
+  | string_of_rewr Rewr_Iff_Symm = "(p1 == p2) = (p2 == p1)"
+  | string_of_rewr Rewr_Iff_Dual = "(p == ~p) = F"
+  | string_of_rewr Rewr_Imp = "(p1 --> p2) = (~p1 | p2)"
+  | string_of_rewr Rewr_Ite_Prop = "(if p1 p2 p2) = ((~p1 | p2) & (p1 | p3) & (p2 | p3))"
+  | string_of_rewr Rewr_Ite_True = "(if T t1 t2) = t1"
+  | string_of_rewr Rewr_Ite_False = "(if F t1 t2) = t2"
+  | string_of_rewr Rewr_Ite_Eq = "(if p t t) = t"
+  | string_of_rewr Rewr_Eq_Refl = "(e = e) = T"
+  | string_of_rewr Rewr_Eq_Symm = "(e1 = e2) = (e2 = e1)"
+  | string_of_rewr Rewr_Neg = "-e = -1 * e"
+  | string_of_rewr (Rewr_Add (p1, p2)) =
+      let
+        fun string_of_monom (n, NONE) = Rat.string_of_rat n
+          | string_of_monom (n, SOME i) =
+              (if n = @1 then "" else Rat.string_of_rat n ^ " * ") ^ "e" ^ string_of_int i
+        fun string_of_polynom ms = space_implode " + " (map string_of_monom ms)
+      in string_of_polynom p1 ^ " = " ^ string_of_polynom p2 end
+  | string_of_rewr Rewr_Sub = "e1 - e2 = e1 + -1 * e2"
+  | string_of_rewr (Rewr_Mul_Nums (n1, n2)) =
+      Rat.string_of_rat n1 ^ " * " ^ Rat.string_of_rat n2 ^ " = " ^ Rat.string_of_rat (n1 * n2)
+  | string_of_rewr Rewr_Mul_Zero = "0 * e = 0"
+  | string_of_rewr Rewr_Mul_One = "1 * e = e"
+  | string_of_rewr Rewr_Mul_Comm = "e1 * e2 = e2 * e1"
+  | string_of_rewr Rewr_Mul_Assoc = "n1 * (n2 * e) = (n1 * n2) * e"
+  | string_of_rewr Rewr_Mul_Sum = "n * (e1 + ... + em) = n * e1 + ... n * em"
+  | string_of_rewr (Rewr_Div_Nums (n1, n2)) =
+      Rat.string_of_rat n1 ^ " / " ^ Rat.string_of_rat n2 ^ " = " ^ Rat.string_of_rat (n1 / n2)
+  | string_of_rewr Rewr_Div_Zero = "0 / e = 0"
+  | string_of_rewr Rewr_Div_One = "e / 1 = e"
+  | string_of_rewr Rewr_Div_Mul = "n / e = n * (1 / e)"
+  | string_of_rewr Rewr_Div_Inv = "e / n = 1/n * e"
+  | string_of_rewr Rewr_Div_Left = "(n * e1) / e2 = n * (e1 / e2)"
+  | string_of_rewr Rewr_Div_Right = "e1 / (n * e2) = 1/n * (e1 / e2)"
+  | string_of_rewr Rewr_Min = "min e1 e2 = (if e1 <= e2 then e1 else e2)"
+  | string_of_rewr Rewr_Max = "max e1 e2 = (if e1 < e2 then e2 else e1)"
+  | string_of_rewr Rewr_Abs = "abs e = (if 0 <= e then e else -e)"
+  | string_of_rewr Rewr_Eq_Le = "(e1 = e2) = (and (e1 <= e2) (e2 <= e1))"
+  | string_of_rewr (Rewr_Ineq_Nums (Le, true)) = "(n1 <= n2) = true"
+  | string_of_rewr (Rewr_Ineq_Nums (Le, false)) = "(n1 <= n2) = false"
+  | string_of_rewr (Rewr_Ineq_Nums (Lt, true)) = "(n1 < n2) = true"
+  | string_of_rewr (Rewr_Ineq_Nums (Lt, false)) = "(n1 < n2) = false"
+  | string_of_rewr (Rewr_Ineq_Add (Le, _)) = "(e1 <= e2) = (e1 + n <= e2 + n)"
+  | string_of_rewr (Rewr_Ineq_Add (Lt, _)) = "(e1 < e2) = (e1 + n < e2 + n)"
+  | string_of_rewr (Rewr_Ineq_Sub Le) = "(e1 <= e2) = (e1 - e2 <= 0)"
+  | string_of_rewr (Rewr_Ineq_Sub Lt) = "(e1 < e2) = (e1 - e2 < 0)"
+  | string_of_rewr (Rewr_Ineq_Mul (Le, _)) = "(e1 <= e2) = (n * e1 <= n * e2)"
+  | string_of_rewr (Rewr_Ineq_Mul (Lt, _)) = "(e1 < e2) = (n * e1 < n * e2)"
+  | string_of_rewr (Rewr_Not_Ineq Le) = "~(e1 <= e2) = (e2 < e1)"
+  | string_of_rewr (Rewr_Not_Ineq Lt) = "~(e1 < e2) = (e2 <= e1)"
+
+fun flatten_then_conv (Then_Conv (c1, c2)) = flatten_then_conv c1 @ flatten_then_conv c2
+  | flatten_then_conv c = [c]
+
+fun string_of_conv Keep_Conv = "_"
+  | string_of_conv (c as Then_Conv _) =
+      space_implode " then " (map (enclose "(" ")" o string_of_conv) (flatten_then_conv c))
+  | string_of_conv (Args_Conv cs) = "args " ^ brackets string_of_conv cs
+  | string_of_conv (Rewr_Conv r) = string_of_rewr r
+
+fun string_of_rule (Axiom i) = "axiom " ^ string_of_int i
+  | string_of_rule (Taut (t, _)) = "tautology: " ^ string_of_taut t
+  | string_of_rule (Conjunct (i, n)) = "conjunct " ^ string_of_int i ^ " of " ^ string_of_int n
+  | string_of_rule (Rewrite c) = "rewrite: " ^ string_of_conv c
+  | string_of_rule (Hyp (i, _)) = "hypothesis " ^ string_of_int i
+  | string_of_rule (Clause is) = "clause " ^ brackets string_of_int is
+  | string_of_rule (Lemma is) = "lemma " ^ brackets string_of_int is
+  | string_of_rule (Unit_Res i) = "unit-resolution " ^ string_of_int i
+  | string_of_rule (Refl _) = "reflexivity"
+  | string_of_rule Symm = "symmetry"
+  | string_of_rule Trans = "transitivity"
+  | string_of_rule Cong = "congruence"
+  | string_of_rule Subst = "substitution"
+  | string_of_rule Linear_Comb = "linear-combination"
+
+
+(* unsatisfiability *)
+
+exception UNSAT of proof
+
+fun unsat p = raise UNSAT p
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_rewr.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,266 @@
+(*  Title:      Tools/Argo/argo_rewr.ML
+    Author:     Sascha Boehme
+
+Bottom-up rewriting of expressions based on rewrite rules and rewrite functions.
+*)
+
+signature ARGO_REWR =
+sig
+  (* patterns *)
+  datatype pattern =
+    V of int |
+    E of Argo_Expr.expr |
+    A of Argo_Expr.kind |
+    P of Argo_Expr.kind * pattern list |
+    M of pattern |
+    X
+
+  (* scanning patterns from strings *)
+  val scan: string -> pattern
+
+  (* pattern matching *)
+  type env
+  val get_all: env -> Argo_Expr.expr list
+  val get: env -> int -> Argo_Expr.expr
+
+  (* conversions *)
+  type conv = Argo_Expr.expr -> Argo_Expr.expr * Argo_Proof.conv
+  val keep: conv
+  val seq: conv list -> conv
+  val args: conv -> conv
+  val rewr: Argo_Proof.rewrite -> Argo_Expr.expr -> conv
+
+  (* context *)
+  type context
+  val context: context
+  val flat: string ->
+    (int -> Argo_Expr.expr list -> (Argo_Proof.rewrite * Argo_Expr.expr) option) ->
+    context -> context
+  val func: string -> (env -> (Argo_Proof.rewrite * pattern) option) -> context -> context
+  val rule: Argo_Proof.rewrite -> string -> string -> context -> context
+
+  (* rewriting *)
+  val rewrite: context -> conv
+  val rewrite_top: context -> conv
+  val with_proof: conv -> Argo_Expr.expr * Argo_Proof.proof -> Argo_Proof.context ->
+    (Argo_Expr.expr * Argo_Proof.proof) * Argo_Proof.context
+end
+
+structure Argo_Rewr: ARGO_REWR =
+struct
+
+(* patterns *)
+
+(*
+  Patterns are used for matching against expressions and as a schema to construct
+  expressions. For matching, only V, E, A and P must be used. For the schema,
+  additionally M and X can be used.
+*)
+
+datatype pattern =
+  V of int | (* indexed placeholder where the index must be greater than 0 *)
+  E of Argo_Expr.expr | (* expression without placeholders *)
+  A of Argo_Expr.kind | (* placeholder for the arguments of an n-ary expression *)
+  P of Argo_Expr.kind * pattern list | (* expression with optional placeholders *)
+  M of pattern | (* mapping operator to iterate over an argument list of an n-ary expression *)
+  X (* closure argument under a mapping operator representing an expression *)
+
+fun int_of_pattern (E _) = 0
+  | int_of_pattern (P _) = 1
+  | int_of_pattern (A _) = 2
+  | int_of_pattern (V _) = 3
+  | int_of_pattern _ = raise Fail "bad pattern"
+
+(*
+  Specific patterns are ordered before generic patterns, since pattern matching
+  performs a linear search for the most suitable pattern.
+*)
+
+fun pattern_ord (E e1, E e2) = Argo_Expr.expr_ord (e1, e2)
+  | pattern_ord (P (k1, ps1), P (k2, ps2)) =
+      (case Argo_Expr.kind_ord (k1, k2) of EQUAL => list_ord pattern_ord (ps1, ps2) | x => x)
+  | pattern_ord (A k1, A k2) = Argo_Expr.kind_ord (k1, k2)
+  | pattern_ord (V i1, V i2) = int_ord (i1, i2)
+  | pattern_ord (p1, p2) = int_ord (int_of_pattern p1, int_of_pattern p2)
+
+
+(* scanning patterns from strings *)
+
+(*
+  The pattern language is cumbersome to use in other structures. Strings are a more
+  lightweight representation. Scanning patterns from strings should be performed at
+  compile time to avoid the parsing overhead at runtime.
+*)
+
+val kind = Scan.many1 Symbol.is_ascii_letter >> (Argo_Expr.kind_of_string o implode)
+val num = Scan.many1 Symbol.is_ascii_digit >> (the o Int.fromString o implode)
+val integer = $$ "-" |-- num >> ~ || num
+val blank = Scan.many1 Symbol.is_ascii_blank >> K ()
+
+fun pattern xs = (
+  kind >> (P o rpair []) ||
+  $$ "!" >> K X ||
+  $$ "(" -- $$ "#" -- blank |-- pattern --| $$ ")" >> M ||
+  $$ "(" -- $$ "?" -- blank |-- num --| $$ ")" >> V ||
+  $$ "(" -- Scan.this_string "num" -- blank |-- integer --| $$ ")" >>
+    (E o Argo_Expr.mk_num o Rat.of_int) ||
+  $$ "(" |-- kind --| blank --| $$ "_" --| $$ ")" >> A ||
+  $$ "(" |-- kind -- Scan.repeat1 (blank |-- pattern) --| $$ ")" >> P) xs
+
+fun scan s = fst (pattern (map str (String.explode s) @ [""]))
+
+
+(* pattern matching *)
+
+exception PATTERN of unit
+
+(*
+  The environment stores the matched expressions for the pattern placeholders.
+  The expression for an indexed placeholder (V i) can be retrieved by "get env i".
+  The argument expressions for an n-ary placeholder (A k) can be retrieved by "get_all env".
+*)
+
+type env = Argo_Expr.expr list Inttab.table
+
+val empty_env: env = Inttab.empty
+fun get_all env = Inttab.lookup_list env 0
+fun get env i = hd (Inttab.lookup_list env i)
+
+fun depth_of (V _) = 0
+  | depth_of (E _) = 0
+  | depth_of (A _) = 1
+  | depth_of (P (_, ps)) = 1 + fold (Integer.max o depth_of) ps 0
+  | depth_of (M p) = depth_of p
+  | depth_of X = 0
+
+fun match_list f k k' env = if k = k' then f env else raise PATTERN ()
+
+fun match (V i) e env = Inttab.update_new (i, [e]) env
+  | match (A k) (Argo_Expr.E (k', es)) env = match_list (Inttab.update_new (0, es)) k k' env
+  | match (P (k, ps)) (Argo_Expr.E (k', es)) env = match_list (fold2 match ps es) k k' env
+  | match _ _ _ = raise Fail "bad pattern"
+
+fun unfold_index env (V i) _ = get env i
+  | unfold_index _ (E e) _ = e
+  | unfold_index env (P (k, ps)) e = Argo_Expr.E (k, map (fn p => unfold_index env p e) ps)
+  | unfold_index _ X e = e
+  | unfold_index _ _ _ = raise Fail "bad pattern"
+
+fun unfold_pattern env (V i) = get env i
+  | unfold_pattern _ (E e) = e
+  | unfold_pattern env (A k) = Argo_Expr.E (k, get_all env)
+  | unfold_pattern env (P (k, [M p])) = Argo_Expr.E (k, map (unfold_index env p) (get_all env))
+  | unfold_pattern env (P (k, ps)) = Argo_Expr.E (k, map (unfold_pattern env) ps)
+  | unfold_pattern _ _ = raise Fail "bad pattern"
+
+
+(* conversions *)
+
+(*
+  Conversions are atomic rewrite steps. For every conversion there is a corresponding
+  inference step.
+*)
+
+type conv = Argo_Expr.expr -> Argo_Expr.expr * Argo_Proof.conv
+
+fun keep e = (e, Argo_Proof.keep_conv)
+
+fun seq [] e = keep e
+  | seq [cv] e = cv e
+  | seq (cv :: cvs) e =
+      let val ((e, c2), c1) = cv e |>> seq cvs
+      in (e, Argo_Proof.mk_then_conv c1 c2) end
+
+fun args cv (Argo_Expr.E (k, es)) =
+  let val (es, cs) = split_list (map cv es)
+  in (Argo_Expr.E (k, es), Argo_Proof.mk_args_conv cs) end
+
+fun rewr r e _ = (e, Argo_Proof.mk_rewr_conv r)
+
+
+(* context *)
+
+(*
+  The context stores functions to flatten expressions and functions to rewrite expressions.
+  Flattening an n-ary expression of kind k produces an expression whose arguments are not
+  of kind k. For instance, flattening (and p (and q r)) produces (and p q r) where p, q and r
+  are not conjunctions.
+*)
+
+structure Pattab = Table(type key = pattern val ord = pattern_ord)
+
+type context = {
+  flats:
+    (Argo_Expr.kind * (int -> Argo_Expr.expr list -> (Argo_Proof.rewrite * Argo_Expr.expr) option))
+      list, (* expressions that should be flattened before rewriting *)
+  rewrs: (env -> (Argo_Proof.rewrite * pattern) option) list Pattab.table}
+    (* Looking up matching rules is O(n). This could be optimized. *)
+
+fun mk_context flats rewrs: context = {flats=flats, rewrs=rewrs}
+val context = mk_context [] Pattab.empty
+
+fun map_context map_flats map_rewrs ({flats, rewrs}: context) =
+  mk_context (map_flats flats) (map_rewrs rewrs)
+
+fun flat lhs f =
+  (case scan lhs of
+    A k => map_context (cons (k, f)) I
+  | _ => raise Fail "bad pattern")
+
+fun func lhs f = map_context I (Pattab.map_default (scan lhs, []) (fn fs => fs @ [f]))
+fun rule r lhs rhs = func lhs (K (SOME (r, scan rhs)))
+
+
+(* rewriting *)
+
+(*
+  Rewriting proceeds bottom-up. Whenever a rewrite rule with placeholders is used,
+  the arguments are again rewritten, but only up to depth of the placeholders within the
+  matched pattern.
+*)
+
+fun rewr_rule r env p = rewr r (unfold_pattern env p)
+
+fun try_apply p e f =
+  let val env = match p e empty_env
+  in (case f env of NONE => NONE | SOME (r, p) => SOME (r, env, p)) end
+  handle PATTERN () => NONE
+
+fun all_args cv k (e as Argo_Expr.E (k', _)) = if k = k' then args (all_args cv k) e else cv e
+fun all_args_of k (e as Argo_Expr.E (k', es)) = if k = k' then maps (all_args_of k) es else [e]
+fun kind_depth_of k (Argo_Expr.E (k', es)) =
+  if k = k' then 1 + fold (Integer.max o kind_depth_of k) es 0 else 0
+
+fun descend cv flats (e as Argo_Expr.E (k, _)) =
+  if AList.defined Argo_Expr.eq_kind flats k then all_args cv k e
+  else args cv e
+
+fun flatten flats (e as Argo_Expr.E (k, _)) =
+  (case AList.lookup Argo_Expr.eq_kind flats k of
+    NONE => keep e
+  | SOME f =>
+      (case f (kind_depth_of k e) (all_args_of k e) of
+        NONE => keep e
+      | SOME (r, e') => rewr r e' e))
+
+fun exhaust cv rewrs e =
+  (case Pattab.get_first (fn (p, fs) => get_first (try_apply p e) fs) rewrs of
+    NONE => keep e
+  | SOME (r, env, p) => seq [rewr_rule r env p, cv (depth_of p)] e)
+
+fun norm (cx as {flats, rewrs}: context) d e =
+  seq [
+    if d = 0 then keep else descend (norm cx (d - 1)) flats,
+    flatten flats,
+    exhaust (norm cx) rewrs] e
+
+fun rewrite cx = norm cx ~1   (* bottom-up rewriting *)
+fun rewrite_top cx = norm cx 0   (* top-down rewriting *)
+
+fun with_proof cv (e, p) prf =
+  let
+    val (e, c) = cv e
+    val (p, prf) = Argo_Proof.mk_rewrite c p prf
+  in ((e, p), prf) end
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_simplex.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,672 @@
+(*  Title:      Tools/Argo/argo_simplex.ML
+    Author:     Sascha Boehme
+
+Linear arithmetic reasoning based on the simplex algorithm. It features:
+
+ * simplification and normalization of arithmetic expressions
+ * decision procedure for reals
+
+These features might be added:
+
+ * propagating implied inequality literals while assuming external knowledge
+ * propagating equalities for fixed variables to all other theory solvers
+ * pruning the tableau after new atoms have been added: eliminate unnecessary
+   variables
+
+The implementation is inspired by:
+
+  Bruno Dutertre and Leonardo de Moura. A fast linear-arithmetic solver
+  for DPLL(T). In Computer Aided Verification, pages 81-94. Springer, 2006.
+*)
+
+signature ARGO_SIMPLEX =
+sig
+  (* simplification *)
+  val simplify: Argo_Rewr.context -> Argo_Rewr.context
+  
+  (* context *)
+  type context
+  val context: context
+
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> Argo_Lit.literal option * context
+
+  (* main operations *)
+  val prepare: context -> context
+  val assume: Argo_Common.literal -> context -> Argo_Lit.literal Argo_Common.implied * context
+  val check: context -> Argo_Lit.literal Argo_Common.implied * context
+  val explain: Argo_Lit.literal -> context -> (Argo_Cls.clause * context) option
+  val add_level: context -> context
+  val backtrack: context -> context
+end
+
+structure Argo_Simplex: ARGO_SIMPLEX =
+struct
+
+(* expression functions *)
+
+fun mk_mul n e =
+  if n = @0 then Argo_Expr.mk_num n
+  else if n = @1 then e
+  else Argo_Expr.mk_mul (Argo_Expr.mk_num n) e
+
+fun dest_mul (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) = (n, e)
+  | dest_mul e = (@1, e)
+
+fun dest_factor e = fst (dest_mul e)
+
+fun mk_add [] = raise Fail "bad addition"
+  | mk_add [e] = e
+  | mk_add es = Argo_Expr.mk_add es
+
+fun dest_ineq (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (false, e1, e2)
+  | dest_ineq (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (true, e1, e2)
+  | dest_ineq _ = NONE
+
+
+(* simplification *)
+
+(*
+  Products are normalized either to a number or to the monomial form
+    a * x
+  where a is a non-zero number and x is a variable. If x is a product,
+  it contains no number factors. The coefficient a is dropped if it is equal
+  to one; instead of 1 * x the expression x is used.
+
+  No further normalization to non-linear products is applied. Hence, the
+  products x * y and y * x will be treated as two different variables by the
+  arithmetic decision procedure.
+*)
+
+val mul_comm_rhs = Argo_Rewr.scan "(mul (? 2) (? 1))"
+
+fun norm_mul env =
+  (case apply2 (Argo_Rewr.get env) (1, 2) of
+    (Argo_Expr.E (Argo_Expr.Num n1, _), Argo_Expr.E (Argo_Expr.Num n2, _)) =>
+      SOME (Argo_Proof.Rewr_Mul_Nums (n1, n2), Argo_Rewr.E (Argo_Expr.mk_num (n1 * n2)))
+  | (Argo_Expr.E (Argo_Expr.Num n, _), _) =>
+      if n = @0 then SOME (Argo_Proof.Rewr_Mul_Zero, Argo_Rewr.V 1)
+      else if n = @1 then SOME (Argo_Proof.Rewr_Mul_One, Argo_Rewr.V 2)
+      else NONE
+  | (_, Argo_Expr.E (Argo_Expr.Num _, _)) => SOME (Argo_Proof.Rewr_Mul_Comm, mul_comm_rhs)
+  | _ => NONE)
+
+(*
+  Quotients are normalized either to a number or to the monomial form
+    a * x
+  where a is a non-zero number and x is a variable. If x is a quotient,
+  both divident and divisor are not a number. The coefficient a is dropped
+  if it is equal to one; instead of 1 * x the expression x is used.
+
+  No further normalization to quotients is applied. Hence, the
+  expressions (x * z) / y and x * (z / y) will be treated as two different
+  variables by the arithmetic decision procedure.
+*)
+
+fun div_mul_rhs n e =
+  let val pat = Argo_Rewr.P (Argo_Expr.Div, [Argo_Rewr.E e, Argo_Rewr.V 2])
+  in Argo_Rewr.P (Argo_Expr.Mul, [Argo_Rewr.E (Argo_Expr.mk_num n), pat])  end
+
+fun div_inv_rhs n pat =
+  let val inv_pat = Argo_Rewr.P (Argo_Expr.Div, map (Argo_Rewr.E o Argo_Expr.mk_num) [@1, n]) 
+  in Argo_Rewr.P (Argo_Expr.Mul, [inv_pat, pat]) end
+
+fun norm_div env =
+  (case apply2 (Argo_Rewr.get env) (1, 2) of
+    (Argo_Expr.E (Argo_Expr.Num n1, _), Argo_Expr.E (Argo_Expr.Num n2, _)) =>
+      if n2 = @0 then NONE
+      else SOME (Argo_Proof.Rewr_Div_Nums (n1, n2), Argo_Rewr.E (Argo_Expr.mk_num (n1 / n2)))
+  | (Argo_Expr.E (Argo_Expr.Num n, _), _) =>
+      if n = @0 then SOME (Argo_Proof.Rewr_Div_Zero, Argo_Rewr.E (Argo_Expr.mk_num @0))
+      else if n = @1 then NONE
+      else SOME (Argo_Proof.Rewr_Div_Mul, div_mul_rhs n (Argo_Expr.mk_num @1))
+  | (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e]), _) =>
+      if n = @1 then NONE
+      else SOME (Argo_Proof.Rewr_Div_Left, div_mul_rhs n e)
+  | (_, Argo_Expr.E (Argo_Expr.Num n, _)) =>
+      if n = @0 then NONE
+      else if n = @1 then SOME (Argo_Proof.Rewr_Div_One, Argo_Rewr.V 1)
+      else SOME (Argo_Proof.Rewr_Div_Inv, div_inv_rhs n (Argo_Rewr.V 1))
+  | (_, Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) =>
+      let val pat = Argo_Rewr.P (Argo_Expr.Div, [Argo_Rewr.V 1, Argo_Rewr.E e])
+      in SOME (Argo_Proof.Rewr_Div_Right, div_inv_rhs n pat) end
+  | _ => NONE)
+
+(*
+  Sums are flatten and normalized to the polynomial form
+    a_0 + a_1 * x_1 + ... + a_n * x_n
+  where all variables x_i are disjoint and where all coefficients a_i are
+  non-zero numbers. Coefficients equal to one are dropped; instead of 1 * x_i
+  the reduced monom x_i is used. The variables x_i are ordered based on the
+  expression order to reduce the number of problem literals by sharing equal
+  expressions.
+*)
+
+fun add_monom_expr i n e (p, s, etab) =
+  let val etab = Argo_Exprtab.map_default (e, (i, @0)) (apsnd (Rat.add n)) etab
+  in ((n, Option.map fst (Argo_Exprtab.lookup etab e)) :: p, s, etab) end
+
+fun add_monom (_, Argo_Expr.E (Argo_Expr.Num n, _)) (p, s, etab) = ((n, NONE) :: p, s + n, etab)
+  | add_monom (i, Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) x =
+      add_monom_expr i n e x
+  | add_monom (i, e) x = add_monom_expr i @1 e x
+
+fun norm_add d es =
+  let
+    val (p1, s, etab) = fold_index add_monom es ([], @0, Argo_Exprtab.empty)
+    val (p2, es) =
+      []
+      |> Argo_Exprtab.fold_rev (fn (e, (i, n)) => n <> @0 ? cons ((n, SOME i), mk_mul n e)) etab
+      |> s <> @0 ? cons ((s, NONE), Argo_Expr.mk_num s)
+      |> (fn [] => [((@0, NONE), Argo_Expr.mk_num @0)] | xs => xs)
+      |> split_list
+    val ps = (rev p1, p2)
+  in
+    if d = 1 andalso eq_list (op =) ps then NONE
+    else SOME (Argo_Proof.Rewr_Add ps, mk_add es)
+  end
+
+(*
+  An equation between two arithmetic expressions is rewritten to a conjunction of two
+  non-strict inequalities.
+*)
+
+val eq_rhs = Argo_Rewr.scan "(and (le (? 1) (? 2)) (le (? 2) (? 1)))"
+fun is_arith e = member (op =) [Argo_Expr.Real] (Argo_Expr.type_of e)
+
+fun norm_eq env =
+  if is_arith (Argo_Rewr.get env 1) then SOME (Argo_Proof.Rewr_Eq_Le, eq_rhs)
+  else NONE
+
+(*
+  Arithmetic inequalities (less and less-than) are normalized to the normal form
+    a_0 + a_1 * x_1 + ... + a_n * x_n ~ b
+  or
+    b ~ a_0 + a_1 * x_1 + ... + a_n * x_n
+  such that most of the coefficients a_i are positive.
+
+  Arithmetic inequalities of the form
+    a * x ~ b
+  or
+    b ~ a * x
+  are normalized to the form
+    x ~ c
+  or
+    c ~ x
+  where c is a number.
+*)
+
+fun mk_num_pat n = Argo_Rewr.E (Argo_Expr.mk_num n)
+fun mk_cmp_pat k ps = Argo_Rewr.P (k, ps)
+
+fun norm_cmp_mul k r n =
+  let
+    fun mult i = Argo_Rewr.P (Argo_Expr.Mul, [mk_num_pat n, Argo_Rewr.V i])
+    val ps = if n > @0 then [mult 1, mult 2] else [mult 2, mult 1]
+  in SOME (Argo_Proof.Rewr_Ineq_Mul (r, n), mk_cmp_pat k ps) end
+
+fun count_factors pred es = fold (fn e => if pred (dest_factor e) then Integer.add 1 else I) es 0
+
+fun norm_cmp_swap k r es =
+  let
+    val pos = count_factors (fn n => n > @0) es
+    val neg = count_factors (fn n => n < @0) es
+    val keep = pos > neg orelse (pos = neg andalso dest_factor (hd es) > @0)
+  in if keep then NONE else norm_cmp_mul k r @~1 end
+
+fun norm_cmp1 k r (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), _])) =
+      norm_cmp_mul k r (Rat.inv n)
+  | norm_cmp1 k r (Argo_Expr.E (Argo_Expr.Add, Argo_Expr.E (Argo_Expr.Num n, _) :: _)) =
+      let fun mk i = Argo_Rewr.P (Argo_Expr.Add, [Argo_Rewr.V i, mk_num_pat (~ n)])
+      in SOME (Argo_Proof.Rewr_Ineq_Add (r, ~ n), mk_cmp_pat k [mk 1, mk 2]) end
+  | norm_cmp1 k r (Argo_Expr.E (Argo_Expr.Add, es)) = norm_cmp_swap k r es
+  | norm_cmp1 _ _ _ = NONE
+
+val nums_true_rhs = Argo_Rewr.scan "true"
+val nums_false_rhs = Argo_Rewr.scan "false"
+
+val cmp_sub_rhs = map Argo_Rewr.scan ["(sub (? 1) (? 2))", "(num 0)"]
+
+fun norm_cmp k r pred env =
+  (case apply2 (Argo_Rewr.get env) (1, 2) of
+    (Argo_Expr.E (Argo_Expr.Num n1, _), Argo_Expr.E (Argo_Expr.Num n2, _)) =>
+      let val b = pred n1 n2
+      in SOME (Argo_Proof.Rewr_Ineq_Nums (r, b), if b then nums_true_rhs else nums_false_rhs) end
+  | (Argo_Expr.E (Argo_Expr.Num _, _), e) => norm_cmp1 k r e
+  | (e, Argo_Expr.E (Argo_Expr.Num _, _)) => norm_cmp1 k r e
+  | _ => SOME (Argo_Proof.Rewr_Ineq_Sub r, mk_cmp_pat k cmp_sub_rhs))
+
+(*
+  Arithmetic expressions are normalized in order to reduce the number of
+  problem literals. Arithmetically equal expressions are normalized into
+  syntactically equal expressions as much as possible.
+*)
+
+val simplify =
+  Argo_Rewr.rule Argo_Proof.Rewr_Neg
+    "(neg (? 1))"
+    "(mul (num -1) (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Sub
+    "(sub (? 1) (? 2))"
+    "(add (? 1) (mul (num -1) (? 2)))" #>
+  Argo_Rewr.func "(mul (? 1) (? 2))" norm_mul #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Mul_Assoc
+    "(mul (? 1) (mul (? 2) (? 3)))"
+    "(mul (mul (? 1) (? 2)) (? 3))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Mul_Sum
+    "(mul (? 1) (add _))"
+    "(add (# (mul (? 1) !)))" #>
+  Argo_Rewr.func "(div (? 1) (? 2))" norm_div #>
+  Argo_Rewr.flat "(add _)" norm_add #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Min
+    "(min (? 1) (? 2))"
+    "(ite (le (? 1) (? 2)) (? 1) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Max
+    "(max (? 1) (? 2))"
+    "(ite (le (? 1) (? 2)) (? 2) (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Abs
+    "(abs (? 1))"
+    "(ite (le (num 0) (? 1)) (? 1) (neg (? 1)))" #>
+  Argo_Rewr.func "(eq (? 1) (? 2))" norm_eq #>
+  Argo_Rewr.func "(le (? 1) (? 2))" (norm_cmp Argo_Expr.Le Argo_Proof.Le Rat.le) #>
+  Argo_Rewr.func "(lt (? 1) (? 2))" (norm_cmp Argo_Expr.Lt Argo_Proof.Lt Rat.lt)
+
+
+(* extended rationals *)
+
+(*
+  Extended rationals (c, k) are reals (c + k * e) where e is some small positive real number.
+  Extended rationals are used to represent a strict inequality by a non-strict inequality:
+    c < x  ~~  c + k * e <= e
+    x < c  ~~  x <= c - k * e
+*)
+
+type erat = Rat.rat * Rat.rat
+
+val erat_zero = (@0, @0)
+
+fun add (c1, k1) (c2, k2) = (c1 + c2, k1 + k2)
+fun sub (c1, k1) (c2, k2) = (c1 - c2, k1 - k2)
+fun mul n (c, k) = (n * c, n * k)
+
+val erat_ord = prod_ord Rat.ord Rat.ord
+
+fun less_eq n1 n2 = (erat_ord (n1, n2) <> GREATER)
+fun less n1 n2 = (erat_ord (n1, n2) = LESS)
+
+
+(* term functions *)
+
+fun dest_monom (Argo_Term.T (_, Argo_Expr.Mul, [Argo_Term.T (_, Argo_Expr.Num n, _), t])) = (t, n)
+  | dest_monom t = (t, @1)
+
+datatype node = Var of Argo_Term.term | Num of Rat.rat
+datatype ineq = Lower of Argo_Term.term * erat | Upper of Argo_Term.term * erat
+
+fun dest_node (Argo_Term.T (_, Argo_Expr.Num n, _)) = Num n
+  | dest_node t = Var t
+
+fun dest_atom true (k as Argo_Expr.Le) t1 t2 = SOME (k, dest_node t1, dest_node t2)
+  | dest_atom true (k as Argo_Expr.Lt) t1 t2 = SOME (k, dest_node t1, dest_node t2)
+  | dest_atom false Argo_Expr.Le t1 t2 = SOME (Argo_Expr.Lt, dest_node t2, dest_node t1)
+  | dest_atom false Argo_Expr.Lt t1 t2 = SOME (Argo_Expr.Le, dest_node t2, dest_node t1)
+  | dest_atom _ _ _ _ = NONE
+
+fun ineq_of pol (Argo_Term.T (_, k, [t1, t2])) =
+      (case dest_atom pol k t1 t2 of
+        SOME (Argo_Expr.Le, Var x, Num n) => SOME (Upper (x, (n, @0)))
+      | SOME (Argo_Expr.Le, Num n, Var x) => SOME (Lower (x, (n, @0)))
+      | SOME (Argo_Expr.Lt, Var x, Num n) => SOME (Upper (x, (n, @~1)))
+      | SOME (Argo_Expr.Lt, Num n, Var x) => SOME (Lower (x, (n, @1)))
+      | _ => NONE)
+  | ineq_of _ _ = NONE
+
+
+(* proofs *)
+
+(*
+  comment missing
+*)
+
+fun mk_ineq is_lt = if is_lt then Argo_Expr.mk_lt else Argo_Expr.mk_le
+fun ineq_rule_of is_lt = if is_lt then Argo_Proof.Lt else Argo_Proof.Le
+
+fun rewrite_top f = Argo_Rewr.rewrite_top (f Argo_Rewr.context)
+
+fun unnegate_conv (e as Argo_Expr.E (Argo_Expr.Not, [Argo_Expr.E (Argo_Expr.Le, [e1, e2])])) =
+      Argo_Rewr.rewr (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) (Argo_Expr.mk_lt e2 e1) e
+  | unnegate_conv (e as Argo_Expr.E (Argo_Expr.Not, [Argo_Expr.E (Argo_Expr.Lt, [e1, e2])])) =
+      Argo_Rewr.rewr (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) (Argo_Expr.mk_le e2 e1) e
+  | unnegate_conv e = Argo_Rewr.keep e
+
+val norm_scale_conv = rewrite_top (
+  Argo_Rewr.rule Argo_Proof.Rewr_Mul_Sum
+    "(mul (? 1) (add _))"
+    "(add (# (mul (? 1) !)))" #>
+  Argo_Rewr.func "(mul (? 1) (? 2))" norm_mul)
+
+fun scale_conv r mk n e1 e2 =
+  let
+    fun scale e = Argo_Expr.mk_mul (Argo_Expr.mk_num n) e
+    val (e1, e2) = if n > @0 then (scale e1, scale e2) else (scale e2, scale e1)
+    val conv = Argo_Rewr.rewr (Argo_Proof.Rewr_Ineq_Mul (r, n)) (mk e1 e2)
+  in Argo_Rewr.seq [conv, Argo_Rewr.args norm_scale_conv] end
+
+fun scale_ineq_conv n e =
+  if n = @1 then Argo_Rewr.keep e
+  else
+    (case dest_ineq e of
+      NONE => raise Fail "bad inequality"
+    | SOME (is_lt, e1, e2) => scale_conv (ineq_rule_of is_lt) (mk_ineq is_lt) n e1 e2 e)
+
+fun simp_lit (n, (lit, p)) =
+  let val conv = Argo_Rewr.seq [unnegate_conv, scale_ineq_conv n]
+  in Argo_Rewr.with_proof conv (Argo_Lit.signed_expr_of lit, p) end
+
+val combine_conv = rewrite_top (Argo_Rewr.flat "(add _)" norm_add)
+fun reduce_conv r = Argo_Rewr.rewr (Argo_Proof.Rewr_Ineq_Nums (r, false)) Argo_Expr.false_expr
+
+fun simp_combine es p prf =
+  let
+    fun dest e (is_lt, (es1, es2)) =
+      let val (is_lt', e1, e2) = the (dest_ineq e)
+      in (is_lt orelse is_lt', (e1 :: es1, e2 :: es2)) end
+    val (is_lt, (es1, es2)) = fold_rev dest es (false, ([], []))
+    val e = uncurry (mk_ineq is_lt) (apply2 Argo_Expr.mk_add (es1, es2))
+    val conv = Argo_Rewr.seq [Argo_Rewr.args combine_conv, reduce_conv (ineq_rule_of is_lt)]
+  in prf |> Argo_Rewr.with_proof conv (e, p) |>> snd end
+
+fun linear_combination nlps prf =
+  let val ((es, ps), prf) = fold_map simp_lit nlps prf |>> split_list
+  in prf |> Argo_Proof.mk_linear_comb ps |-> simp_combine es |-> Argo_Proof.mk_lemma [] end
+
+fun proof_of (lit, SOME p) (ls, prf) = ((lit, p), (ls, prf))
+  | proof_of (lit, NONE) (ls, prf) =
+      let val (p, prf) = Argo_Proof.mk_hyp lit prf
+      in ((lit, p), (Argo_Lit.negate lit :: ls, prf)) end
+
+
+(* tableau *)
+
+(*
+  The tableau consists of equations x_i = a_i1 * x_1 + ... a_ik * x_k where
+  the variable on the left-hand side is called a basic variable and
+  the variables on the right-hand side are called non-basic variables.
+
+  For each basic variable, the polynom on the right-hand side is stored as a map
+  from variables to coefficients. Only variables with non-zero coefficients are stored.
+  The map is sorted by the term order of the variables for a deterministic order when
+  analyzing a polynom.
+
+  Additionally, for each basic variable a boolean flag is kept that, when false,
+  indicates that the current value of the basic variable might be outside its bounds.
+  The value of a non-basic variable is always within its bounds.
+
+  The tableau is stored as a table indexed by variables. For each variable,
+  both basic and non-basic, its current value is stored as extended rational
+  along with either the equations or the occurrences.
+*)
+
+type basic = bool * (Argo_Term.term * Rat.rat) Ord_List.T
+type entry = erat * basic option
+type tableau = entry Argo_Termtab.table
+
+fun dirty ms = SOME (false, ms)
+fun checked ms = SOME (true, ms)
+
+fun basic_entry ms = (erat_zero, dirty ms)
+val non_basic_entry: entry = (erat_zero, NONE)
+
+fun value_of tableau x =
+  (case Argo_Termtab.lookup tableau x of
+    NONE => erat_zero
+  | SOME (v, _) => v)
+
+fun first_unchecked_basic tableau =
+  Argo_Termtab.get_first (fn (y, (v, SOME (false, ms))) => SOME (y, v, ms) | _ => NONE) tableau
+
+local
+
+fun coeff_of ms x = the (AList.lookup Argo_Term.eq_term ms x)
+
+val eq_var = Argo_Term.eq_term
+fun monom_ord sp = prod_ord Argo_Term.term_ord (K EQUAL) sp
+
+fun add_monom m ms = Ord_List.insert monom_ord m ms
+fun update_monom (m as (x, a)) = if a = @0 then AList.delete eq_var x else AList.update eq_var m
+
+fun add_scaled_monom n (x, a) ms =
+  (case AList.lookup eq_var ms x of
+    NONE => add_monom (x, n * a) ms
+  | SOME b => update_monom (x, n * a + b) ms)
+
+fun replace_polynom x n ms' ms = fold (add_scaled_monom n) ms' (AList.delete eq_var x ms)
+
+fun map_basic f (v, SOME (_, ms)) = f v ms
+  | map_basic _ e = e
+
+fun map_basic_entries x f =
+  let
+    fun apply (e as (v, SOME (_, ms))) = if AList.defined eq_var ms x then f v ms else e
+      | apply ve = ve
+  in Argo_Termtab.map (K apply) end
+
+fun put_entry x e = Argo_Termtab.update (x, e)
+
+fun add_new_entry (y as Argo_Term.T (_, Argo_Expr.Add, ts)) tableau =
+      let val ms = Ord_List.make monom_ord (map dest_monom ts)
+      in fold (fn (x, _) => put_entry x non_basic_entry) ms (put_entry y (basic_entry ms) tableau) end
+  | add_new_entry x tableau = put_entry x non_basic_entry tableau
+
+fun with_non_basic update_basic x f tableau =
+  (case Argo_Termtab.lookup tableau x of
+    NONE => tableau
+  | SOME (v, NONE) => f v tableau
+  | SOME (v, SOME (_, ms)) => if update_basic then put_entry x (v, dirty ms) tableau else tableau)
+
+in
+
+fun add_entry x tableau =
+  if Argo_Termtab.defined tableau x then tableau
+  else add_new_entry x tableau
+
+fun basic_within_bounds y = Argo_Termtab.map_entry y (map_basic (fn v => fn ms => (v, checked ms)))
+
+fun eliminate _ tableau = tableau
+
+fun update_non_basic pred x v' = with_non_basic true x (fn v =>
+  let fun update_basic n v ms = (add v (mul (coeff_of ms x) n), dirty ms)
+  in pred v ? put_entry x (v', NONE) o map_basic_entries x (update_basic (sub v' v)) end)
+
+fun update_pivot y vy ms x c v = with_non_basic false x (fn vx =>
+  let
+    val a = Rat.inv c
+    val v' = mul a (sub v vy)
+
+    fun scale_or_drop (x', b) = if Argo_Term.eq_term (x', x) then NONE else SOME (x', ~ a * b)
+    val ms = add_monom (y, a) (map_filter scale_or_drop ms)
+
+    fun update_basic v ms' =
+      let val n = coeff_of ms' x
+      in (add v (mul n v'), dirty (replace_polynom x n ms ms')) end
+  in
+    put_entry x (add vx v', dirty ms) #>
+    put_entry y (v, NONE) #>
+    map_basic_entries x update_basic
+  end)
+
+end
+
+
+(* bounds *)
+
+(*
+  comment missing
+*)
+
+type bound = (erat * Argo_Common.literal) option
+type atoms = (erat * Argo_Term.term) list
+type bounds_atoms = ((bound * bound) * (atoms * atoms))
+type bounds = bounds_atoms Argo_Termtab.table
+
+val empty_bounds_atoms: bounds_atoms = ((NONE, NONE), ([], []))
+
+fun on_some pred (SOME (n, _)) = pred n
+  | on_some _ NONE = false
+
+fun none_or_some pred (SOME (n, _)) = pred n
+  | none_or_some _ NONE = true
+
+fun bound_of (SOME (n, _)) = n
+  | bound_of NONE = raise Fail "bad bound"
+
+fun reason_of (SOME (_, r)) = r
+  | reason_of NONE = raise Fail "bad reason"
+
+fun bounds_atoms_of bounds x = the_default empty_bounds_atoms (Argo_Termtab.lookup bounds x)
+fun bounds_of bounds x = fst (bounds_atoms_of bounds x)
+
+fun put_bounds x bs bounds = Argo_Termtab.map_default (x, empty_bounds_atoms) (apfst (K bs)) bounds
+
+fun has_bound_atoms bounds x =
+  (case Argo_Termtab.lookup bounds x of
+    NONE => false
+  | SOME (_, ([], [])) => false
+  | _ => true)
+
+fun add_new_atom f x n t =
+  let val ins = f (insert (eq_snd Argo_Term.eq_term) (n, t))
+  in Argo_Termtab.map_default (x, empty_bounds_atoms) (apsnd ins) end
+
+fun del_atom x t =
+  let fun eq_atom (t1, (_, t2)) = Argo_Term.eq_term (t1, t2)
+  in Argo_Termtab.map_entry x (apsnd (apply2 (remove eq_atom t))) end
+
+
+(* context *)
+
+type context = {
+  tableau: tableau, (* values of variables and tableau entries for each variable *)
+  bounds: bounds, (* bounds and unassigned atoms for each variable *)
+  prf: Argo_Proof.context, (* proof context *)
+  back: bounds list} (* stack storing previous bounds and unassigned atoms *)
+
+fun mk_context tableau bounds prf back: context =
+  {tableau=tableau, bounds=bounds, prf=prf, back=back}
+
+val context = mk_context Argo_Termtab.empty Argo_Termtab.empty Argo_Proof.simplex_context []
+
+
+(* declaring atoms *)
+
+fun add_ineq_atom f t x n ({tableau, bounds, prf, back}: context) =
+  (* TODO: check whether the atom is already known to hold *)
+  (NONE, mk_context (add_entry x tableau) (add_new_atom f x n t bounds) prf back)
+
+fun add_atom t cx =
+  (case ineq_of true t of
+    SOME (Lower (x, n)) => add_ineq_atom apfst t x n cx
+  | SOME (Upper (x, n)) => add_ineq_atom apsnd t x n cx
+  | NONE => (NONE, cx))
+
+
+(* preparing the solver after new atoms have been added *)
+
+(*
+  Variables that do not directly occur in atoms can be eliminated from the tableau
+  since no bounds will ever limit their value. This can reduce the tableau size
+  substantially.
+*)
+
+fun prepare ({tableau, bounds, prf, back}: context) =
+  let fun drop (xe as (x, _)) = not (has_bound_atoms bounds x) ? eliminate xe
+  in mk_context (Argo_Termtab.fold drop tableau tableau) bounds prf back end
+
+
+(* assuming external knowledge *)
+
+fun bounds_conflict r1 r2 ({tableau, bounds, prf, back}: context) =
+  let
+    val ((lp2, lp1), (lits, prf)) = ([], prf) |> proof_of r2 ||>> proof_of r1
+    val (p, prf) = linear_combination [(@~1, lp1), (@1, lp2)] prf
+  in (Argo_Common.Conflict (lits, p), mk_context tableau bounds prf back) end
+
+fun assume_bounds order x c bs ({tableau, bounds, prf, back}: context) =
+  let
+    val lits = []
+    val bounds = put_bounds x bs bounds
+    val tableau = update_non_basic (fn v => erat_ord (v, c) = order) x c tableau
+  in (Argo_Common.Implied lits, mk_context tableau bounds prf back) end
+
+fun assume_lower r x c (low, upp) cx =
+  if on_some (fn l => less_eq c l) low then (Argo_Common.Implied [], cx)
+  else if on_some (fn u => less u c) upp then bounds_conflict r (reason_of upp) cx
+  else assume_bounds LESS x c (SOME (c, r), upp) cx
+
+fun assume_upper r x c (low, upp) cx =
+  if on_some (fn u => less_eq u c) upp then (Argo_Common.Implied [], cx)
+  else if on_some (fn l => less c l) low then bounds_conflict (reason_of low) r cx
+  else assume_bounds GREATER x c (low, SOME (c, r)) cx
+
+fun with_bounds r t f x n ({tableau, bounds, prf, back}: context) =
+  f r x n (bounds_of bounds x) (mk_context tableau (del_atom x t bounds) prf back)
+
+fun choose f (SOME (Lower (x, n))) cx = f assume_lower x n cx
+  | choose f (SOME (Upper (x, n))) cx = f assume_upper x n cx
+  | choose _ NONE cx = (Argo_Common.Implied [], cx)
+
+fun assume (r as (lit, _)) cx =
+  let val (t, pol) = Argo_Lit.dest lit
+  in choose (with_bounds r t) (ineq_of pol t) cx end
+
+
+(* checking for consistency and pending implications *)
+
+fun basic_bounds_conflict lower y ms ({tableau, bounds, prf, back}: context) =
+  let
+    val (a, low, upp) = if lower then (@1, fst, snd) else (@~1, snd, fst)
+    fun coeff_proof f a x = apfst (pair a) o proof_of (reason_of (f (bounds_of bounds x)))
+    fun monom_proof (x, a) = coeff_proof (if a < @0 then low else upp) a x
+    val ((alp, alps), (lits, prf)) = ([], prf) |> coeff_proof low a y ||>> fold_map monom_proof ms
+    val (p, prf) = linear_combination (alp :: alps) prf
+  in (Argo_Common.Conflict (lits, p), mk_context tableau bounds prf back) end
+
+fun can_compensate ord tableau bounds (x, a) =
+  let val (low, upp) = bounds_of bounds x
+  in
+    if Rat.ord (a, @0) = ord then none_or_some (fn u => less (value_of tableau x) u) upp
+    else none_or_some (fn l => less l (value_of tableau x)) low
+  end
+
+fun check (cx as {tableau, bounds, prf, back}: context) =
+  (case first_unchecked_basic tableau of
+    NONE => (Argo_Common.Implied [], cx)
+  | SOME (y, v, ms) =>
+      let val (low, upp) = bounds_of bounds y
+      in
+        if on_some (fn l => less v l) low then adjust GREATER true y v ms (bound_of low) cx
+        else if on_some (fn u => less u v) upp then adjust LESS false y v ms (bound_of upp) cx
+        else check (mk_context (basic_within_bounds y tableau) bounds prf back)
+      end)
+
+and adjust ord lower y vy ms v (cx as {tableau, bounds, prf, back}: context) =
+  (case find_first (can_compensate ord tableau bounds) ms of
+    NONE => basic_bounds_conflict lower y ms cx
+  | SOME (x, a) => check (mk_context (update_pivot y vy ms x a v tableau) bounds prf back))
+
+
+(* explanations *)
+
+fun explain _ _ = NONE
+
+
+(* backtracking *)
+
+fun add_level ({tableau, bounds, prf, back}: context) =
+  mk_context tableau bounds prf (bounds :: back)
+
+fun backtrack ({back=[], ...}: context) = raise Empty
+  | backtrack ({tableau, prf, back=bounds :: back, ...}: context) =
+      mk_context tableau bounds prf back
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_solver.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,168 @@
+(*  Title:      Tools/Argo/argo_solver.ML
+    Author:     Sascha Boehme
+
+The main interface to the Argo solver.
+
+The solver performs satisfiability checking for a given set of assertions. If these assertions
+are unsatisfiable, a proof trace is returned. If these assertions are satisfiable, the computed
+model can be queried or further assertions may be added.
+*)
+
+signature ARGO_SOLVER =
+sig
+  type context
+  val context: context
+  val assert: Argo_Expr.expr list -> context -> context (* raises Argo_Expr.TYPE, Argo_Expr.EXPR
+    and Argo_Proof.UNSAT *)
+  val model_of: context -> string * Argo_Expr.typ -> bool option
+end
+
+structure Argo_Solver: ARGO_SOLVER =
+struct
+
+(* context *)
+
+type context = {
+  next_axiom: int,
+  prf: Argo_Proof.context,
+  core: Argo_Core.context}
+
+fun mk_context next_axiom prf core: context = {next_axiom=next_axiom, prf=prf, core=core}
+
+val context = mk_context 0 Argo_Proof.solver_context Argo_Core.context
+
+
+(* negation normal form *)
+
+fun nnf_nary r rhs env = SOME (r (length (Argo_Rewr.get_all env)), rhs)
+
+val not_and_rhs = Argo_Rewr.scan "(or (# (not !)))"
+val not_or_rhs = Argo_Rewr.scan "(and (# (not !)))"
+
+val nnf =
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_True "(not true)" "false" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_False "(not false)" "true" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Not "(not (not (? 1)))" "(? 1)" #>
+  Argo_Rewr.func "(not (and _))" (nnf_nary Argo_Proof.Rewr_Not_And not_and_rhs) #>
+  Argo_Rewr.func "(not (or _))" (nnf_nary Argo_Proof.Rewr_Not_Or not_or_rhs) #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Iff "(not (iff (not (? 1)) (? 2)))" "(iff (? 1) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Iff "(not (iff (? 1) (not (? 2))))" "(iff (? 1) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Not_Iff "(not (iff (? 1) (? 2)))" "(iff (not (? 1)) (? 2))"
+
+
+(* propositional normalization *)
+
+(*
+  Propositional expressions are transformed into literals in the clausifier. Having
+  fewer literals results in faster solver execution. Normalizing propositional expressions
+  turns similar expressions into equal expressions, for which the same literal can be used.
+  The clausifier expects that only negation, disjunction, conjunction and equivalence form
+  propositional expressions. Expressions may be simplified to truth or falsity, but both
+  truth and falsity eventually occur nowhere inside expressions.
+*)
+
+val e_true = Argo_Expr.true_expr
+val e_false = Argo_Expr.false_expr
+
+fun first_index pred xs =
+  let val i = find_index pred xs
+  in if i >= 0 then SOME i else NONE end
+
+fun find_zero r_zero zero es =
+  Option.map (fn i => (r_zero i, zero)) (first_index (curry Argo_Expr.eq_expr zero) es)
+
+fun find_duals _ _ _ [] = NONE
+  | find_duals _ _ _ [_] = NONE
+  | find_duals r_dual zero i (e :: es) =
+      (case first_index (Argo_Expr.dual_expr e) es of
+        SOME i' => SOME (r_dual (i, i + i' + 1), zero)
+      | NONE => find_duals r_dual zero (i + 1) es)
+
+fun sort_nary r_sort one mk l es =
+  let
+    val n = length es
+    fun add (i, e) = if Argo_Expr.eq_expr (e, one) then I else Argo_Exprtab.cons_list (e, i)
+    fun dest (e, i) (es, is) = (e :: es, i :: is)
+    val (es, iss) = Argo_Exprtab.fold_rev dest (fold_index add es Argo_Exprtab.empty) ([], [])
+    fun identity is = length is = n andalso forall (op =) (map_index I is)
+  in if l = 1 andalso identity (map hd iss) then NONE else (SOME (r_sort (n, iss), mk es)) end
+
+fun apply_first fs es = get_first (fn f => f es) fs
+
+fun norm_nary r_zero r_dual r_sort zero one mk l =
+  apply_first [find_zero r_zero zero, find_duals r_dual zero 0, sort_nary r_sort one mk l]
+
+val norm_and = norm_nary Argo_Proof.Rewr_And_False Argo_Proof.Rewr_And_Dual Argo_Proof.Rewr_And_Sort
+  e_false e_true Argo_Expr.mk_and
+val norm_or = norm_nary Argo_Proof.Rewr_Or_True Argo_Proof.Rewr_Or_Dual Argo_Proof.Rewr_Or_Sort
+  e_true e_false Argo_Expr.mk_or
+
+fun norm_iff env =
+  let val e1 = Argo_Rewr.get env 1 and e2 = Argo_Rewr.get env 2
+  in
+    if Argo_Expr.dual_expr e1 e2 then SOME (Argo_Proof.Rewr_Iff_Dual, Argo_Rewr.E e_false)
+    else
+      (case Argo_Expr.expr_ord (e1, e2) of
+        EQUAL => SOME (Argo_Proof.Rewr_Iff_Refl, Argo_Rewr.E e_true)
+      | LESS => NONE
+      | GREATER => SOME (Argo_Proof.Rewr_Iff_Symm, Argo_Rewr.E (Argo_Expr.mk_iff e2 e1)))
+  end
+
+val norm_prop =
+  Argo_Rewr.flat "(and _)" norm_and #>
+  Argo_Rewr.flat "(or _)" norm_or #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Imp "(imp (? 1) (? 2))" "(or (not (? 1)) (? 2))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_True "(iff true (? 1))" "(? 1)" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_False "(iff false (? 1))" "(not (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_True "(iff (? 1) true)" "(? 1)" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_False "(iff (? 1) false)" "(not (? 1))" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Iff_Not_Not "(iff (not (? 1)) (not (? 2)))" "(iff (? 1) (? 2))" #>
+  Argo_Rewr.func "(iff (? 1) (? 2))" norm_iff
+
+
+(* normalization of if-then-else expressions *)
+
+val simp_prop_ite_result =
+  Argo_Rewr.scan "(and (or (not (? 1)) (? 2)) (or (? 1) (? 3)) (or (? 2) (? 3)))"
+
+val simp_ite_eq_result = Argo_Rewr.scan "(? 2)"
+
+fun simp_ite env =
+  if Argo_Expr.type_of (Argo_Rewr.get env 2) = Argo_Expr.Bool then
+    SOME (Argo_Proof.Rewr_Ite_Prop, simp_prop_ite_result)
+  else if Argo_Expr.eq_expr (Argo_Rewr.get env 2, Argo_Rewr.get env 3) then
+    SOME (Argo_Proof.Rewr_Ite_Eq, simp_ite_eq_result)
+  else NONE
+
+val norm_ite =
+  Argo_Rewr.rule Argo_Proof.Rewr_Ite_True "(ite true (? 1) (? 2))" "(? 1)" #>
+  Argo_Rewr.rule Argo_Proof.Rewr_Ite_False "(ite false (? 1) (? 2))" "(? 2)" #>
+  Argo_Rewr.func "(ite (? 1) (? 2) (? 3))" simp_ite
+
+
+(* rewriting and normalizing axioms *)
+
+val simp_context = Argo_Rewr.context |> nnf |> norm_prop |> norm_ite |> Argo_Thy.simplify
+val simp_axiom = Argo_Rewr.with_proof (Argo_Rewr.rewrite simp_context)
+
+
+(* asserting axioms *)
+
+fun add_axiom e ({next_axiom, prf, core}: context) =
+  let
+    val _ = Argo_Expr.check e
+    val (p, prf) = Argo_Proof.mk_axiom next_axiom prf
+    val (ep, prf) = simp_axiom (e, p) prf 
+    val (prf, core) = Argo_Clausify.clausify simp_context ep (prf, core)
+  in mk_context (next_axiom + 1) prf core end
+
+fun assert es cx =
+  let val {next_axiom, prf, core}: context = fold add_axiom es cx
+  in mk_context next_axiom prf (Argo_Core.run core) end
+
+
+(* models *)
+
+fun model_of ({core, ...}: context) = Argo_Core.model_of core
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_term.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,152 @@
+(*  Title:      Tools/Argo/argo_term.ML
+    Author:     Sascha Boehme
+
+Internal language of the Argo solver.
+
+Terms are fully-shared via hash-consing. Alpha-equivalent terms have the same identifier.
+*)
+
+signature ARGO_TERM =
+sig
+  (* data types *)
+  type meta
+  datatype term = T of meta * Argo_Expr.kind * term list
+
+  (* term operations *)
+  val id_of: term -> int
+  val expr_of: term -> Argo_Expr.expr
+  val type_of: term -> Argo_Expr.typ
+  val eq_term: term * term -> bool
+  val term_ord: term * term -> order
+
+  (* context *)
+  type context
+  val context: context
+
+  (* identifying expressions *)
+  datatype item = Expr of Argo_Expr.expr | Term of term
+  datatype identified = New of term | Known of term
+  val identify_item: item -> context -> identified * context
+end
+
+structure Argo_Term: ARGO_TERM =
+struct
+
+(* data types *)
+
+(*
+  The type meta is intentionally hidden to prevent that functions outside of this structure
+  are able to build terms. Meta stores the identifier of the term as well as the complete
+  expression underlying the term.
+*)
+
+datatype meta = M of int * Argo_Expr.expr
+datatype term = T of meta * Argo_Expr.kind * term list
+
+
+(* term operations *)
+
+fun id_of (T (M (id, _), _, _)) = id
+fun expr_of (T (M (_, e), _, _)) = e
+fun type_of t = Argo_Expr.type_of (expr_of t)
+
+(*
+  Comparing terms is fast as only the identifiers are compared. No expressions need to
+  be taken into account.
+*)
+
+fun eq_term (t1, t2) = (id_of t1 = id_of t2)
+fun term_ord (t1, t2) = int_ord (id_of t1, id_of t2)
+
+
+(* sharing of terms *)
+
+(*
+  Kinds are short representation of expressions. Constants and numbers carry additional
+  information and have no arguments. Their kind is hence similar to them. All other expressions
+  are stored in a flat way with identifiers of shared terms as arguments instead of expression
+  as arguments.
+*)
+
+datatype kind =
+  Con of string * Argo_Expr.typ |
+  Num of Rat.rat |
+  Exp of int list
+
+fun kind_of (Argo_Expr.E (Argo_Expr.Con c, _)) _ = Con c
+  | kind_of (Argo_Expr.E (Argo_Expr.Num n, _)) _ = Num n
+  | kind_of (Argo_Expr.E (k, _)) is = Exp (Argo_Expr.int_of_kind k :: is)
+
+fun int_of_kind (Con _) = 1
+  | int_of_kind (Num _) = 2
+  | int_of_kind (Exp _) = 3
+
+fun kind_ord (Con c1, Con c2) = Argo_Expr.con_ord (c1, c2)
+  | kind_ord (Num n1, Num n2) = Rat.ord (n1, n2)
+  | kind_ord (Exp is1, Exp is2) = dict_ord int_ord (is1, is2)
+  | kind_ord (k1, k2) = int_ord (int_of_kind k1, int_of_kind k2)
+
+structure Kindtab = Table(type key = kind val ord = kind_ord)
+
+(*
+  The context keeps track of the next unused identifier as well as all shared terms,
+  which are indexed by their unique kind. For each term, a boolean marker flag is stored.
+  When set to true on an atom, the atom is already asserted to the solver core. When set to
+  true on an if-then-else term, the term has already been lifted.
+
+  Zero is intentionally avoided as identifier, since literals use term identifiers
+  with a sign as literal identifiers.
+*)
+
+type context = {
+  next_id: int,
+  terms: (term * bool) Kindtab.table}
+
+fun mk_context next_id terms: context = {next_id=next_id, terms=terms}
+
+val context = mk_context 1 Kindtab.empty
+
+fun note_atom true kind (t, false) ({next_id, terms}: context) =
+      mk_context next_id (Kindtab.update (kind, (t, true)) terms)
+  | note_atom _ _ _ cx = cx
+
+fun with_unique_id kind is_atom (e as Argo_Expr.E (k, _)) ts ({next_id, terms}: context) =
+  let val t = T (M (next_id, e), k, ts)
+  in ((t, false), mk_context (next_id + 1) (Kindtab.update (kind, (t, is_atom)) terms)) end
+
+fun unique kind is_atom e ts (cx as {terms, ...}: context) =
+  (case Kindtab.lookup terms kind of
+    SOME tp => (tp, note_atom is_atom kind tp cx)
+  | NONE => with_unique_id kind is_atom e ts cx)
+
+
+(* identifying expressions *)
+
+(*
+  Only atoms, i.e., boolean propositons, and if-then-else expressions need to be identified.
+  Other terms are identified implicitly. The identification process works bottom-up.
+
+  The solver core needs to know whether an atom has already been added. Likewise, the clausifier
+  needs to know whether an if-then-else expression has already been lifted. Therefore,
+  the identified term is marked as either "new" when identified for the first time or
+  "known" when it has already been identified before.
+*)
+
+datatype item = Expr of Argo_Expr.expr | Term of term
+datatype identified = New of term | Known of term
+
+fun identify_head is_atom e (ts, cx) = unique (kind_of e (map id_of ts)) is_atom e ts cx
+
+fun identify is_atom (e as Argo_Expr.E (_, es)) cx =
+  identify_head is_atom e (fold_map (apfst fst oo identify false) es cx)
+
+fun identified (t, true) = Known t
+  | identified (t, false) = New t
+
+fun identify_item (Expr e) cx = identify true e cx |>> identified
+  | identify_item (Term (t as T (_, _, ts))) cx =
+      identify_head true (expr_of t) (ts, cx) |>> identified
+
+end
+
+structure Argo_Termtab = Table(type key = Argo_Term.term val ord = Argo_Term.term_ord)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Argo/argo_thy.ML	Thu Sep 29 20:54:44 2016 +0200
@@ -0,0 +1,125 @@
+(*  Title:      Tools/Argo/argo_theory.ML
+    Author:     Sascha Boehme
+
+Combination of all theory solvers.
+
+Currently, it is assumed that theories have distinct domains. Theory solvers do not
+exchange knowledge among each other. This should be changed in the future. Relevant work is:
+
+  Greg Nelson and Derek C. Oppen. Simplification by cooperating decision procedures. In ACM
+  Transactions on Programming Languages and Systems, 1(2):245-257, 1979.
+
+  Leonardo de Moura and Nikolaj Bj/orner. Model-based Theory Combination. In Electronic Notes
+  in Theoretical Computer Science, volume 198(2), pages 37-49, 2008.
+*)
+
+signature ARGO_THY =
+sig
+  (* context *)
+  type context
+  val context: context
+
+  (* simplification *)
+  val simplify: Argo_Rewr.context -> Argo_Rewr.context
+  
+  (* enriching the context *)
+  val add_atom: Argo_Term.term -> context -> Argo_Lit.literal option * context
+
+  (* main operations *)
+  val prepare: context -> context
+  val assume: Argo_Common.literal -> context -> Argo_Lit.literal Argo_Common.implied * context
+  val check: context -> Argo_Lit.literal Argo_Common.implied * context
+  val explain: Argo_Lit.literal -> context -> Argo_Cls.clause * context
+  val add_level: context -> context
+  val backtrack: context -> context
+end
+
+structure Argo_Thy: ARGO_THY =
+struct
+
+(* context *)
+
+type context = Argo_Cc.context * Argo_Simplex.context
+
+val context = (Argo_Cc.context, Argo_Simplex.context)
+
+fun map_cc f (cc, simplex) =
+  let val (x, cc) = f cc
+  in (x, (cc, simplex)) end
+
+fun map_simplex f (cc, simplex) =
+  let val (x, simplex) = f simplex
+  in (x, (cc, simplex)) end
+
+
+(* simplification *)
+
+val simplify = Argo_Cc.simplify #> Argo_Simplex.simplify
+
+
+(* enriching the context *)
+
+fun add_atom t (cc, simplex) =
+  let
+    val (lit1, cc) = Argo_Cc.add_atom t cc
+    val (lit2, simplex) = Argo_Simplex.add_atom t simplex
+  in
+    (case fold (union Argo_Lit.eq_lit o the_list) [lit1, lit2] [] of
+      [] => (NONE, (cc, simplex))
+    | [lit] => (SOME lit, (cc, simplex))
+    | _ => raise Fail "unsynchronized theory solvers")
+  end
+    
+
+
+(* main operations *)
+
+fun prepare (cc, simplex) = (cc, Argo_Simplex.prepare simplex)
+
+local
+
+exception CONFLICT of Argo_Cls.clause * context
+
+datatype tag = All | Cc | Simplex
+
+fun apply f cx =
+  (case f cx of
+    (Argo_Common.Conflict cls, cx) => raise CONFLICT (cls, cx)
+  | (Argo_Common.Implied lits, cx) => (lits, cx))
+
+fun with_lits tag f (txs, lits, cx) =
+  let val (lits', cx) = f cx
+  in (fold (fn l => cons (tag, (l, NONE))) lits' txs, union Argo_Lit.eq_lit lits' lits, cx) end
+
+fun apply0 (tag, f) = with_lits tag (apply f)
+fun apply1 (tag, f) (tag', x) = if tag <> tag' then with_lits tag (apply (f x)) else I
+
+val assumes = [(Cc, map_cc o Argo_Cc.assume), (Simplex, map_simplex o Argo_Simplex.assume)]
+val checks = [(Cc, map_cc Argo_Cc.check), (Simplex, map_simplex Argo_Simplex.check)]
+
+fun propagate ([], lits, cx) = (Argo_Common.Implied lits, cx)
+  | propagate (txs, lits, cx) = propagate (fold_product apply1 assumes txs ([], lits, cx))
+
+in
+
+fun assume lp cx = propagate ([(All, lp)], [], cx)
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+fun check cx = propagate (fold apply0 checks ([], [], cx))
+  handle CONFLICT (cls, cx) => (Argo_Common.Conflict cls, cx)
+
+end
+
+fun explain lit (cc, simplex) =
+  (case Argo_Cc.explain lit cc of
+    SOME (cls, cc) => (cls, (cc, simplex))
+  | NONE =>
+      (case Argo_Simplex.explain lit simplex of
+        SOME (cls, simplex) => (cls, (cc, simplex))
+      | NONE => raise Fail "bad literal without explanation"))
+
+fun add_level (cc, simplex) = (Argo_Cc.add_level cc, Argo_Simplex.add_level simplex)
+
+fun backtrack (cc, simplex) = (Argo_Cc.backtrack cc, Argo_Simplex.backtrack simplex)
+
+end