merged
authorhaftmann
Wed, 17 Feb 2010 09:51:46 +0100
changeset 35161 be96405ccaf3
parent 35155 3011b2149089 (diff)
parent 35160 6eb2b6c1d2d5 (current diff)
child 35162 ea99593b44a5
child 35163 2e0966d6f951
child 35182 4c39632b811f
merged
--- a/NEWS	Wed Feb 17 09:48:53 2010 +0100
+++ b/NEWS	Wed Feb 17 09:51:46 2010 +0100
@@ -9,6 +9,11 @@
 * Code generator: details of internal data cache have no impact on
 the user space functionality any longer.
 
+* Discontinued unnamed infix syntax (legacy feature for many years) --
+need to specify constant name and syntax separately.  Internal ML
+datatype constructors have been renamed from InfixName to Infix etc.
+Minor INCOMPATIBILITY.
+
 
 *** HOL ***
 
--- a/etc/components	Wed Feb 17 09:48:53 2010 +0100
+++ b/etc/components	Wed Feb 17 09:51:46 2010 +0100
@@ -13,6 +13,7 @@
 #misc components
 src/Tools/Code
 src/Tools/WWW_Find
+src/Tools/Cache_IO
 src/HOL/Tools/ATP_Manager
 src/HOL/Mirabelle
 src/HOL/Library/Sum_Of_Squares
--- a/src/FOLP/IFOLP.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/FOLP/IFOLP.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -26,14 +26,14 @@
  EqProof        ::   "[p,p,o]=>prop"    ("(3_ /= _ :/ _)" [10,10,10] 5)
 
       (*** Logical Connectives -- Type Formers ***)
- "="            ::      "['a,'a] => o"  (infixl 50)
+ "op ="         ::      "['a,'a] => o"  (infixl "=" 50)
  True           ::      "o"
  False          ::      "o"
  Not            ::      "o => o"        ("~ _" [40] 40)
- "&"            ::      "[o,o] => o"    (infixr 35)
- "|"            ::      "[o,o] => o"    (infixr 30)
- "-->"          ::      "[o,o] => o"    (infixr 25)
- "<->"          ::      "[o,o] => o"    (infixr 25)
+ "op &"         ::      "[o,o] => o"    (infixr "&" 35)
+ "op |"         ::      "[o,o] => o"    (infixr "|" 30)
+ "op -->"       ::      "[o,o] => o"    (infixr "-->" 25)
+ "op <->"       ::      "[o,o] => o"    (infixr "<->" 25)
       (*Quantifiers*)
  All            ::      "('a => o) => o"        (binder "ALL " 10)
  Ex             ::      "('a => o) => o"        (binder "EX " 10)
@@ -53,9 +53,9 @@
  inr            :: "p=>p"
  when           :: "[p, p=>p, p=>p]=>p"
  lambda         :: "(p => p) => p"      (binder "lam " 55)
- "`"            :: "[p,p]=>p"           (infixl 60)
+ "op `"         :: "[p,p]=>p"           (infixl "`" 60)
  alll           :: "['a=>p]=>p"         (binder "all " 55)
- "^"            :: "[p,'a]=>p"          (infixl 55)
+ "op ^"         :: "[p,'a]=>p"          (infixl "^" 55)
  exists         :: "['a,p]=>p"          ("(1[_,/_])")
  xsplit         :: "[p,['a,p]=>p]=>p"
  ideq           :: "'a=>p"
--- a/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Dijkstra.certs	Wed Feb 17 09:51:46 2010 +0100
@@ -1,4 +1,4 @@
-2/jIbDaU00KSkSih1o9sXg 193550
+JinTdmjIiorL0/vvOyf3+w 6542 0
 #2 := false
 decl up_6 :: (-> T4 T2 bool)
 decl ?x47!7 :: (-> T2 T2)
@@ -6541,4 +6541,3 @@
 #23081 := [unit-resolution #19916 #27207]: #17029
 [unit-resolution #23081 #23182 #18055 #27235]: false
 unsat
-
--- a/src/HOL/Boogie/Examples/Boogie_Max.certs	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Boogie/Examples/Boogie_Max.certs	Wed Feb 17 09:51:46 2010 +0100
@@ -1,4 +1,4 @@
-yJC0k+R1r4pWViX9DxewEQ 62526
+iks4GfP7O/NgNFyGZ4ynjQ 2224 0
 #2 := false
 #4 := 0::int
 decl uf_3 :: (-> int int)
@@ -2223,4 +2223,3 @@
 #2015 := [unit-resolution #2013 #2021]: #2041
 [th-lemma #2015 #2047 #2043]: false
 unsat
-
--- a/src/HOL/Boogie/Examples/VCC_Max.certs	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Boogie/Examples/VCC_Max.certs	Wed Feb 17 09:51:46 2010 +0100
@@ -1,4 +1,4 @@
-8ZKUEUSWY0Pcw6t0NqCjrQ 253722
+6Q8QWdFv463DpfVfkr0XnA 7790 0
 #2 := false
 decl uf_110 :: (-> T4 T5 int)
 decl uf_66 :: (-> T5 int T3 T5)
@@ -7789,4 +7789,3 @@
 #30656 := [unit-resolution #30273 #30655 #30643]: #30496
 [unit-resolution #30529 #30656 #30639]: false
 unsat
-
--- a/src/HOL/Boogie/Examples/VCC_Max.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Boogie/Examples/VCC_Max.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -57,4 +57,4 @@
 
 boogie_end
 
-end
\ No newline at end of file
+end
--- a/src/HOL/Hoare_Parallel/OG_Syntax.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -94,17 +94,6 @@
           annquote_tr' (Syntax.const name) (r :: t :: ts)
       | annbexp_tr' _ _ = raise Match;
 
-    fun upd_tr' (x_upd, T) =
-      (case try (unsuffix Record.updateN) x_upd of
-        SOME x => (x, if T = dummyT then T else Term.domain_type T)
-      | NONE => raise Match);
-
-    fun update_name_tr' (Free x) = Free (upd_tr' x)
-      | update_name_tr' ((c as Const (@{syntax_const "_free"}, _)) $ Free x) =
-          c $ Free (upd_tr' x)
-      | update_name_tr' (Const x) = Const (upd_tr' x)
-      | update_name_tr' _ = raise Match;
-
     fun K_tr' (Abs (_, _, t)) =
           if null (loose_bnos t) then t else raise Match
       | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
@@ -112,12 +101,12 @@
       | K_tr' _ = raise Match;
 
     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const @{syntax_const "_Assign"} $ update_name_tr' f)
+          quote_tr' (Syntax.const @{syntax_const "_Assign"} $ Syntax.update_name_tr' f)
             (Abs (x, dummyT, K_tr' k) :: ts)
       | assign_tr' _ = raise Match;
 
     fun annassign_tr' (r :: Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const @{syntax_const "_AnnAssign"} $ r $ update_name_tr' f)
+          quote_tr' (Syntax.const @{syntax_const "_AnnAssign"} $ r $ Syntax.update_name_tr' f)
             (Abs (x, dummyT, K_tr' k) :: ts)
       | annassign_tr' _ = raise Match;
 
--- a/src/HOL/Hoare_Parallel/RG_Syntax.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -68,17 +68,6 @@
           quote_tr' (Syntax.const name) (t :: ts)
       | bexp_tr' _ _ = raise Match;
 
-    fun upd_tr' (x_upd, T) =
-      (case try (unsuffix Record.updateN) x_upd of
-        SOME x => (x, if T = dummyT then T else Term.domain_type T)
-      | NONE => raise Match);
-
-    fun update_name_tr' (Free x) = Free (upd_tr' x)
-      | update_name_tr' ((c as Const (@{syntax_const "_free"}, _)) $ Free x) =
-          c $ Free (upd_tr' x)
-      | update_name_tr' (Const x) = Const (upd_tr' x)
-      | update_name_tr' _ = raise Match;
-
     fun K_tr' (Abs (_, _, t)) =
           if null (loose_bnos t) then t else raise Match
       | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
@@ -86,7 +75,7 @@
       | K_tr' _ = raise Match;
 
     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const @{syntax_const "_Assign"} $ update_name_tr' f)
+          quote_tr' (Syntax.const @{syntax_const "_Assign"} $ Syntax.update_name_tr' f)
             (Abs (x, dummyT, K_tr' k) :: ts)
       | assign_tr' _ = raise Match;
   in
--- a/src/HOL/Import/xmlconv.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Import/xmlconv.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -218,12 +218,9 @@
   | xml_of_mixfix Structure = wrap_empty "structure"
   | xml_of_mixfix (Mixfix args) = wrap "mixfix" (xml_of_triple xml_of_string (xml_of_list xml_of_int) xml_of_int) args
   | xml_of_mixfix (Delimfix s) = wrap "delimfix" xml_of_string s
-  | xml_of_mixfix (InfixName args) = wrap "infixname" (xml_of_pair xml_of_string xml_of_int) args
-  | xml_of_mixfix (InfixlName args) = wrap "infixlname" (xml_of_pair xml_of_string xml_of_int) args
-  | xml_of_mixfix (InfixrName args) = wrap "infixrname" (xml_of_pair xml_of_string xml_of_int) args
-  | xml_of_mixfix (Infix i) = wrap "infix" xml_of_int i
-  | xml_of_mixfix (Infixl i) = wrap "infixl" xml_of_int i
-  | xml_of_mixfix (Infixr i) = wrap "infixr" xml_of_int i
+  | xml_of_mixfix (Infix args) = wrap "infix" (xml_of_pair xml_of_string xml_of_int) args
+  | xml_of_mixfix (Infixl args) = wrap "infixl" (xml_of_pair xml_of_string xml_of_int) args
+  | xml_of_mixfix (Infixr args) = wrap "infixr" (xml_of_pair xml_of_string xml_of_int) args
   | xml_of_mixfix (Binder args) = wrap "binder" (xml_of_triple xml_of_string xml_of_int xml_of_int) args
                                   
 fun mixfix_of_xml e = 
@@ -232,12 +229,9 @@
        | "structure" => unwrap_empty Structure 
        | "mixfix" => unwrap Mixfix (triple_of_xml string_of_xml (list_of_xml int_of_xml) int_of_xml)
        | "delimfix" => unwrap Delimfix string_of_xml
-       | "infixname" => unwrap InfixName (pair_of_xml string_of_xml int_of_xml) 
-       | "infixlname" => unwrap InfixlName (pair_of_xml string_of_xml int_of_xml)  
-       | "infixrname" => unwrap InfixrName (pair_of_xml string_of_xml int_of_xml)
-       | "infix" => unwrap Infix int_of_xml
-       | "infixl" => unwrap Infixl int_of_xml 
-       | "infixr" => unwrap Infixr int_of_xml
+       | "infix" => unwrap Infix (pair_of_xml string_of_xml int_of_xml) 
+       | "infixl" => unwrap Infixl (pair_of_xml string_of_xml int_of_xml)  
+       | "infixr" => unwrap Infixr (pair_of_xml string_of_xml int_of_xml)
        | "binder" => unwrap Binder (triple_of_xml string_of_xml int_of_xml int_of_xml)
        | _ => parse_err "mixfix"
     ) e
--- a/src/HOL/IsaMakefile	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/IsaMakefile	Wed Feb 17 09:51:46 2010 +0100
@@ -1226,7 +1226,8 @@
   SMT/Tools/smt_solver.ML SMT/Tools/cvc3_solver.ML                      \
   SMT/Tools/yices_solver.ML SMT/Tools/z3_proof_terms.ML                 \
   SMT/Tools/z3_proof_rules.ML SMT/Tools/z3_proof.ML                     \
-  SMT/Tools/z3_model.ML SMT/Tools/z3_interface.ML SMT/Tools/z3_solver.ML
+  SMT/Tools/z3_model.ML SMT/Tools/z3_interface.ML			\
+  SMT/Tools/z3_solver.ML $(SRC)/Tools/Cache_IO/cache_io.ML
 	@cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT
 
 
--- a/src/HOL/Isar_Examples/Hoare.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Isar_Examples/Hoare.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -260,23 +260,14 @@
           quote_tr' (Syntax.const name) (t :: ts)
       | bexp_tr' _ _ = raise Match;
 
-    fun upd_tr' (x_upd, T) =
-      (case try (unsuffix Record.updateN) x_upd of
-        SOME x => (x, if T = dummyT then T else Term.domain_type T)
-      | NONE => raise Match);
-
-    fun update_name_tr' (Free x) = Free (upd_tr' x)
-      | update_name_tr' ((c as Const (@{syntax_const "_free"}, _)) $ Free x) =
-          c $ Free (upd_tr' x)
-      | update_name_tr' (Const x) = Const (upd_tr' x)
-      | update_name_tr' _ = raise Match;
-
-    fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
+    fun K_tr' (Abs (_, _, t)) =
+          if null (loose_bnos t) then t else raise Match
+      | K_tr' (Abs (_, _, Abs (_, _, t) $ Bound 0)) =
+          if null (loose_bnos t) then t else raise Match
       | K_tr' _ = raise Match;
 
     fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const @{syntax_const "_Assign"} $ update_name_tr' f)
+          quote_tr' (Syntax.const @{syntax_const "_Assign"} $ Syntax.update_name_tr' f)
             (Abs (x, dummyT, K_tr' k) :: ts)
       | assign_tr' _ = raise Match;
   in
--- a/src/HOL/Multivariate_Analysis/Determinants.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Determinants.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -121,7 +121,7 @@
 (* Basic determinant properties.                                             *)
 (* ------------------------------------------------------------------------- *)
 
-lemma det_transp: "det (transp A) = det (A::'a::comm_ring_1 ^'n^'n)"
+lemma det_transpose: "det (transpose A) = det (A::'a::comm_ring_1 ^'n^'n)"
 proof-
   let ?di = "\<lambda>A i j. A$i$j"
   let ?U = "(UNIV :: 'n set)"
@@ -133,18 +133,18 @@
     from permutes_inj[OF pU]
     have pi: "inj_on p ?U" by (blast intro: subset_inj_on)
     from permutes_image[OF pU]
-    have "setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U = setprod (\<lambda>i. ?di (transp A) i (inv p i)) (p ` ?U)" by simp
-    also have "\<dots> = setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U"
+    have "setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U = setprod (\<lambda>i. ?di (transpose A) i (inv p i)) (p ` ?U)" by simp
+    also have "\<dots> = setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) o p) ?U"
       unfolding setprod_reindex[OF pi] ..
     also have "\<dots> = setprod (\<lambda>i. ?di A i (p i)) ?U"
     proof-
       {fix i assume i: "i \<in> ?U"
         from i permutes_inv_o[OF pU] permutes_in_image[OF pU]
-        have "((\<lambda>i. ?di (transp A) i (inv p i)) o p) i = ?di A i (p i)"
-          unfolding transp_def by (simp add: expand_fun_eq)}
-      then show "setprod ((\<lambda>i. ?di (transp A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
+        have "((\<lambda>i. ?di (transpose A) i (inv p i)) o p) i = ?di A i (p i)"
+          unfolding transpose_def by (simp add: expand_fun_eq)}
+      then show "setprod ((\<lambda>i. ?di (transpose A) i (inv p i)) o p) ?U = setprod (\<lambda>i. ?di A i (p i)) ?U" by (auto intro: setprod_cong)
     qed
-    finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transp A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
+    finally have "of_int (sign (inv p)) * (setprod (\<lambda>i. ?di (transpose A) i (inv p i)) ?U) = of_int (sign p) * (setprod (\<lambda>i. ?di A i (p i)) ?U)" using sth
       by simp}
   then show ?thesis unfolding det_def apply (subst setsum_permutations_inverse)
   apply (rule setsum_cong2) by blast
@@ -267,12 +267,12 @@
   shows "det(\<chi> i j. A$i$ p j :: 'a^'n^'n) = of_int (sign p) * det A"
 proof-
   let ?Ap = "\<chi> i j. A$i$ p j :: 'a^'n^'n"
-  let ?At = "transp A"
-  have "of_int (sign p) * det A = det (transp (\<chi> i. transp A $ p i))"
-    unfolding det_permute_rows[OF p, of ?At] det_transp ..
+  let ?At = "transpose A"
+  have "of_int (sign p) * det A = det (transpose (\<chi> i. transpose A $ p i))"
+    unfolding det_permute_rows[OF p, of ?At] det_transpose ..
   moreover
-  have "?Ap = transp (\<chi> i. transp A $ p i)"
-    by (simp add: transp_def Cart_eq)
+  have "?Ap = transpose (\<chi> i. transpose A $ p i)"
+    by (simp add: transpose_def Cart_eq)
   ultimately show ?thesis by simp
 qed
 
@@ -299,9 +299,9 @@
   assumes ij: "i \<noteq> j"
   and r: "column i A = column j A"
   shows "det A = 0"
-apply (subst det_transp[symmetric])
+apply (subst det_transpose[symmetric])
 apply (rule det_identical_rows[OF ij])
-by (metis row_transp r)
+by (metis row_transpose r)
 
 lemma det_zero_row:
   fixes A :: "'a::{idom, ring_char_0}^'n^'n"
@@ -317,9 +317,9 @@
   fixes A :: "'a::{idom,ring_char_0}^'n^'n"
   assumes r: "column i A = 0"
   shows "det A = 0"
-  apply (subst det_transp[symmetric])
+  apply (subst det_transpose[symmetric])
   apply (rule det_zero_row [of i])
-  by (metis row_transp r)
+  by (metis row_transpose r)
 
 lemma det_row_add:
   fixes a b c :: "'n::finite \<Rightarrow> _ ^ 'n"
@@ -489,7 +489,7 @@
 qed
 
 lemma det_dependent_columns: assumes d: "dependent(columns (A::'a::linordered_idom^'n^'n))" shows "det A = 0"
-by (metis d det_dependent_rows rows_transp det_transp)
+by (metis d det_dependent_rows rows_transpose det_transpose)
 
 (* ------------------------------------------------------------------------- *)
 (* Multilinearity and the multiplication formula.                            *)
@@ -760,7 +760,7 @@
 (* Cramer's rule.                                                            *)
 (* ------------------------------------------------------------------------- *)
 
-lemma cramer_lemma_transp:
+lemma cramer_lemma_transpose:
   fixes A:: "'a::linordered_idom^'n^'n" and x :: "'a ^'n"
   shows "det ((\<chi> i. if i = k then setsum (\<lambda>i. x$i *s row i A) (UNIV::'n set)
                            else row i A)::'a^'n^'n) = x$k * det A"
@@ -801,13 +801,13 @@
   shows "det((\<chi> i j. if j = k then (A *v x)$i else A$i$j):: 'a^'n^'n) = x$k * det A"
 proof-
   let ?U = "UNIV :: 'n set"
-  have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transp A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
-    by (auto simp add: row_transp intro: setsum_cong2)
+  have stupid: "\<And>c. setsum (\<lambda>i. c i *s row i (transpose A)) ?U = setsum (\<lambda>i. c i *s column i A) ?U"
+    by (auto simp add: row_transpose intro: setsum_cong2)
   show ?thesis  unfolding matrix_mult_vsum
-  unfolding cramer_lemma_transp[of k x "transp A", unfolded det_transp, symmetric]
+  unfolding cramer_lemma_transpose[of k x "transpose A", unfolded det_transpose, symmetric]
   unfolding stupid[of "\<lambda>i. x$i"]
-  apply (subst det_transp[symmetric])
-  apply (rule cong[OF refl[of det]]) by (vector transp_def column_def row_def)
+  apply (subst det_transpose[symmetric])
+  apply (rule cong[OF refl[of det]]) by (vector transpose_def column_def row_def)
 qed
 
 lemma cramer:
@@ -840,13 +840,13 @@
   apply (simp add: real_vector_norm_def)
   by (simp add: dot_norm  linear_add[symmetric])
 
-definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transp Q ** Q = mat 1 \<and> Q ** transp Q = mat 1"
+definition "orthogonal_matrix (Q::'a::semiring_1^'n^'n) \<longleftrightarrow> transpose Q ** Q = mat 1 \<and> Q ** transpose Q = mat 1"
 
-lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n)  \<longleftrightarrow> transp Q ** Q = mat 1"
+lemma orthogonal_matrix: "orthogonal_matrix (Q:: real ^'n^'n)  \<longleftrightarrow> transpose Q ** Q = mat 1"
   by (metis matrix_left_right_inverse orthogonal_matrix_def)
 
 lemma orthogonal_matrix_id: "orthogonal_matrix (mat 1 :: _^'n^'n)"
-  by (simp add: orthogonal_matrix_def transp_mat matrix_mul_lid)
+  by (simp add: orthogonal_matrix_def transpose_mat matrix_mul_lid)
 
 lemma orthogonal_matrix_mul:
   fixes A :: "real ^'n^'n"
@@ -854,7 +854,7 @@
   and oB: "orthogonal_matrix B"
   shows "orthogonal_matrix(A ** B)"
   using oA oB
-  unfolding orthogonal_matrix matrix_transp_mul
+  unfolding orthogonal_matrix matrix_transpose_mul
   apply (subst matrix_mul_assoc)
   apply (subst matrix_mul_assoc[symmetric])
   by (simp add: matrix_mul_rid)
@@ -873,7 +873,7 @@
     from ot have lf: "linear f" and fd: "\<forall>v w. f v \<bullet> f w = v \<bullet> w"
       unfolding  orthogonal_transformation_def orthogonal_matrix by blast+
     {fix i j
-      let ?A = "transp ?mf ** ?mf"
+      let ?A = "transpose ?mf ** ?mf"
       have th0: "\<And>b (x::'a::comm_ring_1). (if b then 1 else 0)*x = (if b then x else 0)"
         "\<And>b (x::'a::comm_ring_1). x*(if b then 1 else 0) = (if b then x else 0)"
         by simp_all
@@ -908,9 +908,9 @@
     also have "\<dots> \<longleftrightarrow> x = 1 \<or> x = - 1" unfolding th0 th1 by simp
     finally show "?ths x" ..
   qed
-  from oQ have "Q ** transp Q = mat 1" by (metis orthogonal_matrix_def)
-  hence "det (Q ** transp Q) = det (mat 1:: 'a^'n^'n)" by simp
-  hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transp)
+  from oQ have "Q ** transpose Q = mat 1" by (metis orthogonal_matrix_def)
+  hence "det (Q ** transpose Q) = det (mat 1:: 'a^'n^'n)" by simp
+  hence "det Q * det Q = 1" by (simp add: det_mul det_I det_transpose)
   then show ?thesis unfolding th .
 qed
 
--- a/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Multivariate_Analysis/Euclidean_Space.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -2029,7 +2029,8 @@
   where "v v* m == (\<chi> j. setsum (\<lambda>i. ((m$i)$j) * (v$i)) (UNIV :: 'm set)) :: 'a^'n"
 
 definition "(mat::'a::zero => 'a ^'n^'n) k = (\<chi> i j. if i = j then k else 0)"
-definition "(transp::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
+definition transpose where 
+  "(transpose::'a^'n^'m \<Rightarrow> 'a^'m^'n) A = (\<chi> i j. ((A$j)$i))"
 definition "(row::'m => 'a ^'n^'m \<Rightarrow> 'a ^'n) i A = (\<chi> j. ((A$i)$j))"
 definition "(column::'n =>'a^'n^'m =>'a^'m) j A = (\<chi> i. ((A$i)$j))"
 definition "rows(A::'a^'n^'m) = { row i A | i. i \<in> (UNIV :: 'm set)}"
@@ -2071,8 +2072,8 @@
   by (simp add: cond_value_iff cond_application_beta
     setsum_delta' cong del: if_weak_cong)
 
-lemma matrix_transp_mul: "transp(A ** B) = transp B ** transp (A::'a::comm_semiring_1^_^_)"
-  by (simp add: matrix_matrix_mult_def transp_def Cart_eq mult_commute)
+lemma matrix_transpose_mul: "transpose(A ** B) = transpose B ** transpose (A::'a::comm_semiring_1^_^_)"
+  by (simp add: matrix_matrix_mult_def transpose_def Cart_eq mult_commute)
 
 lemma matrix_eq:
   fixes A B :: "'a::semiring_1 ^ 'n ^ 'm"
@@ -2094,26 +2095,26 @@
   apply (subst setsum_commute)
   by simp
 
-lemma transp_mat: "transp (mat n) = mat n"
-  by (vector transp_def mat_def)
-
-lemma transp_transp: "transp(transp A) = A"
-  by (vector transp_def)
-
-lemma row_transp:
+lemma transpose_mat: "transpose (mat n) = mat n"
+  by (vector transpose_def mat_def)
+
+lemma transpose_transpose: "transpose(transpose A) = A"
+  by (vector transpose_def)
+
+lemma row_transpose:
   fixes A:: "'a::semiring_1^_^_"
-  shows "row i (transp A) = column i A"
-  by (simp add: row_def column_def transp_def Cart_eq)
-
-lemma column_transp:
+  shows "row i (transpose A) = column i A"
+  by (simp add: row_def column_def transpose_def Cart_eq)
+
+lemma column_transpose:
   fixes A:: "'a::semiring_1^_^_"
-  shows "column i (transp A) = row i A"
-  by (simp add: row_def column_def transp_def Cart_eq)
-
-lemma rows_transp: "rows(transp (A::'a::semiring_1^_^_)) = columns A"
-by (auto simp add: rows_def columns_def row_transp intro: set_ext)
-
-lemma columns_transp: "columns(transp (A::'a::semiring_1^_^_)) = rows A" by (metis transp_transp rows_transp)
+  shows "column i (transpose A) = row i A"
+  by (simp add: row_def column_def transpose_def Cart_eq)
+
+lemma rows_transpose: "rows(transpose (A::'a::semiring_1^_^_)) = columns A"
+by (auto simp add: rows_def columns_def row_transpose intro: set_ext)
+
+lemma columns_transpose: "columns(transpose (A::'a::semiring_1^_^_)) = rows A" by (metis transpose_transpose rows_transpose)
 
 text{* Two sometimes fruitful ways of looking at matrix-vector multiplication. *}
 
@@ -2176,19 +2177,19 @@
   using lf lg linear_compose[OF lf lg] matrix_works[OF linear_compose[OF lf lg]]
   by (simp  add: matrix_eq matrix_works matrix_vector_mul_assoc[symmetric] o_def)
 
-lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^_) *v x = setsum (\<lambda>i. (x$i) *s ((transp A)$i)) (UNIV:: 'n set)"
-  by (simp add: matrix_vector_mult_def transp_def Cart_eq mult_commute)
-
-lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n^'m) *v x) = (\<lambda>x. transp A *v x)"
+lemma matrix_vector_column:"(A::'a::comm_semiring_1^'n^_) *v x = setsum (\<lambda>i. (x$i) *s ((transpose A)$i)) (UNIV:: 'n set)"
+  by (simp add: matrix_vector_mult_def transpose_def Cart_eq mult_commute)
+
+lemma adjoint_matrix: "adjoint(\<lambda>x. (A::'a::comm_ring_1^'n^'m) *v x) = (\<lambda>x. transpose A *v x)"
   apply (rule adjoint_unique[symmetric])
   apply (rule matrix_vector_mul_linear)
-  apply (simp add: transp_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)
+  apply (simp add: transpose_def dot_def matrix_vector_mult_def setsum_left_distrib setsum_right_distrib)
   apply (subst setsum_commute)
   apply (auto simp add: mult_ac)
   done
 
 lemma matrix_adjoint: assumes lf: "linear (f :: 'a::comm_ring_1^'n \<Rightarrow> 'a ^'m)"
-  shows "matrix(adjoint f) = transp(matrix f)"
+  shows "matrix(adjoint f) = transpose(matrix f)"
   apply (subst matrix_vector_mul[OF lf])
   unfolding adjoint_matrix matrix_of_matrix_vector_mul ..
 
@@ -4317,13 +4318,13 @@
 
 (* Detailed theorems about left and right invertibility in general case.     *)
 
-lemma left_invertible_transp:
-  "(\<exists>(B). B ** transp (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
-  by (metis matrix_transp_mul transp_mat transp_transp)
-
-lemma right_invertible_transp:
-  "(\<exists>(B). transp (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
-  by (metis matrix_transp_mul transp_mat transp_transp)
+lemma left_invertible_transpose:
+  "(\<exists>(B). B ** transpose (A) = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). A ** B = mat 1)"
+  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
+
+lemma right_invertible_transpose:
+  "(\<exists>(B). transpose (A) ** B = mat (1::'a::comm_semiring_1)) \<longleftrightarrow> (\<exists>(B). B ** A = mat 1)"
+  by (metis matrix_transpose_mul transpose_mat transpose_transpose)
 
 lemma linear_injective_left_inverse:
   assumes lf: "linear (f::real ^'n \<Rightarrow> real ^'m)" and fi: "inj f"
@@ -4438,9 +4439,9 @@
 lemma matrix_right_invertible_independent_rows:
   fixes A :: "real^'n^'m"
   shows "(\<exists>(B::real^'m^'n). A ** B = mat 1) \<longleftrightarrow> (\<forall>c. setsum (\<lambda>i. c i *s row i A) (UNIV :: 'm set) = 0 \<longrightarrow> (\<forall>i. c i = 0))"
-  unfolding left_invertible_transp[symmetric]
+  unfolding left_invertible_transpose[symmetric]
     matrix_left_invertible_independent_columns
-  by (simp add: column_transp)
+  by (simp add: column_transpose)
 
 lemma matrix_right_invertible_span_columns:
   "(\<exists>(B::real ^'n^'m). (A::real ^'m^'n) ** B = mat 1) \<longleftrightarrow> span (columns A) = UNIV" (is "?lhs = ?rhs")
@@ -4506,8 +4507,8 @@
 
 lemma matrix_left_invertible_span_rows:
   "(\<exists>(B::real^'m^'n). B ** (A::real^'n^'m) = mat 1) \<longleftrightarrow> span (rows A) = UNIV"
-  unfolding right_invertible_transp[symmetric]
-  unfolding columns_transp[symmetric]
+  unfolding right_invertible_transpose[symmetric]
+  unfolding columns_transpose[symmetric]
   unfolding matrix_right_invertible_span_columns
  ..
 
@@ -4728,12 +4729,12 @@
 
 definition "columnvector v = (\<chi> i j. (v$i))"
 
-lemma transp_columnvector:
- "transp(columnvector v) = rowvector v"
-  by (simp add: transp_def rowvector_def columnvector_def Cart_eq)
-
-lemma transp_rowvector: "transp(rowvector v) = columnvector v"
-  by (simp add: transp_def columnvector_def rowvector_def Cart_eq)
+lemma transpose_columnvector:
+ "transpose(columnvector v) = rowvector v"
+  by (simp add: transpose_def rowvector_def columnvector_def Cart_eq)
+
+lemma transpose_rowvector: "transpose(rowvector v) = columnvector v"
+  by (simp add: transpose_def columnvector_def rowvector_def Cart_eq)
 
 lemma dot_rowvector_columnvector:
   "columnvector (A *v v) = A ** columnvector v"
@@ -4745,9 +4746,9 @@
 lemma dot_matrix_vector_mul:
   fixes A B :: "real ^'n ^'n" and x y :: "real ^'n"
   shows "(A *v x) \<bullet> (B *v y) =
-      (((rowvector x :: real^'n^1) ** ((transp A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
-unfolding dot_matrix_product transp_columnvector[symmetric]
-  dot_rowvector_columnvector matrix_transp_mul matrix_mul_assoc ..
+      (((rowvector x :: real^'n^1) ** ((transpose A ** B) ** (columnvector y :: real ^1^'n)))$1)$1"
+unfolding dot_matrix_product transpose_columnvector[symmetric]
+  dot_rowvector_columnvector matrix_transpose_mul matrix_mul_assoc ..
 
 (* Infinity norm.                                                            *)
 
--- a/src/HOL/Record.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Record.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -79,48 +79,64 @@
 
 subsection {* Operators and lemmas for types isomorphic to tuples *}
 
-datatype ('a, 'b, 'c) tuple_isomorphism = Tuple_Isomorphism "'a \<Rightarrow> 'b \<times> 'c" "'b \<times> 'c \<Rightarrow> 'a"
+datatype ('a, 'b, 'c) tuple_isomorphism =
+  Tuple_Isomorphism "'a \<Rightarrow> 'b \<times> 'c" "'b \<times> 'c \<Rightarrow> 'a"
 
-primrec repr :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'c" where
+primrec
+  repr :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'b \<times> 'c" where
   "repr (Tuple_Isomorphism r a) = r"
 
-primrec abst :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'b \<times> 'c \<Rightarrow> 'a" where
+primrec
+  abst :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'b \<times> 'c \<Rightarrow> 'a" where
   "abst (Tuple_Isomorphism r a) = a"
 
-definition iso_tuple_fst :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'b" where
+definition
+  iso_tuple_fst :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'b" where
   "iso_tuple_fst isom = fst \<circ> repr isom"
 
-definition iso_tuple_snd :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'c" where
+definition
+  iso_tuple_snd :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'a \<Rightarrow> 'c" where
   "iso_tuple_snd isom = snd \<circ> repr isom"
 
-definition iso_tuple_fst_update :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)" where
+definition
+  iso_tuple_fst_update ::
+    "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)" where
   "iso_tuple_fst_update isom f = abst isom \<circ> apfst f \<circ> repr isom"
 
-definition iso_tuple_snd_update :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> ('c \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'a)" where
+definition
+  iso_tuple_snd_update ::
+    "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> ('c \<Rightarrow> 'c) \<Rightarrow> ('a \<Rightarrow> 'a)" where
   "iso_tuple_snd_update isom f = abst isom \<circ> apsnd f \<circ> repr isom"
 
-definition iso_tuple_cons :: "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'a" where
+definition
+  iso_tuple_cons ::
+    "('a, 'b, 'c) tuple_isomorphism \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'a" where
   "iso_tuple_cons isom = curry (abst isom)"
 
 
 subsection {* Logical infrastructure for records *}
 
-definition iso_tuple_surjective_proof_assist :: "'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
+definition
+  iso_tuple_surjective_proof_assist :: "'a \<Rightarrow> 'b \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
   "iso_tuple_surjective_proof_assist x y f \<longleftrightarrow> f x = y"
 
-definition iso_tuple_update_accessor_cong_assist :: "(('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
-  "iso_tuple_update_accessor_cong_assist upd acc \<longleftrightarrow> 
+definition
+  iso_tuple_update_accessor_cong_assist ::
+    "(('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool" where
+  "iso_tuple_update_accessor_cong_assist upd acc \<longleftrightarrow>
      (\<forall>f v. upd (\<lambda>x. f (acc v)) v = upd f v) \<and> (\<forall>v. upd id v = v)"
 
-definition iso_tuple_update_accessor_eq_assist :: "(('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
+definition
+  iso_tuple_update_accessor_eq_assist ::
+    "(('b \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'a)) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> ('b \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool" where
   "iso_tuple_update_accessor_eq_assist upd acc v f v' x \<longleftrightarrow>
      upd f v = v' \<and> acc v = x \<and> iso_tuple_update_accessor_cong_assist upd acc"
 
 lemma update_accessor_congruence_foldE:
   assumes uac: "iso_tuple_update_accessor_cong_assist upd acc"
-  and       r: "r = r'" and v: "acc r' = v'"
-  and       f: "\<And>v. v' = v \<Longrightarrow> f v = f' v"
-  shows        "upd f r = upd f' r'"
+    and r: "r = r'" and v: "acc r' = v'"
+    and f: "\<And>v. v' = v \<Longrightarrow> f v = f' v"
+  shows "upd f r = upd f' r'"
   using uac r v [symmetric]
   apply (subgoal_tac "upd (\<lambda>x. f (acc r')) r' = upd (\<lambda>x. f' (acc r')) r'")
    apply (simp add: iso_tuple_update_accessor_cong_assist_def)
@@ -128,8 +144,9 @@
   done
 
 lemma update_accessor_congruence_unfoldE:
-  "iso_tuple_update_accessor_cong_assist upd acc \<Longrightarrow> r = r' \<Longrightarrow> acc r' = v' \<Longrightarrow> (\<And>v. v = v' \<Longrightarrow> f v = f' v)
-     \<Longrightarrow> upd f r = upd f' r'"
+  "iso_tuple_update_accessor_cong_assist upd acc \<Longrightarrow>
+    r = r' \<Longrightarrow> acc r' = v' \<Longrightarrow> (\<And>v. v = v' \<Longrightarrow> f v = f' v) \<Longrightarrow>
+    upd f r = upd f' r'"
   apply (erule(2) update_accessor_congruence_foldE)
   apply simp
   done
@@ -140,15 +157,16 @@
 
 lemma update_accessor_noopE:
   assumes uac: "iso_tuple_update_accessor_cong_assist upd acc"
-      and acc: "f (acc x) = acc x"
-  shows        "upd f x = x"
-using uac by (simp add: acc iso_tuple_update_accessor_cong_assist_id [OF uac, unfolded id_def]
-  cong: update_accessor_congruence_unfoldE [OF uac])
+    and acc: "f (acc x) = acc x"
+  shows "upd f x = x"
+  using uac
+  by (simp add: acc iso_tuple_update_accessor_cong_assist_id [OF uac, unfolded id_def]
+    cong: update_accessor_congruence_unfoldE [OF uac])
 
 lemma update_accessor_noop_compE:
   assumes uac: "iso_tuple_update_accessor_cong_assist upd acc"
-  assumes acc: "f (acc x) = acc x"
-  shows      "upd (g \<circ> f) x = upd g x"
+    and acc: "f (acc x) = acc x"
+  shows "upd (g \<circ> f) x = upd g x"
   by (simp add: acc cong: update_accessor_congruence_unfoldE[OF uac])
 
 lemma update_accessor_cong_assist_idI:
@@ -156,7 +174,8 @@
   by (simp add: iso_tuple_update_accessor_cong_assist_def)
 
 lemma update_accessor_cong_assist_triv:
-  "iso_tuple_update_accessor_cong_assist upd acc \<Longrightarrow> iso_tuple_update_accessor_cong_assist upd acc"
+  "iso_tuple_update_accessor_cong_assist upd acc \<Longrightarrow>
+    iso_tuple_update_accessor_cong_assist upd acc"
   by assumption
 
 lemma update_accessor_accessor_eqE:
@@ -172,11 +191,13 @@
   by (simp add: iso_tuple_update_accessor_eq_assist_def update_accessor_cong_assist_idI)
 
 lemma iso_tuple_update_accessor_eq_assist_triv:
-  "iso_tuple_update_accessor_eq_assist upd acc v f v' x \<Longrightarrow> iso_tuple_update_accessor_eq_assist upd acc v f v' x"
+  "iso_tuple_update_accessor_eq_assist upd acc v f v' x \<Longrightarrow>
+    iso_tuple_update_accessor_eq_assist upd acc v f v' x"
   by assumption
 
 lemma iso_tuple_update_accessor_cong_from_eq:
-  "iso_tuple_update_accessor_eq_assist upd acc v f v' x \<Longrightarrow> iso_tuple_update_accessor_cong_assist upd acc"
+  "iso_tuple_update_accessor_eq_assist upd acc v f v' x \<Longrightarrow>
+    iso_tuple_update_accessor_cong_assist upd acc"
   by (simp add: iso_tuple_update_accessor_eq_assist_def)
 
 lemma iso_tuple_surjective_proof_assistI:
@@ -190,124 +211,139 @@
 locale isomorphic_tuple =
   fixes isom :: "('a, 'b, 'c) tuple_isomorphism"
   assumes repr_inv: "\<And>x. abst isom (repr isom x) = x"
-  assumes abst_inv: "\<And>y. repr isom (abst isom y) = y"
+    and abst_inv: "\<And>y. repr isom (abst isom y) = y"
 begin
 
-lemma repr_inj:
-  "repr isom x = repr isom y \<longleftrightarrow> x = y"
-  by (auto dest: arg_cong [of "repr isom x" "repr isom y" "abst isom"] simp add: repr_inv)
+lemma repr_inj: "repr isom x = repr isom y \<longleftrightarrow> x = y"
+  by (auto dest: arg_cong [of "repr isom x" "repr isom y" "abst isom"]
+    simp add: repr_inv)
 
-lemma abst_inj:
-  "abst isom x = abst isom y \<longleftrightarrow> x = y"
-  by (auto dest: arg_cong [of "abst isom x" "abst isom y" "repr isom"] simp add: abst_inv)
+lemma abst_inj: "abst isom x = abst isom y \<longleftrightarrow> x = y"
+  by (auto dest: arg_cong [of "abst isom x" "abst isom y" "repr isom"]
+    simp add: abst_inv)
 
 lemmas simps = Let_def repr_inv abst_inv repr_inj abst_inj
 
 lemma iso_tuple_access_update_fst_fst:
   "f o h g = j o f \<Longrightarrow>
-    (f o iso_tuple_fst isom) o (iso_tuple_fst_update isom o h) g
-          = j o (f o iso_tuple_fst isom)"
+    (f o iso_tuple_fst isom) o (iso_tuple_fst_update isom o h) g =
+      j o (f o iso_tuple_fst isom)"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_fst_def simps
-             intro!: ext elim!: o_eq_elim)
+    intro!: ext elim!: o_eq_elim)
 
 lemma iso_tuple_access_update_snd_snd:
   "f o h g = j o f \<Longrightarrow>
-    (f o iso_tuple_snd isom) o (iso_tuple_snd_update isom o h) g
-          = j o (f o iso_tuple_snd isom)"
+    (f o iso_tuple_snd isom) o (iso_tuple_snd_update isom o h) g =
+      j o (f o iso_tuple_snd isom)"
   by (clarsimp simp: iso_tuple_snd_update_def iso_tuple_snd_def simps
-             intro!: ext elim!: o_eq_elim)
+    intro!: ext elim!: o_eq_elim)
 
 lemma iso_tuple_access_update_fst_snd:
-  "(f o iso_tuple_fst isom) o (iso_tuple_snd_update isom o h) g
-          = id o (f o iso_tuple_fst isom)"
+  "(f o iso_tuple_fst isom) o (iso_tuple_snd_update isom o h) g =
+    id o (f o iso_tuple_fst isom)"
   by (clarsimp simp: iso_tuple_snd_update_def iso_tuple_fst_def simps
-             intro!: ext elim!: o_eq_elim)
+    intro!: ext elim!: o_eq_elim)
 
 lemma iso_tuple_access_update_snd_fst:
-  "(f o iso_tuple_snd isom) o (iso_tuple_fst_update isom o h) g
-          = id o (f o iso_tuple_snd isom)"
+  "(f o iso_tuple_snd isom) o (iso_tuple_fst_update isom o h) g =
+    id o (f o iso_tuple_snd isom)"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_def simps
-             intro!: ext elim!: o_eq_elim)
+    intro!: ext elim!: o_eq_elim)
 
 lemma iso_tuple_update_swap_fst_fst:
   "h f o j g = j g o h f \<Longrightarrow>
-    (iso_tuple_fst_update isom o h) f o (iso_tuple_fst_update isom o j) g
-          = (iso_tuple_fst_update isom o j) g o (iso_tuple_fst_update isom o h) f"
+    (iso_tuple_fst_update isom o h) f o (iso_tuple_fst_update isom o j) g =
+      (iso_tuple_fst_update isom o j) g o (iso_tuple_fst_update isom o h) f"
   by (clarsimp simp: iso_tuple_fst_update_def simps apfst_compose intro!: ext)
 
 lemma iso_tuple_update_swap_snd_snd:
   "h f o j g = j g o h f \<Longrightarrow>
-    (iso_tuple_snd_update isom o h) f o (iso_tuple_snd_update isom o j) g
-          = (iso_tuple_snd_update isom o j) g o (iso_tuple_snd_update isom o h) f"
+    (iso_tuple_snd_update isom o h) f o (iso_tuple_snd_update isom o j) g =
+      (iso_tuple_snd_update isom o j) g o (iso_tuple_snd_update isom o h) f"
   by (clarsimp simp: iso_tuple_snd_update_def simps apsnd_compose intro!: ext)
 
 lemma iso_tuple_update_swap_fst_snd:
-  "(iso_tuple_snd_update isom o h) f o (iso_tuple_fst_update isom o j) g
-          = (iso_tuple_fst_update isom o j) g o (iso_tuple_snd_update isom o h) f"
-  by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def simps intro!: ext)
+  "(iso_tuple_snd_update isom o h) f o (iso_tuple_fst_update isom o j) g =
+    (iso_tuple_fst_update isom o j) g o (iso_tuple_snd_update isom o h) f"
+  by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def
+    simps intro!: ext)
 
 lemma iso_tuple_update_swap_snd_fst:
-  "(iso_tuple_fst_update isom o h) f o (iso_tuple_snd_update isom o j) g
-          = (iso_tuple_snd_update isom o j) g o (iso_tuple_fst_update isom o h) f"
+  "(iso_tuple_fst_update isom o h) f o (iso_tuple_snd_update isom o j) g =
+    (iso_tuple_snd_update isom o j) g o (iso_tuple_fst_update isom o h) f"
   by (clarsimp simp: iso_tuple_fst_update_def iso_tuple_snd_update_def simps intro!: ext)
 
 lemma iso_tuple_update_compose_fst_fst:
   "h f o j g = k (f o g) \<Longrightarrow>
-    (iso_tuple_fst_update isom o h) f o (iso_tuple_fst_update isom o j) g
-          = (iso_tuple_fst_update isom o k) (f o g)"
+    (iso_tuple_fst_update isom o h) f o (iso_tuple_fst_update isom o j) g =
+      (iso_tuple_fst_update isom o k) (f o g)"
   by (clarsimp simp: iso_tuple_fst_update_def simps apfst_compose intro!: ext)
 
 lemma iso_tuple_update_compose_snd_snd:
   "h f o j g = k (f o g) \<Longrightarrow>
-    (iso_tuple_snd_update isom o h) f o (iso_tuple_snd_update isom o j) g
-          = (iso_tuple_snd_update isom o k) (f o g)"
+    (iso_tuple_snd_update isom o h) f o (iso_tuple_snd_update isom o j) g =
+      (iso_tuple_snd_update isom o k) (f o g)"
   by (clarsimp simp: iso_tuple_snd_update_def simps apsnd_compose intro!: ext)
 
 lemma iso_tuple_surjective_proof_assist_step:
   "iso_tuple_surjective_proof_assist v a (iso_tuple_fst isom o f) \<Longrightarrow>
-     iso_tuple_surjective_proof_assist v b (iso_tuple_snd isom o f)
-      \<Longrightarrow> iso_tuple_surjective_proof_assist v (iso_tuple_cons isom a b) f"
+    iso_tuple_surjective_proof_assist v b (iso_tuple_snd isom o f) \<Longrightarrow>
+    iso_tuple_surjective_proof_assist v (iso_tuple_cons isom a b) f"
   by (clarsimp simp: iso_tuple_surjective_proof_assist_def simps
     iso_tuple_fst_def iso_tuple_snd_def iso_tuple_cons_def)
 
 lemma iso_tuple_fst_update_accessor_cong_assist:
   assumes "iso_tuple_update_accessor_cong_assist f g"
-  shows "iso_tuple_update_accessor_cong_assist (iso_tuple_fst_update isom o f) (g o iso_tuple_fst isom)"
+  shows "iso_tuple_update_accessor_cong_assist
+    (iso_tuple_fst_update isom o f) (g o iso_tuple_fst isom)"
 proof -
-  from assms have "f id = id" by (rule iso_tuple_update_accessor_cong_assist_id)
-  with assms show ?thesis by (clarsimp simp: iso_tuple_update_accessor_cong_assist_def simps
-    iso_tuple_fst_update_def iso_tuple_fst_def)
+  from assms have "f id = id"
+    by (rule iso_tuple_update_accessor_cong_assist_id)
+  with assms show ?thesis
+    by (clarsimp simp: iso_tuple_update_accessor_cong_assist_def simps
+      iso_tuple_fst_update_def iso_tuple_fst_def)
 qed
 
 lemma iso_tuple_snd_update_accessor_cong_assist:
   assumes "iso_tuple_update_accessor_cong_assist f g"
-  shows "iso_tuple_update_accessor_cong_assist (iso_tuple_snd_update isom o f) (g o iso_tuple_snd isom)"
+  shows "iso_tuple_update_accessor_cong_assist
+    (iso_tuple_snd_update isom o f) (g o iso_tuple_snd isom)"
 proof -
-  from assms have "f id = id" by (rule iso_tuple_update_accessor_cong_assist_id)
-  with assms show ?thesis by (clarsimp simp: iso_tuple_update_accessor_cong_assist_def simps
-    iso_tuple_snd_update_def iso_tuple_snd_def)
+  from assms have "f id = id"
+    by (rule iso_tuple_update_accessor_cong_assist_id)
+  with assms show ?thesis
+    by (clarsimp simp: iso_tuple_update_accessor_cong_assist_def simps
+      iso_tuple_snd_update_def iso_tuple_snd_def)
 qed
 
 lemma iso_tuple_fst_update_accessor_eq_assist:
   assumes "iso_tuple_update_accessor_eq_assist f g a u a' v"
-  shows "iso_tuple_update_accessor_eq_assist (iso_tuple_fst_update isom o f) (g o iso_tuple_fst isom)
+  shows "iso_tuple_update_accessor_eq_assist
+    (iso_tuple_fst_update isom o f) (g o iso_tuple_fst isom)
     (iso_tuple_cons isom a b) u (iso_tuple_cons isom a' b) v"
 proof -
   from assms have "f id = id"
-    by (auto simp add: iso_tuple_update_accessor_eq_assist_def intro: iso_tuple_update_accessor_cong_assist_id)
-  with assms show ?thesis by (clarsimp simp: iso_tuple_update_accessor_eq_assist_def
-    iso_tuple_fst_update_def iso_tuple_fst_def iso_tuple_update_accessor_cong_assist_def iso_tuple_cons_def simps)
+    by (auto simp add: iso_tuple_update_accessor_eq_assist_def
+      intro: iso_tuple_update_accessor_cong_assist_id)
+  with assms show ?thesis
+    by (clarsimp simp: iso_tuple_update_accessor_eq_assist_def
+      iso_tuple_fst_update_def iso_tuple_fst_def
+      iso_tuple_update_accessor_cong_assist_def iso_tuple_cons_def simps)
 qed
 
 lemma iso_tuple_snd_update_accessor_eq_assist:
   assumes "iso_tuple_update_accessor_eq_assist f g b u b' v"
-  shows "iso_tuple_update_accessor_eq_assist (iso_tuple_snd_update isom o f) (g o iso_tuple_snd isom)
+  shows "iso_tuple_update_accessor_eq_assist
+    (iso_tuple_snd_update isom o f) (g o iso_tuple_snd isom)
     (iso_tuple_cons isom a b) u (iso_tuple_cons isom a b') v"
 proof -
   from assms have "f id = id"
-    by (auto simp add: iso_tuple_update_accessor_eq_assist_def intro: iso_tuple_update_accessor_cong_assist_id)
-  with assms show ?thesis by (clarsimp simp: iso_tuple_update_accessor_eq_assist_def
-    iso_tuple_snd_update_def iso_tuple_snd_def iso_tuple_update_accessor_cong_assist_def iso_tuple_cons_def simps)
+    by (auto simp add: iso_tuple_update_accessor_eq_assist_def
+      intro: iso_tuple_update_accessor_cong_assist_id)
+  with assms show ?thesis
+    by (clarsimp simp: iso_tuple_update_accessor_eq_assist_def
+      iso_tuple_snd_update_def iso_tuple_snd_def
+      iso_tuple_update_accessor_cong_assist_def iso_tuple_cons_def simps)
 qed
 
 lemma iso_tuple_cons_conj_eqI:
@@ -316,37 +352,39 @@
   by (clarsimp simp: iso_tuple_cons_def simps)
 
 lemmas intros =
-    iso_tuple_access_update_fst_fst
-    iso_tuple_access_update_snd_snd
-    iso_tuple_access_update_fst_snd
-    iso_tuple_access_update_snd_fst
-    iso_tuple_update_swap_fst_fst
-    iso_tuple_update_swap_snd_snd
-    iso_tuple_update_swap_fst_snd
-    iso_tuple_update_swap_snd_fst
-    iso_tuple_update_compose_fst_fst
-    iso_tuple_update_compose_snd_snd
-    iso_tuple_surjective_proof_assist_step
-    iso_tuple_fst_update_accessor_eq_assist
-    iso_tuple_snd_update_accessor_eq_assist
-    iso_tuple_fst_update_accessor_cong_assist
-    iso_tuple_snd_update_accessor_cong_assist
-    iso_tuple_cons_conj_eqI
+  iso_tuple_access_update_fst_fst
+  iso_tuple_access_update_snd_snd
+  iso_tuple_access_update_fst_snd
+  iso_tuple_access_update_snd_fst
+  iso_tuple_update_swap_fst_fst
+  iso_tuple_update_swap_snd_snd
+  iso_tuple_update_swap_fst_snd
+  iso_tuple_update_swap_snd_fst
+  iso_tuple_update_compose_fst_fst
+  iso_tuple_update_compose_snd_snd
+  iso_tuple_surjective_proof_assist_step
+  iso_tuple_fst_update_accessor_eq_assist
+  iso_tuple_snd_update_accessor_eq_assist
+  iso_tuple_fst_update_accessor_cong_assist
+  iso_tuple_snd_update_accessor_cong_assist
+  iso_tuple_cons_conj_eqI
 
 end
 
 lemma isomorphic_tuple_intro:
   fixes repr abst
   assumes repr_inj: "\<And>x y. repr x = repr y \<longleftrightarrow> x = y"
-     and abst_inv: "\<And>z. repr (abst z) = z"
-  assumes v: "v \<equiv> Tuple_Isomorphism repr abst"
+    and abst_inv: "\<And>z. repr (abst z) = z"
+    and v: "v \<equiv> Tuple_Isomorphism repr abst"
   shows "isomorphic_tuple v"
 proof
-  have "\<And>x. repr (abst (repr x)) = repr x"
+  fix x have "repr (abst (repr x)) = repr x"
     by (simp add: abst_inv)
-  then show "\<And>x. Record.abst v (Record.repr v x) = x"
+  then show "Record.abst v (Record.repr v x) = x"
     by (simp add: v repr_inj)
-  show P: "\<And>y. Record.repr v (Record.abst v y) = y"
+next
+  fix y
+  show "Record.repr v (Record.abst v y) = y"
     by (simp add: v) (fact abst_inv)
 qed
 
@@ -357,8 +395,7 @@
   "isomorphic_tuple tuple_iso_tuple"
   by (simp add: isomorphic_tuple_intro [OF _ _ reflexive] tuple_iso_tuple_def)
 
-lemma refl_conj_eq:
-  "Q = R \<Longrightarrow> P \<and> Q \<longleftrightarrow> P \<and> R"
+lemma refl_conj_eq: "Q = R \<Longrightarrow> P \<and> Q \<longleftrightarrow> P \<and> R"
   by simp
 
 lemma iso_tuple_UNIV_I: "x \<in> UNIV \<equiv> True"
@@ -370,50 +407,47 @@
 lemma prop_subst: "s = t \<Longrightarrow> PROP P t \<Longrightarrow> PROP P s"
   by simp
 
-lemma K_record_comp: "(\<lambda>x. c) \<circ> f = (\<lambda>x. c)" 
+lemma K_record_comp: "(\<lambda>x. c) \<circ> f = (\<lambda>x. c)"
   by (simp add: comp_def)
 
-lemma o_eq_dest_lhs:
-  "a o b = c \<Longrightarrow> a (b v) = c v"
+lemma o_eq_dest_lhs: "a o b = c \<Longrightarrow> a (b v) = c v"
   by clarsimp
 
-lemma o_eq_id_dest:
-  "a o b = id o c \<Longrightarrow> a (b v) = c v"
+lemma o_eq_id_dest: "a o b = id o c \<Longrightarrow> a (b v) = c v"
   by clarsimp
 
 
 subsection {* Concrete record syntax *}
 
 nonterminals
-  ident field_type field_types field fields update updates
+  ident field_type field_types field fields field_update field_updates
 syntax
   "_constify"           :: "id => ident"                        ("_")
   "_constify"           :: "longid => ident"                    ("_")
 
-  "_field_type"         :: "[ident, type] => field_type"        ("(2_ ::/ _)")
+  "_field_type"         :: "ident => type => field_type"        ("(2_ ::/ _)")
   ""                    :: "field_type => field_types"          ("_")
-  "_field_types"        :: "[field_type, field_types] => field_types"    ("_,/ _")
+  "_field_types"        :: "field_type => field_types => field_types"    ("_,/ _")
   "_record_type"        :: "field_types => type"                ("(3'(| _ |'))")
-  "_record_type_scheme" :: "[field_types, type] => type"        ("(3'(| _,/ (2... ::/ _) |'))")
+  "_record_type_scheme" :: "field_types => type => type"        ("(3'(| _,/ (2... ::/ _) |'))")
 
-  "_field"              :: "[ident, 'a] => field"               ("(2_ =/ _)")
+  "_field"              :: "ident => 'a => field"               ("(2_ =/ _)")
   ""                    :: "field => fields"                    ("_")
-  "_fields"             :: "[field, fields] => fields"          ("_,/ _")
+  "_fields"             :: "field => fields => fields"          ("_,/ _")
   "_record"             :: "fields => 'a"                       ("(3'(| _ |'))")
-  "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3'(| _,/ (2... =/ _) |'))")
+  "_record_scheme"      :: "fields => 'a => 'a"                 ("(3'(| _,/ (2... =/ _) |'))")
 
-  "_update_name"        :: idt
-  "_update"             :: "[ident, 'a] => update"              ("(2_ :=/ _)")
-  ""                    :: "update => updates"                  ("_")
-  "_updates"            :: "[update, updates] => updates"       ("_,/ _")
-  "_record_update"      :: "['a, updates] => 'b"                ("_/(3'(| _ |'))" [900,0] 900)
+  "_field_update"       :: "ident => 'a => field_update"        ("(2_ :=/ _)")
+  ""                    :: "field_update => field_updates"      ("_")
+  "_field_updates"      :: "field_update => field_updates => field_updates"  ("_,/ _")
+  "_record_update"      :: "'a => field_updates => 'b"          ("_/(3'(| _ |'))" [900, 0] 900)
 
 syntax (xsymbols)
   "_record_type"        :: "field_types => type"                ("(3\<lparr>_\<rparr>)")
-  "_record_type_scheme" :: "[field_types, type] => type"        ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
-  "_record"             :: "fields => 'a"                               ("(3\<lparr>_\<rparr>)")
-  "_record_scheme"      :: "[fields, 'a] => 'a"                 ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
-  "_record_update"      :: "['a, updates] => 'b"                ("_/(3\<lparr>_\<rparr>)" [900,0] 900)
+  "_record_type_scheme" :: "field_types => type => type"        ("(3\<lparr>_,/ (2\<dots> ::/ _)\<rparr>)")
+  "_record"             :: "fields => 'a"                       ("(3\<lparr>_\<rparr>)")
+  "_record_scheme"      :: "fields => 'a => 'a"                 ("(3\<lparr>_,/ (2\<dots> =/ _)\<rparr>)")
+  "_record_update"      :: "'a => field_updates => 'b"          ("_/(3\<lparr>_\<rparr>)" [900, 0] 900)
 
 
 subsection {* Record package *}
--- a/src/HOL/SMT/Examples/SMT_Examples.certs	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/SMT/Examples/SMT_Examples.certs	Wed Feb 17 09:51:46 2010 +0100
@@ -1,4 +1,4 @@
-yknPpoG47N1CXnUaEL9RvQ 133
+Fg1W6egjwo9zhhAmUXOW+w 8 0
 #2 := false
 #1 := true
 #4 := (not true)
@@ -7,8 +7,7 @@
 #20 := [asserted]: #4
 [mp #20 #22]: false
 unsat
-
-ymB2ZiCSl9aUjMXP3tIpZA 359
+2+cndY9nzS72l7VvBCGRAw 19 0
 #2 := false
 decl up_1 :: bool
 #4 := up_1
@@ -28,8 +27,7 @@
 #23 := [asserted]: #7
 [mp #23 #32]: false
 unsat
-
-XC3j0tGm4Y5klDm8sMkzVg 510
+0vJQrobUDcQ9PkGJO8aM8g 25 0
 #2 := false
 decl up_1 :: bool
 #4 := up_1
@@ -55,8 +53,7 @@
 #23 := [asserted]: #7
 [mp #23 #38]: false
 unsat
-
-y5d/Jtt47lXm/wUEvH5fHw 794
+AGGnpwEv208Vqxly7wTWHA 38 0
 #2 := false
 decl up_2 :: bool
 #5 := up_2
@@ -95,11 +92,9 @@
 #30 := [and-elim #31]: #6
 [mp #30 #52]: false
 unsat
-
-mDaukNkyA4glYbkfEOtcAw 7
-unsat
-
-TmB9YjKjdtDMIh0j9rMVPw 1530
+wakXeIy1uoPgglzOQGFhJQ 1 0
+unsat
+cpSlDe0l7plVktRNxGU5dA 71 0
 #2 := false
 decl up_1 :: bool
 #4 := up_1
@@ -171,8 +166,7 @@
 #31 := [asserted]: #15
 [mp #31 #82]: false
 unsat
-
-olufSxMlwMAAqyArYWPVOA 1300
+pg19mjJfV75T2QDrgWd4JA 57 0
 #2 := false
 decl up_1 :: bool
 #4 := up_1
@@ -230,8 +224,7 @@
 #30 := [asserted]: #14
 [mp #30 #70]: false
 unsat
-
-agKJ550QwZ1mvlK8gw+tjQ 4627
+Mj1B8X1MaN7xU/W4Kz3FoQ 194 0
 #2 := false
 decl up_1 :: bool
 #4 := up_1
@@ -426,8 +419,7 @@
 #237 := [mp #83 #236]: #75
 [mp #237 #247]: false
 unsat
-
-+R/oj2I+uFZAw/Eu+3ULdw 1160
+JkhYJB8FDavTZkizO1/9IA 52 0
 #2 := false
 decl uf_1 :: (-> T1 T1 T1)
 decl uf_2 :: T1
@@ -480,8 +472,7 @@
 #113 := [quant-inst]: #199
 [unit-resolution #113 #536 #49]: false
 unsat
-
-c67Ar5f1aFkzAZ2wYy62Wg 56943
+0ZdSZH2DbtjHNTyrDkZmXg 1667 0
 #2 := false
 decl up_54 :: bool
 #126 := up_54
@@ -2149,8 +2140,7 @@
 #2371 := [unit-resolution #891 #2369]: #166
 [unit-resolution #2159 #2371 #2370 #2358 #2357]: false
 unsat
-
-NdJwa8pysYWhn57eCXiqFA 1731
+R3pmBDBlU9XdUrxJXhj7nA 78 0
 #2 := false
 decl up_1 :: (-> int bool)
 decl ?x1!0 :: int
@@ -2229,8 +2219,7 @@
 #82 := [and-elim #81]: #55
 [unit-resolution #82 #95]: false
 unsat
-
-dWWXDEA5bUZjEafLPXbSkA 3321
+IBRj/loev6P6r0J+HOit6A 135 0
 #2 := false
 decl up_1 :: (-> T1 T2 bool)
 #5 := (:var 0 T2)
@@ -2366,9 +2355,7 @@
 #176 := [quant-inst]: #538
 [unit-resolution #176 #252 #535]: false
 unsat
-
-iGZ7b2jaCnn82lPL6oIDZA 3465
-WARNING: failed to find a pattern for quantifier (quantifier id: k!11)
+72504KVBixGB/87pOYiU/A 135 2
 #2 := false
 decl up_1 :: (-> T1 T2 bool)
 #5 := (:var 0 T2)
@@ -2504,8 +2491,9 @@
 #235 := [quant-inst]: #597
 [unit-resolution #235 #311 #594]: false
 unsat
+WARNING: failed to find a pattern for quantifier (quantifier id: k!12)
 
-eTjcfu6S5eyz+xNJ7SVluw 1246
+RaQLz4GxtUICnOD5WoYnzQ 56 0
 #2 := false
 decl up_1 :: (-> T1 bool)
 decl uf_2 :: T1
@@ -2562,8 +2550,7 @@
 #120 := [quant-inst]: #206
 [unit-resolution #120 #542 #41]: false
 unsat
-
-anG1bKU/YVTHEmc1Eh/ZXw 331
+NPQIgVPhSpgSLeS+u/EatQ 17 0
 #2 := false
 #4 := 3::int
 #5 := (= 3::int 3::int)
@@ -2581,8 +2568,7 @@
 #22 := [asserted]: #6
 [mp #22 #31]: false
 unsat
-
-lHpRCTa744ODgmii2zARAw 334
+Lc9NwVtwY2Wo0G7UbFD1oA 17 0
 #2 := false
 #4 := 3::real
 #5 := (= 3::real 3::real)
@@ -2600,8 +2586,7 @@
 #22 := [asserted]: #6
 [mp #22 #31]: false
 unsat
-
-AinXomcY4W1L/t0ZtkDhBg 524
+pYVrUflpYrrZEWALJDnvPw 26 0
 #2 := false
 #7 := 4::int
 #5 := 1::int
@@ -2628,8 +2613,7 @@
 #25 := [asserted]: #9
 [mp #25 #40]: false
 unsat
-
-WxMdOusjxqQwBPorpXBkFQ 815
+FIqzVlbN8RT0iWarmBEpjw 41 0
 #2 := false
 decl uf_1 :: int
 #4 := uf_1
@@ -2671,8 +2655,7 @@
 #28 := [asserted]: #12
 [mp #28 #52]: false
 unsat
-
-K7g37p4yZoVyQcabYS4I2w 754
+HWVNtxMa8xktQsg8pHG+1w 35 0
 #2 := false
 #5 := 3::int
 #6 := 8::int
@@ -2708,8 +2691,7 @@
 #26 := [asserted]: #10
 [mp #26 #51]: false
 unsat
-
-eCmVy21SUmWImXZDJNOfzA 6496
+M71YYpEc8u/aEIH3MOQrcg 250 0
 #2 := false
 #7 := 0::real
 decl uf_2 :: real
@@ -2960,8 +2942,7 @@
 #294 := [unit-resolution #190 #293]: #188
 [th-lemma #280 #294]: false
 unsat
-
-eBRZKSXriNPK3BNu3AWMmQ 3017
+G00bTqBjtW66EmwIZbXbOg 124 0
 #2 := false
 decl uf_1 :: (-> T1 T2)
 decl uf_3 :: T1
@@ -3086,8 +3067,7 @@
 #34 := [asserted]: #11
 [unit-resolution #34 #536]: false
 unsat
-
-CjDkjvq1e9i+SJ3L9ESARg 1146
+6QdzkSy/RtEjUu+wUKIKqA 54 0
 #2 := false
 #9 := 1::int
 decl uf_1 :: int
@@ -3142,8 +3122,7 @@
 #28 := [asserted]: #12
 [mp #28 #67]: false
 unsat
-
-nonk4MmmwlBqL8YtiJY/Qw 1339
+xoSwaSeELbR0PHe0zb/BSg 63 0
 #2 := false
 #11 := 0::int
 decl uf_2 :: int
@@ -3207,8 +3186,7 @@
 #76 := [mp #52 #75]: #63
 [mp #76 #84]: false
 unsat
-
-dCX9jxibjKl6gmr8okzk0w 727
+ciHqmDSmPpA15rO932dhvA 35 0
 #2 := false
 #6 := 5::int
 #4 := 2::int
@@ -3244,8 +3222,7 @@
 #25 := [asserted]: #9
 [mp #25 #49]: false
 unsat
-
-/kLzs8f/jQjEM38PdppYPA 912
+HzwFy7SRHqpspkYnzyeF4w 45 0
 #2 := false
 #11 := 4::real
 decl uf_2 :: real
@@ -3291,8 +3268,7 @@
 #60 := [mp #36 #59]: #51
 [th-lemma #60 #47 #38]: false
 unsat
-
-iT8vKYi503k30rQLczD7yw 1245
+XW7QIWmzYjfQXaHHPc98eA 59 0
 #2 := false
 #16 := (not false)
 decl uf_2 :: int
@@ -3352,8 +3328,94 @@
 #34 := [asserted]: #18
 [mp #34 #71]: false
 unsat
-
-6R4JcV7tL9QRH7WWPAKM5g 5413
+ZGL00TLLioiLlWFiXUnbxg 86 0
+#2 := false
+decl uf_1 :: int
+#5 := uf_1
+#7 := 2::int
+#29 := (* 2::int uf_1)
+#4 := 0::int
+#54 := (= 0::int #29)
+#55 := (not #54)
+#61 := (= #29 0::int)
+#104 := (not #61)
+#110 := (iff #104 #55)
+#108 := (iff #61 #54)
+#109 := [commutativity]: #108
+#111 := [monotonicity #109]: #110
+#62 := (<= #29 0::int)
+#100 := (not #62)
+#30 := (<= uf_1 0::int)
+#31 := (not #30)
+#6 := (< 0::int uf_1)
+#32 := (iff #6 #31)
+#33 := [rewrite]: #32
+#27 := [asserted]: #6
+#34 := [mp #27 #33]: #31
+#101 := (or #100 #30)
+#102 := [th-lemma]: #101
+#103 := [unit-resolution #102 #34]: #100
+#105 := (or #104 #62)
+#106 := [th-lemma]: #105
+#107 := [unit-resolution #106 #103]: #104
+#112 := [mp #107 #111]: #55
+#56 := (= uf_1 #29)
+#57 := (not #56)
+#53 := (= 0::int uf_1)
+#50 := (not #53)
+#58 := (and #50 #55 #57)
+#69 := (not #58)
+#42 := (distinct 0::int uf_1 #29)
+#47 := (not #42)
+#9 := (- uf_1 uf_1)
+#8 := (* uf_1 2::int)
+#10 := (distinct uf_1 #8 #9)
+#11 := (not #10)
+#48 := (iff #11 #47)
+#45 := (iff #10 #42)
+#39 := (distinct uf_1 #29 0::int)
+#43 := (iff #39 #42)
+#44 := [rewrite]: #43
+#40 := (iff #10 #39)
+#37 := (= #9 0::int)
+#38 := [rewrite]: #37
+#35 := (= #8 #29)
+#36 := [rewrite]: #35
+#41 := [monotonicity #36 #38]: #40
+#46 := [trans #41 #44]: #45
+#49 := [monotonicity #46]: #48
+#28 := [asserted]: #11
+#52 := [mp #28 #49]: #47
+#80 := (or #42 #69)
+#81 := [def-axiom]: #80
+#82 := [unit-resolution #81 #52]: #69
+#59 := (= uf_1 0::int)
+#83 := (not #59)
+#89 := (iff #83 #50)
+#87 := (iff #59 #53)
+#88 := [commutativity]: #87
+#90 := [monotonicity #88]: #89
+#84 := (or #83 #30)
+#85 := [th-lemma]: #84
+#86 := [unit-resolution #85 #34]: #83
+#91 := [mp #86 #90]: #50
+#64 := -1::int
+#65 := (* -1::int #29)
+#66 := (+ uf_1 #65)
+#68 := (>= #66 0::int)
+#92 := (not #68)
+#93 := (or #92 #30)
+#94 := [th-lemma]: #93
+#95 := [unit-resolution #94 #34]: #92
+#96 := (or #57 #68)
+#97 := [th-lemma]: #96
+#98 := [unit-resolution #97 #95]: #57
+#76 := (or #58 #53 #54 #56)
+#77 := [def-axiom]: #76
+#99 := [unit-resolution #77 #98 #91 #82]: #54
+[unit-resolution #99 #112]: false
+unsat
+DWt5rIK6NWlI4vrw+691Zg 212 0
 #2 := false
 decl uf_4 :: T1
 #13 := uf_4
@@ -3566,96 +3628,7 @@
 #519 := [unit-resolution #521 #518]: #547
 [unit-resolution #519 #522]: false
 unsat
-
-eOXl5Nf4A2Sq4Q+Wh5XNfA 2026
-#2 := false
-decl uf_1 :: int
-#5 := uf_1
-#7 := 2::int
-#29 := (* 2::int uf_1)
-#4 := 0::int
-#54 := (= 0::int #29)
-#55 := (not #54)
-#61 := (= #29 0::int)
-#104 := (not #61)
-#110 := (iff #104 #55)
-#108 := (iff #61 #54)
-#109 := [commutativity]: #108
-#111 := [monotonicity #109]: #110
-#62 := (<= #29 0::int)
-#100 := (not #62)
-#30 := (<= uf_1 0::int)
-#31 := (not #30)
-#6 := (< 0::int uf_1)
-#32 := (iff #6 #31)
-#33 := [rewrite]: #32
-#27 := [asserted]: #6
-#34 := [mp #27 #33]: #31
-#101 := (or #100 #30)
-#102 := [th-lemma]: #101
-#103 := [unit-resolution #102 #34]: #100
-#105 := (or #104 #62)
-#106 := [th-lemma]: #105
-#107 := [unit-resolution #106 #103]: #104
-#112 := [mp #107 #111]: #55
-#56 := (= uf_1 #29)
-#57 := (not #56)
-#53 := (= 0::int uf_1)
-#50 := (not #53)
-#58 := (and #50 #55 #57)
-#69 := (not #58)
-#42 := (distinct 0::int uf_1 #29)
-#47 := (not #42)
-#9 := (- uf_1 uf_1)
-#8 := (* uf_1 2::int)
-#10 := (distinct uf_1 #8 #9)
-#11 := (not #10)
-#48 := (iff #11 #47)
-#45 := (iff #10 #42)
-#39 := (distinct uf_1 #29 0::int)
-#43 := (iff #39 #42)
-#44 := [rewrite]: #43
-#40 := (iff #10 #39)
-#37 := (= #9 0::int)
-#38 := [rewrite]: #37
-#35 := (= #8 #29)
-#36 := [rewrite]: #35
-#41 := [monotonicity #36 #38]: #40
-#46 := [trans #41 #44]: #45
-#49 := [monotonicity #46]: #48
-#28 := [asserted]: #11
-#52 := [mp #28 #49]: #47
-#80 := (or #42 #69)
-#81 := [def-axiom]: #80
-#82 := [unit-resolution #81 #52]: #69
-#59 := (= uf_1 0::int)
-#83 := (not #59)
-#89 := (iff #83 #50)
-#87 := (iff #59 #53)
-#88 := [commutativity]: #87
-#90 := [monotonicity #88]: #89
-#84 := (or #83 #30)
-#85 := [th-lemma]: #84
-#86 := [unit-resolution #85 #34]: #83
-#91 := [mp #86 #90]: #50
-#64 := -1::int
-#65 := (* -1::int #29)
-#66 := (+ uf_1 #65)
-#68 := (>= #66 0::int)
-#92 := (not #68)
-#93 := (or #92 #30)
-#94 := [th-lemma]: #93
-#95 := [unit-resolution #94 #34]: #92
-#96 := (or #57 #68)
-#97 := [th-lemma]: #96
-#98 := [unit-resolution #97 #95]: #57
-#76 := (or #58 #53 #54 #56)
-#77 := [def-axiom]: #76
-#99 := [unit-resolution #77 #98 #91 #82]: #54
-[unit-resolution #99 #112]: false
-unsat
-
-TwJgkTydtls9Q94iw4a3jw 17377
+PaSeDRf7Set5ywlblDOoTg 673 0
 #2 := false
 #169 := 0::int
 decl uf_2 :: int
@@ -4329,8 +4302,7 @@
 #410 := [mp #349 #409]: #405
 [unit-resolution #410 #710 #709 #697 #706]: false
 unsat
-
-ib5n9nvBAC5jXuKIpV/54g 82870
+U7jSPEM53XYq3qs03aUczw 2291 0
 #2 := false
 #6 := 0::int
 decl z3name!0 :: int
@@ -6622,8 +6594,7 @@
 #2323 := [unit-resolution #918 #2322]: #100
 [unit-resolution #917 #2323 #2318]: false
 unsat
-
-SqgPFdiZeq8SOFuTISQN5g 1109
+eqE7IAqFr0UIBuUsVDgHvw 52 0
 #2 := false
 #8 := 1::real
 decl uf_1 :: real
@@ -6676,26 +6647,15 @@
 #29 := [asserted]: #13
 [mp #29 #65]: false
 unsat
-
-rOkYPs+Q++z/M3OPR/88JQ 1654
+ADs4ZPiuUr7Xu7tk71NnEw 59 0
 #2 := false
 #55 := 0::int
 #7 := 2::int
 decl uf_1 :: int
 #4 := uf_1
 #8 := (mod uf_1 2::int)
-#67 := (>= #8 0::int)
-#1 := true
-#156 := (iff #67 true)
-#158 := (iff #156 #67)
-#159 := [rewrite]: #158
-#28 := [true-axiom]: true
-#153 := (or false #67)
-#154 := [th-lemma]: #153
-#155 := [unit-resolution #154 #28]: #67
-#157 := [iff-true #155]: #156
-#160 := [mp #157 #159]: #67
-#70 := (not #67)
+#58 := (>= #8 0::int)
+#61 := (not #58)
 #5 := 1::int
 #9 := (* 2::int #8)
 #10 := (+ #9 1::int)
@@ -6703,35 +6663,24 @@
 #6 := (+ uf_1 1::int)
 #12 := (<= #6 #11)
 #13 := (not #12)
-#77 := (iff #13 #70)
+#66 := (iff #13 #61)
 #39 := (+ uf_1 #9)
 #40 := (+ 1::int #39)
 #30 := (+ 1::int uf_1)
 #45 := (<= #30 #40)
 #48 := (not #45)
-#75 := (iff #48 #70)
-#58 := (* 1::int #8)
-#59 := (>= #58 0::int)
-#62 := (not #59)
-#71 := (iff #62 #70)
-#68 := (iff #59 #67)
-#65 := (= #58 #8)
-#66 := [rewrite]: #65
-#69 := [monotonicity #66]: #68
-#72 := [monotonicity #69]: #71
-#73 := (iff #48 #62)
+#64 := (iff #48 #61)
 #56 := (>= #9 0::int)
 #51 := (not #56)
-#63 := (iff #51 #62)
-#60 := (iff #56 #59)
-#61 := [rewrite]: #60
-#64 := [monotonicity #61]: #63
+#62 := (iff #51 #61)
+#59 := (iff #56 #58)
+#60 := [rewrite]: #59
+#63 := [monotonicity #60]: #62
 #52 := (iff #48 #51)
 #53 := (iff #45 #56)
 #54 := [rewrite]: #53
 #57 := [monotonicity #54]: #52
-#74 := [trans #57 #64]: #73
-#76 := [trans #74 #72]: #75
+#65 := [trans #57 #63]: #64
 #49 := (iff #13 #48)
 #46 := (iff #12 #45)
 #43 := (= #11 #40)
@@ -6748,38 +6697,39 @@
 #32 := [rewrite]: #31
 #47 := [monotonicity #32 #44]: #46
 #50 := [monotonicity #47]: #49
-#78 := [trans #50 #76]: #77
+#67 := [trans #50 #65]: #66
 #29 := [asserted]: #13
-#79 := [mp #29 #78]: #70
-[unit-resolution #79 #160]: false
-unsat
-
-oSBTeiF6GKDeHPsMxXHXtQ 1064
+#68 := [mp #29 #67]: #61
+#1 := true
+#28 := [true-axiom]: true
+#142 := (or false #58)
+#143 := [th-lemma]: #142
+#144 := [unit-resolution #143 #28]: #58
+[unit-resolution #144 #68]: false
+unsat
+x2NmsblNl28xPXP2EG22rA 54 0
 #2 := false
 #5 := 2::int
 decl uf_1 :: int
 #4 := uf_1
 #6 := (mod uf_1 2::int)
-#122 := (>= #6 2::int)
-#123 := (not #122)
-#1 := true
-#27 := [true-axiom]: true
-#133 := (or false #123)
-#134 := [th-lemma]: #133
-#135 := [unit-resolution #134 #27]: #123
+#55 := (>= #6 2::int)
 #9 := 3::int
-#29 := (* 2::int #6)
-#48 := (>= #29 3::int)
 #10 := (+ uf_1 3::int)
 #7 := (+ #6 #6)
 #8 := (+ uf_1 #7)
 #11 := (< #8 #10)
 #12 := (not #11)
-#55 := (iff #12 #48)
+#60 := (iff #12 #55)
 #35 := (+ 3::int uf_1)
+#29 := (* 2::int #6)
 #32 := (+ uf_1 #29)
 #38 := (< #32 #35)
 #41 := (not #38)
+#58 := (iff #41 #55)
+#48 := (>= #29 3::int)
+#56 := (iff #48 #55)
+#57 := [rewrite]: #56
 #53 := (iff #41 #48)
 #46 := (not #48)
 #45 := (not #46)
@@ -6790,6 +6740,7 @@
 #44 := [rewrite]: #47
 #50 := [monotonicity #44]: #49
 #54 := [trans #50 #52]: #53
+#59 := [trans #54 #57]: #58
 #42 := (iff #12 #41)
 #39 := (iff #11 #38)
 #36 := (= #10 #35)
@@ -6800,13 +6751,18 @@
 #34 := [monotonicity #31]: #33
 #40 := [monotonicity #34 #37]: #39
 #43 := [monotonicity #40]: #42
-#56 := [trans #43 #54]: #55
+#61 := [trans #43 #59]: #60
 #28 := [asserted]: #12
-#57 := [mp #28 #56]: #48
-[th-lemma #57 #135]: false
-unsat
-
-hqH9yBHvmZgES3pBkvsqVQ 2715
+#62 := [mp #28 #61]: #55
+#127 := (not #55)
+#1 := true
+#27 := [true-axiom]: true
+#137 := (or false #127)
+#138 := [th-lemma]: #137
+#139 := [unit-resolution #138 #27]: #127
+[unit-resolution #139 #62]: false
+unsat
+kfLiOGBz3RZx9wt+FS+hfg 118 0
 #2 := false
 #5 := 0::real
 decl uf_1 :: real
@@ -6925,8 +6881,7 @@
 #126 := [mp #101 #125]: #115
 [unit-resolution #126 #132 #129]: false
 unsat
-
-ar4IlNF9IylWgGSPOf9paw 5159
+FPTJq9aN3ES4iIrHgaTv+A 208 0
 #2 := false
 #9 := 0::int
 #11 := 4::int
@@ -7135,8 +7090,7 @@
 #418 := [unit-resolution #417 #36]: #298
 [th-lemma #418 #415 #412 #411 #99 #400 #397 #396]: false
 unsat
-
-o9WM91Nq0O5f08PEA0qA6w 515
+yN0aj3KferzvOSp2KlyNwg 24 0
 #2 := false
 #4 := (exists (vars (?x1 int)) false)
 #5 := (not #4)
@@ -7161,8 +7115,7 @@
 #22 := [asserted]: #6
 [mp #22 #38]: false
 unsat
-
-SJxvvXYE4z1G4iLRBCPerg 516
+7iMPasu6AIeHm45slLCByA 24 0
 #2 := false
 #4 := (exists (vars (?x1 real)) false)
 #5 := (not #4)
@@ -7187,17 +7140,13 @@
 #22 := [asserted]: #6
 [mp #22 #38]: false
 unsat
-
-Fr3hfDrqfMuGDpDYbXAHwg 7
-unsat
-
-bFuFCRUoQSRWyp0iCwo+PA 7
-unsat
-
-NJEgv3p5NO4/yEATNBBDNw 7
-unsat
-
-RC1LWjyqEEsh1xhocCgPmQ 1633
+cv2pC2I0gIUYtVwtXngvXg 1 0
+unsat
+4r8/IxBBDH1ZqF0YfzLLTg 1 0
+unsat
+uj7n+C4nG462DNJy9Divrg 1 0
+unsat
+dn/LVwy1BXEOmtqdUBNhLw 73 0
 #2 := false
 #5 := 0::int
 #8 := 1::int
@@ -7271,8 +7220,7 @@
 #144 := [trans #142 #82]: #143
 [mp #144 #146]: false
 unsat
-
-6pnpFuE9SN6Jr5Upml9T0A 1896
+VzZ1W5SEEis1AJp1qZz86g 82 0
 #2 := false
 #5 := (:var 0 int)
 #7 := 0::int
@@ -7355,100 +7303,52 @@
 #30 := [asserted]: #14
 [mp #30 #96]: false
 unsat
-
-sHpY0IFBgZUNhP56aRB+/w 2968
+UoXgZh5LkmyNCmQEfEtnig 78 0
 #2 := false
-#9 := 1::int
-decl ?x1!1 :: int
-#91 := ?x1!1
-#68 := -2::int
-#129 := (* -2::int ?x1!1)
-decl ?x2!0 :: int
-#90 := ?x2!0
+#5 := (:var 0 int)
 #7 := 2::int
-#128 := (* 2::int ?x2!0)
-#130 := (+ #128 #129)
-#131 := (<= #130 1::int)
-#136 := (not #131)
-#55 := 0::int
-#53 := -1::int
-#115 := (* -1::int ?x1!1)
-#116 := (+ ?x2!0 #115)
-#117 := (<= #116 0::int)
-#139 := (or #117 #136)
-#142 := (not #139)
-#92 := (* -2::int ?x2!0)
-#93 := (* 2::int ?x1!1)
-#94 := (+ #93 #92)
-#95 := (>= #94 -1::int)
-#96 := (not #95)
-#97 := (* -1::int ?x2!0)
-#98 := (+ ?x1!1 #97)
-#99 := (>= #98 0::int)
-#100 := (or #99 #96)
-#101 := (not #100)
-#143 := (iff #101 #142)
-#140 := (iff #100 #139)
-#137 := (iff #96 #136)
-#134 := (iff #95 #131)
-#122 := (+ #92 #93)
-#125 := (>= #122 -1::int)
-#132 := (iff #125 #131)
-#133 := [rewrite]: #132
-#126 := (iff #95 #125)
-#123 := (= #94 #122)
-#124 := [rewrite]: #123
-#127 := [monotonicity #124]: #126
-#135 := [trans #127 #133]: #134
-#138 := [monotonicity #135]: #137
-#120 := (iff #99 #117)
-#109 := (+ #97 ?x1!1)
-#112 := (>= #109 0::int)
-#118 := (iff #112 #117)
-#119 := [rewrite]: #118
-#113 := (iff #99 #112)
-#110 := (= #98 #109)
-#111 := [rewrite]: #110
-#114 := [monotonicity #111]: #113
-#121 := [trans #114 #119]: #120
-#141 := [monotonicity #121 #138]: #140
-#144 := [monotonicity #141]: #143
-#5 := (:var 0 int)
-#71 := (* -2::int #5)
+#11 := (* 2::int #5)
+#9 := 1::int
 #4 := (:var 1 int)
 #8 := (* 2::int #4)
-#72 := (+ #8 #71)
-#70 := (>= #72 -1::int)
-#69 := (not #70)
-#57 := (* -1::int #5)
-#58 := (+ #4 #57)
-#56 := (>= #58 0::int)
-#75 := (or #56 #69)
-#78 := (forall (vars (?x1 int) (?x2 int)) #75)
-#81 := (not #78)
-#102 := (~ #81 #101)
-#103 := [sk]: #102
-#11 := (* 2::int #5)
 #10 := (+ #8 1::int)
 #12 := (< #10 #11)
 #6 := (< #4 #5)
 #13 := (implies #6 #12)
 #14 := (forall (vars (?x1 int) (?x2 int)) #13)
 #15 := (not #14)
-#84 := (iff #15 #81)
+#91 := (iff #15 false)
 #32 := (+ 1::int #8)
 #35 := (< #32 #11)
 #41 := (not #6)
 #42 := (or #41 #35)
 #47 := (forall (vars (?x1 int) (?x2 int)) #42)
 #50 := (not #47)
-#82 := (iff #50 #81)
-#79 := (iff #47 #78)
-#76 := (iff #42 #75)
-#73 := (iff #35 #69)
+#89 := (iff #50 false)
+#1 := true
+#84 := (not true)
+#87 := (iff #84 false)
+#88 := [rewrite]: #87
+#85 := (iff #50 #84)
+#82 := (iff #47 true)
+#77 := (forall (vars (?x1 int) (?x2 int)) true)
+#80 := (iff #77 true)
+#81 := [elim-unused]: #80
+#78 := (iff #47 #77)
+#75 := (iff #42 true)
+#55 := 0::int
+#53 := -1::int
+#57 := (* -1::int #5)
+#58 := (+ #4 #57)
+#56 := (>= #58 0::int)
+#54 := (not #56)
+#69 := (or #56 #54)
+#73 := (iff #69 true)
 #74 := [rewrite]: #73
+#71 := (iff #42 #69)
+#70 := (iff #35 #54)
+#68 := [rewrite]: #70
 #66 := (iff #41 #56)
-#54 := (not #56)
 #61 := (not #54)
 #64 := (iff #61 #56)
 #65 := [rewrite]: #64
@@ -7457,9 +7357,12 @@
 #60 := [rewrite]: #59
 #63 := [monotonicity #60]: #62
 #67 := [trans #63 #65]: #66
-#77 := [monotonicity #67 #74]: #76
-#80 := [quant-intro #77]: #79
-#83 := [monotonicity #80]: #82
+#72 := [monotonicity #67 #68]: #71
+#76 := [trans #72 #74]: #75
+#79 := [quant-intro #76]: #78
+#83 := [trans #79 #81]: #82
+#86 := [monotonicity #83]: #85
+#90 := [trans #86 #88]: #89
 #51 := (iff #15 #50)
 #48 := (iff #14 #47)
 #45 := (iff #13 #42)
@@ -7475,89 +7378,53 @@
 #46 := [trans #40 #44]: #45
 #49 := [quant-intro #46]: #48
 #52 := [monotonicity #49]: #51
-#85 := [trans #52 #83]: #84
+#92 := [trans #52 #90]: #91
 #31 := [asserted]: #15
-#86 := [mp #31 #85]: #81
-#106 := [mp~ #86 #103]: #101
-#107 := [mp #106 #144]: #142
-#146 := [not-or-elim #107]: #131
-#108 := (not #117)
-#145 := [not-or-elim #107]: #108
-[th-lemma #145 #146]: false
-unsat
-
-7GSX+dyn9XwHWNcjJ4X1ww 2292
+[mp #31 #92]: false
+unsat
+Qv4gVhCmOzC39uufV9ZpDA 61 0
 #2 := false
-#7 := 1::int
-decl ?x1!1 :: int
-#74 := ?x1!1
-#51 := -2::int
-#96 := (* -2::int ?x1!1)
-decl ?x2!0 :: int
-#73 := ?x2!0
+#9 := (:var 0 int)
 #4 := 2::int
-#95 := (* 2::int ?x2!0)
-#97 := (+ #95 #96)
-#166 := (<= #97 1::int)
-#94 := (= #97 1::int)
-#53 := -1::int
-#75 := (* -2::int ?x2!0)
-#76 := (* 2::int ?x1!1)
-#77 := (+ #76 #75)
-#78 := (= #77 -1::int)
-#79 := (not #78)
-#80 := (not #79)
-#110 := (iff #80 #94)
-#102 := (not #94)
-#105 := (not #102)
-#108 := (iff #105 #94)
-#109 := [rewrite]: #108
-#106 := (iff #80 #105)
-#103 := (iff #79 #102)
-#100 := (iff #78 #94)
-#88 := (+ #75 #76)
-#91 := (= #88 -1::int)
-#98 := (iff #91 #94)
-#99 := [rewrite]: #98
-#92 := (iff #78 #91)
-#89 := (= #77 #88)
-#90 := [rewrite]: #89
-#93 := [monotonicity #90]: #92
-#101 := [trans #93 #99]: #100
-#104 := [monotonicity #101]: #103
-#107 := [monotonicity #104]: #106
-#111 := [trans #107 #109]: #110
-#9 := (:var 0 int)
-#55 := (* -2::int #9)
+#10 := (* 2::int #9)
+#7 := 1::int
 #5 := (:var 1 int)
 #6 := (* 2::int #5)
-#56 := (+ #6 #55)
-#54 := (= #56 -1::int)
-#58 := (not #54)
-#61 := (forall (vars (?x1 int) (?x2 int)) #58)
-#64 := (not #61)
-#81 := (~ #64 #80)
-#82 := [sk]: #81
-#10 := (* 2::int #9)
 #8 := (+ #6 1::int)
 #11 := (= #8 #10)
 #12 := (not #11)
 #13 := (forall (vars (?x1 int) (?x2 int)) #12)
 #14 := (not #13)
-#67 := (iff #14 #64)
+#74 := (iff #14 false)
 #31 := (+ 1::int #6)
 #37 := (= #10 #31)
 #42 := (not #37)
 #45 := (forall (vars (?x1 int) (?x2 int)) #42)
 #48 := (not #45)
-#65 := (iff #48 #64)
-#62 := (iff #45 #61)
-#59 := (iff #42 #58)
-#52 := (iff #37 #54)
-#57 := [rewrite]: #52
-#60 := [monotonicity #57]: #59
-#63 := [quant-intro #60]: #62
-#66 := [monotonicity #63]: #65
+#72 := (iff #48 false)
+#1 := true
+#67 := (not true)
+#70 := (iff #67 false)
+#71 := [rewrite]: #70
+#68 := (iff #48 #67)
+#65 := (iff #45 true)
+#60 := (forall (vars (?x1 int) (?x2 int)) true)
+#63 := (iff #60 true)
+#64 := [elim-unused]: #63
+#61 := (iff #45 #60)
+#58 := (iff #42 true)
+#51 := (not false)
+#56 := (iff #51 true)
+#57 := [rewrite]: #56
+#52 := (iff #42 #51)
+#53 := (iff #37 false)
+#54 := [rewrite]: #53
+#55 := [monotonicity #54]: #52
+#59 := [trans #55 #57]: #58
+#62 := [quant-intro #59]: #61
+#66 := [trans #62 #64]: #65
+#69 := [monotonicity #66]: #68
+#73 := [trans #69 #71]: #72
 #49 := (iff #14 #48)
 #46 := (iff #13 #45)
 #43 := (iff #12 #42)
@@ -7573,22 +7440,11 @@
 #44 := [monotonicity #41]: #43
 #47 := [quant-intro #44]: #46
 #50 := [monotonicity #47]: #49
-#68 := [trans #50 #66]: #67
+#75 := [trans #50 #73]: #74
 #30 := [asserted]: #14
-#69 := [mp #30 #68]: #64
-#85 := [mp~ #69 #82]: #80
-#86 := [mp #85 #111]: #94
-#168 := (or #102 #166)
-#169 := [th-lemma]: #168
-#170 := [unit-resolution #169 #86]: #166
-#167 := (>= #97 1::int)
-#171 := (or #102 #167)
-#172 := [th-lemma]: #171
-#173 := [unit-resolution #172 #86]: #167
-[th-lemma #173 #170]: false
-unsat
-
-oieid3+1h5s04LTQ9d796Q 2636
+[mp #30 #75]: false
+unsat
++j+tSj7aUImWej2XcTL9dw 111 0
 #2 := false
 #4 := 2::int
 decl ?x1!1 :: int
@@ -7700,8 +7556,7 @@
 #184 := [th-lemma]: #183
 [unit-resolution #184 #127 #125 #126]: false
 unsat
-
-+RiWXCcHPvuSeYUjZ4Ky/g 2113
+kQRsBd9oowc7exsvsEgTUg 89 0
 #2 := false
 #4 := 0::int
 decl ?x1!0 :: int
@@ -7791,9 +7646,7 @@
 #167 := [unit-resolution #154 #90]: #166
 [unit-resolution #167 #165 #162]: false
 unsat
-
-hlG1vHDJCcXbyvxKYDWifg 2036
-WARNING: failed to find a pattern for quantifier (quantifier id: k!1)
+VPjD8BtzPcTZKIRT4SA3Nw 83 2
 #2 := false
 #5 := 0::int
 #4 := (:var 0 int)
@@ -7877,9 +7730,9 @@
 #62 := [mp~ #54 #61]: #49
 [unit-resolution #62 #174]: false
 unsat
+WARNING: failed to find a pattern for quantifier (quantifier id: k!2)
 
-oOC8ghGUYboMezGio2exAg 4464
-WARNING: failed to find a pattern for quantifier (quantifier id: k!1)
+DCV5zpDW3cC2A61VghqFkA 180 2
 #2 := false
 #4 := 0::int
 #5 := (:var 0 int)
@@ -8060,63 +7913,40 @@
 #585 := [unit-resolution #128 #581]: #55
 [unit-resolution #585 #307]: false
 unsat
+WARNING: failed to find a pattern for quantifier (quantifier id: k!2)
 
-4Dtb5Y1TTRPWZcHG9Griog 2407
+lYXJpXHB9nLXJbOsr9VH1w 68 0
 #2 := false
-#104 := -1::int
-decl ?x1!1 :: int
-#86 := ?x1!1
-#106 := -4::int
-#107 := (* -4::int ?x1!1)
-decl ?x2!0 :: int
-#85 := ?x2!0
-#7 := 6::int
-#105 := (* 6::int ?x2!0)
-#108 := (+ #105 #107)
-#168 := (<= #108 -1::int)
-#109 := (= #108 -1::int)
 #12 := 1::int
-#33 := -6::int
-#87 := (* -6::int ?x2!0)
-#4 := 4::int
-#88 := (* 4::int ?x1!1)
-#89 := (+ #88 #87)
-#90 := (= #89 1::int)
-#112 := (iff #90 #109)
-#98 := (+ #87 #88)
-#101 := (= #98 1::int)
-#110 := (iff #101 #109)
-#111 := [rewrite]: #110
-#102 := (iff #90 #101)
-#99 := (= #89 #98)
-#100 := [rewrite]: #99
-#103 := [monotonicity #100]: #102
-#113 := [trans #103 #111]: #112
-#53 := (:var 0 int)
-#54 := (* -6::int #53)
 #9 := (:var 1 int)
-#55 := (* 4::int #9)
-#56 := (+ #55 #54)
-#76 := (= #56 1::int)
-#74 := (exists (vars (?x1 int) (?x2 int)) #76)
-#91 := (~ #74 #90)
-#92 := [sk]: #91
+#7 := 6::int
 #8 := (- 6::int)
 #10 := (* #8 #9)
 #5 := (:var 2 int)
+#4 := 4::int
 #6 := (* 4::int #5)
 #11 := (+ #6 #10)
 #13 := (= #11 1::int)
 #14 := (exists (vars (?x1 int) (?x2 int) (?x3 int)) #13)
 #15 := (not #14)
 #16 := (not #15)
-#79 := (iff #16 #74)
+#82 := (iff #16 false)
+#53 := (:var 0 int)
+#33 := -6::int
+#54 := (* -6::int #53)
+#55 := (* 4::int #9)
+#56 := (+ #55 #54)
 #57 := (= 1::int #56)
 #58 := (exists (vars (?x1 int) (?x2 int)) #57)
-#77 := (iff #58 #74)
-#75 := (iff #57 #76)
-#73 := [rewrite]: #75
-#78 := [quant-intro #73]: #77
+#80 := (iff #58 false)
+#76 := (exists (vars (?x1 int) (?x2 int)) false)
+#78 := (iff #76 false)
+#79 := [elim-unused]: #78
+#77 := (iff #58 #76)
+#73 := (iff #57 false)
+#74 := [rewrite]: #73
+#75 := [quant-intro #74]: #77
+#81 := [trans #75 #79]: #80
 #71 := (iff #16 #58)
 #63 := (not #58)
 #66 := (not #63)
@@ -8150,23 +7980,11 @@
 #65 := [monotonicity #62]: #64
 #68 := [monotonicity #65]: #67
 #72 := [trans #68 #70]: #71
-#80 := [trans #72 #78]: #79
+#83 := [trans #72 #81]: #82
 #32 := [asserted]: #16
-#81 := [mp #32 #80]: #74
-#95 := [mp~ #81 #92]: #90
-#96 := [mp #95 #113]: #109
-#170 := (not #109)
-#171 := (or #170 #168)
-#172 := [th-lemma]: #171
-#173 := [unit-resolution #172 #96]: #168
-#169 := (>= #108 -1::int)
-#174 := (or #170 #169)
-#175 := [th-lemma]: #174
-#176 := [unit-resolution #175 #96]: #169
-[th-lemma #176 #173]: false
-unsat
-
-dbOre63OdVavsqL3lG2ttw 2516
+[mp #32 #83]: false
+unsat
+jNvpOd8qnh73F8B6mQDrRw 107 0
 #2 := false
 #4 := 0::int
 decl ?x2!1 :: int
@@ -8274,8 +8092,7 @@
 #123 := [and-elim #101]: #88
 [th-lemma #123 #124 #125]: false
 unsat
-
-LtM5szEGH9QAF1TwsVtH4w 2764
+QWWPBUGjgvTCpxqJ9oPGdQ 117 0
 #2 := false
 #4 := 0::int
 decl ?x2!1 :: int
@@ -8393,8 +8210,7 @@
 #188 := [unit-resolution #187 #110]: #98
 [unit-resolution #188 #130]: false
 unsat
-
-ibIqbnIUB+oyERADdbFW6w 3631
+3r4MsKEvDJc1RWnNRxu/3Q 148 0
 #2 := false
 #144 := (not false)
 #7 := 0::int
@@ -8543,8 +8359,7 @@
 #158 := [mp #126 #157]: #153
 [mp #158 #181]: false
 unsat
-
-1HbJvLWS/aG8IZEVLDIWyA 1506
+Q+cnHyqIFLGWsSlQkp3fEg 67 0
 #2 := false
 #4 := (:var 0 int)
 #5 := (pattern #4)
@@ -8581,12 +8396,12 @@
 #46 := (+ #4 #45)
 #44 := (>= #46 0::int)
 #42 := (not #44)
-#60 := (or #44 #42)
-#61 := (iff #60 true)
+#57 := (or #44 #42)
+#61 := (iff #57 true)
 #62 := [rewrite]: #61
-#56 := (iff #32 #60)
+#59 := (iff #32 #57)
 #58 := (iff #11 #42)
-#59 := [rewrite]: #58
+#56 := [rewrite]: #58
 #54 := (iff #31 #44)
 #49 := (not #42)
 #52 := (iff #49 #44)
@@ -8596,8 +8411,8 @@
 #48 := [rewrite]: #47
 #51 := [monotonicity #48]: #50
 #55 := [trans #51 #53]: #54
-#57 := [monotonicity #55 #59]: #56
-#64 := [trans #57 #62]: #63
+#60 := [monotonicity #55 #56]: #59
+#64 := [trans #60 #62]: #63
 #67 := [quant-intro #64]: #66
 #71 := [trans #67 #69]: #70
 #74 := [monotonicity #71]: #73
@@ -8612,11 +8427,9 @@
 #30 := [asserted]: #14
 [mp #30 #80]: false
 unsat
-
-K7kWge9RT44bPFRy+hxaqg 7
-unsat
-
-+rwKUm5bOzD9paEkmogLyw 1562
+Q7HDzu4ER2dw+lHHM6YgFg 1 0
+unsat
+saejIG5KeeVxOolEIo3gtw 75 0
 #2 := false
 #6 := 1::int
 decl uf_3 :: int
@@ -8692,8 +8505,7 @@
 #32 := [asserted]: #16
 [mp #32 #86]: false
 unsat
-
-iRJ30NP1Enylq9tZfpCPTA 1288
+PPaoU5CzQFYr3LRpOsGPhQ 62 0
 #2 := false
 decl uf_2 :: real
 #6 := uf_2
@@ -8756,8 +8568,7 @@
 #32 := [asserted]: #16
 [mp #32 #74]: false
 unsat
-
-Ff1vqDiuUnet19j/x+mXkA 3029
+hXKzem5+KYZMOj+GKxjszQ 141 0
 #2 := false
 decl uf_4 :: int
 #9 := uf_4
@@ -8899,8 +8710,7 @@
 #45 := [asserted]: #29
 [mp #45 #150]: false
 unsat
-
-iPZ87GNdi7uQw4ehEpbTPQ 7012
+3D8WhjZTO7T824d7mwXcCA 252 0
 #2 := false
 #9 := 0::int
 decl uf_2 :: (-> T1 int)
@@ -8915,19 +8725,19 @@
 #295 := -1::int
 #274 := (* -1::int #293)
 #610 := (+ #24 #274)
-#315 := (<= #610 0::int)
+#594 := (<= #610 0::int)
 #612 := (= #610 0::int)
-#255 := (>= #23 0::int)
-#317 := (= #293 0::int)
-#522 := (not #317)
-#577 := (<= #293 0::int)
-#519 := (not #577)
+#606 := (>= #23 0::int)
+#237 := (= #293 0::int)
+#549 := (not #237)
+#588 := (<= #293 0::int)
+#457 := (not #588)
 #26 := 1::int
-#553 := (>= #293 1::int)
-#552 := (= #293 1::int)
+#558 := (>= #293 1::int)
+#555 := (= #293 1::int)
 #27 := (uf_1 1::int)
-#420 := (uf_2 #27)
-#565 := (= #420 1::int)
+#589 := (uf_2 #27)
+#301 := (= #589 1::int)
 #10 := (:var 0 int)
 #12 := (uf_1 #10)
 #626 := (pattern #12)
@@ -8983,57 +8793,57 @@
 #87 := [mp #51 #86]: #82
 #146 := [mp~ #87 #130]: #82
 #632 := [mp #146 #631]: #627
-#237 := (not #627)
-#442 := (or #237 #565)
-#578 := (>= 1::int 0::int)
-#419 := (not #578)
-#421 := (= 1::int #420)
-#563 := (or #421 #419)
-#443 := (or #237 #563)
-#550 := (iff #443 #442)
-#547 := (iff #442 #442)
-#548 := [rewrite]: #547
-#559 := (iff #563 #565)
-#554 := (or #565 false)
-#558 := (iff #554 #565)
-#556 := [rewrite]: #558
-#555 := (iff #563 #554)
-#400 := (iff #419 false)
+#609 := (not #627)
+#578 := (or #609 #301)
+#311 := (>= 1::int 0::int)
+#585 := (not #311)
+#586 := (= 1::int #589)
+#590 := (or #586 #585)
+#419 := (or #609 #590)
+#421 := (iff #419 #578)
+#564 := (iff #578 #578)
+#565 := [rewrite]: #564
+#577 := (iff #590 #301)
+#574 := (or #301 false)
+#571 := (iff #574 #301)
+#576 := [rewrite]: #571
+#575 := (iff #590 #574)
+#584 := (iff #585 false)
 #1 := true
-#567 := (not true)
-#569 := (iff #567 false)
-#398 := [rewrite]: #569
-#568 := (iff #419 #567)
-#560 := (iff #578 true)
-#561 := [rewrite]: #560
-#562 := [monotonicity #561]: #568
-#401 := [trans #562 #398]: #400
-#564 := (iff #421 #565)
-#566 := [rewrite]: #564
-#557 := [monotonicity #566 #401]: #555
-#441 := [trans #557 #556]: #559
-#452 := [monotonicity #441]: #550
-#551 := [trans #452 #548]: #550
-#402 := [quant-inst]: #443
-#436 := [mp #402 #551]: #442
-#524 := [unit-resolution #436 #632]: #565
-#526 := (= #293 #420)
+#582 := (not true)
+#583 := (iff #582 false)
+#580 := [rewrite]: #583
+#296 := (iff #585 #582)
+#303 := (iff #311 true)
+#581 := [rewrite]: #303
+#579 := [monotonicity #581]: #296
+#573 := [trans #579 #580]: #584
+#300 := (iff #586 #301)
+#302 := [rewrite]: #300
+#570 := [monotonicity #302 #573]: #575
+#572 := [trans #570 #576]: #577
+#563 := [monotonicity #572]: #421
+#566 := [trans #563 #565]: #421
+#420 := [quant-inst]: #419
+#560 := [mp #420 #566]: #578
+#442 := [unit-resolution #560 #632]: #301
+#443 := (= #293 #589)
 #28 := (= #25 #27)
 #129 := [asserted]: #28
-#527 := [monotonicity #129]: #526
-#528 := [trans #527 #524]: #552
-#529 := (not #552)
-#525 := (or #529 #553)
-#530 := [th-lemma]: #525
-#516 := [unit-resolution #530 #528]: #553
-#517 := (not #553)
-#520 := (or #517 #519)
-#521 := [th-lemma]: #520
-#518 := [unit-resolution #521 #516]: #519
-#502 := (or #522 #577)
-#503 := [th-lemma]: #502
-#505 := [unit-resolution #503 #518]: #522
-#300 := (or #255 #317)
+#436 := [monotonicity #129]: #443
+#451 := [trans #436 #442]: #555
+#453 := (not #555)
+#454 := (or #453 #558)
+#447 := [th-lemma]: #454
+#455 := [unit-resolution #447 #451]: #558
+#456 := (not #558)
+#458 := (or #456 #457)
+#459 := [th-lemma]: #458
+#552 := [unit-resolution #459 #455]: #457
+#553 := (or #549 #588)
+#540 := [th-lemma]: #553
+#542 := [unit-resolution #540 #552]: #549
+#603 := (or #237 #606)
 #18 := (= #13 0::int)
 #118 := (or #18 #70)
 #633 := (forall (vars (?x3 int)) (:pat #626) #118)
@@ -9090,95 +8900,70 @@
 #128 := [mp #88 #127]: #123
 #149 := [mp~ #128 #134]: #123
 #638 := [mp #149 #637]: #633
-#582 := (not #633)
-#296 := (or #582 #255 #317)
+#604 := (not #633)
+#602 := (or #604 #237 #606)
 #204 := (>= #24 0::int)
-#210 := (or #317 #204)
-#579 := (or #582 #210)
-#570 := (iff #579 #296)
-#580 := (or #582 #300)
-#574 := (iff #580 #296)
-#575 := [rewrite]: #574
-#584 := (iff #579 #580)
-#303 := (iff #210 #300)
-#606 := (* 1::int #23)
-#279 := (>= #606 0::int)
-#311 := (or #279 #317)
-#301 := (iff #311 #300)
-#256 := (iff #279 #255)
-#251 := (= #606 #23)
-#593 := [rewrite]: #251
-#257 := [monotonicity #593]: #256
-#302 := [monotonicity #257]: #301
-#586 := (iff #210 #311)
-#587 := (or #317 #279)
-#585 := (iff #587 #311)
-#589 := [rewrite]: #585
-#588 := (iff #210 #587)
-#280 := (iff #204 #279)
-#613 := [rewrite]: #280
-#310 := [monotonicity #613]: #588
-#590 := [trans #310 #589]: #586
-#581 := [trans #590 #302]: #303
-#573 := [monotonicity #581]: #584
-#571 := [trans #573 #575]: #570
-#583 := [quant-inst]: #579
-#576 := [mp #583 #571]: #296
-#506 := [unit-resolution #576 #638]: #300
-#507 := [unit-resolution #506 #505]: #255
-#258 := (not #255)
-#597 := (or #258 #612)
-#601 := (or #237 #258 #612)
+#601 := (or #237 #204)
+#605 := (or #604 #601)
+#317 := (iff #605 #602)
+#592 := (or #604 #603)
+#315 := (iff #592 #602)
+#316 := [rewrite]: #315
+#299 := (iff #605 #592)
+#242 := (iff #601 #603)
+#279 := (iff #204 #606)
+#280 := [rewrite]: #279
+#243 := [monotonicity #280]: #242
+#314 := [monotonicity #243]: #299
+#210 := [trans #314 #316]: #317
+#591 := [quant-inst]: #605
+#587 := [mp #591 #210]: #602
+#534 := [unit-resolution #587 #638]: #603
+#531 := [unit-resolution #534 #542]: #606
+#613 := (not #606)
+#607 := (or #613 #612)
+#251 := (or #609 #613 #612)
 #289 := (not #204)
 #294 := (= #24 #293)
 #291 := (or #294 #289)
-#603 := (or #237 #291)
-#592 := (iff #603 #601)
-#243 := (or #237 #597)
-#605 := (iff #243 #601)
-#591 := [rewrite]: #605
-#604 := (iff #603 #243)
-#594 := (iff #291 #597)
-#614 := (not #279)
-#266 := (or #614 #612)
-#598 := (iff #266 #597)
-#595 := (iff #614 #258)
-#596 := [monotonicity #257]: #595
-#599 := [monotonicity #596]: #598
-#267 := (iff #291 #266)
-#611 := (or #612 #614)
-#271 := (iff #611 #266)
-#608 := [rewrite]: #271
-#617 := (iff #291 #611)
-#615 := (iff #289 #614)
-#616 := [monotonicity #613]: #615
+#593 := (or #609 #291)
+#597 := (iff #593 #251)
+#256 := (or #609 #607)
+#595 := (iff #256 #251)
+#596 := [rewrite]: #595
+#257 := (iff #593 #256)
+#608 := (iff #291 #607)
+#616 := (or #612 #613)
+#266 := (iff #616 #607)
+#271 := [rewrite]: #266
+#611 := (iff #291 #616)
+#614 := (iff #289 #613)
+#615 := [monotonicity #280]: #614
 #268 := (iff #294 #612)
 #399 := [rewrite]: #268
-#607 := [monotonicity #399 #616]: #617
-#609 := [trans #607 #608]: #267
-#600 := [trans #609 #599]: #594
-#602 := [monotonicity #600]: #604
-#299 := [trans #602 #591]: #592
-#242 := [quant-inst]: #603
-#314 := [mp #242 #299]: #601
-#508 := [unit-resolution #314 #632]: #597
-#509 := [unit-resolution #508 #507]: #612
-#510 := (not #612)
-#511 := (or #510 #315)
-#512 := [th-lemma]: #511
-#513 := [unit-resolution #512 #509]: #315
-#316 := (>= #610 0::int)
-#514 := (or #510 #316)
-#504 := [th-lemma]: #514
-#515 := [unit-resolution #504 #509]: #316
-#549 := (<= #293 1::int)
-#493 := (or #529 #549)
-#494 := [th-lemma]: #493
-#496 := [unit-resolution #494 #528]: #549
-[th-lemma #516 #496 #515 #513]: false
-unsat
-
-kDuOn7kAggfP4Um8ghhv5A 6068
+#617 := [monotonicity #399 #615]: #611
+#267 := [trans #617 #271]: #608
+#258 := [monotonicity #267]: #257
+#598 := [trans #258 #596]: #597
+#255 := [quant-inst]: #593
+#599 := [mp #255 #598]: #251
+#533 := [unit-resolution #599 #632]: #607
+#543 := [unit-resolution #533 #531]: #612
+#544 := (not #612)
+#545 := (or #544 #594)
+#541 := [th-lemma]: #545
+#546 := [unit-resolution #541 #543]: #594
+#600 := (>= #610 0::int)
+#535 := (or #544 #600)
+#536 := [th-lemma]: #535
+#537 := [unit-resolution #536 #543]: #600
+#557 := (<= #293 1::int)
+#538 := (or #453 #557)
+#532 := [th-lemma]: #538
+#539 := [unit-resolution #532 #451]: #557
+[th-lemma #455 #539 #537 #546]: false
+unsat
+kyphS4o71h68g2YhvYbQQQ 223 0
 #2 := false
 #23 := 3::int
 decl uf_2 :: (-> T1 int)
@@ -9201,13 +8986,13 @@
 #632 := -1::int
 #634 := (* -1::int #28)
 #290 := (+ #26 #634)
-#609 := (>= #290 0::int)
+#623 := (>= #290 0::int)
 #421 := (= #290 0::int)
-#273 := (>= #22 0::int)
-#610 := (= #28 0::int)
-#588 := (not #610)
-#441 := (<= #28 0::int)
-#443 := (not #441)
+#302 := (>= #22 0::int)
+#625 := (= #28 0::int)
+#318 := (not #625)
+#322 := (<= #28 0::int)
+#324 := (not #322)
 #29 := 7::int
 #143 := (>= #28 7::int)
 #30 := (< #28 7::int)
@@ -9224,12 +9009,12 @@
 #151 := [trans #147 #149]: #150
 #133 := [asserted]: #31
 #152 := [mp #133 #151]: #143
-#585 := (or #443 #141)
-#586 := [th-lemma]: #585
-#587 := [unit-resolution #586 #152]: #443
-#582 := (or #588 #441)
-#583 := [th-lemma]: #582
-#589 := [unit-resolution #583 #587]: #588
+#325 := (or #324 #141)
+#603 := [th-lemma]: #325
+#604 := [unit-resolution #603 #152]: #324
+#601 := (or #318 #322)
+#605 := [th-lemma]: #601
+#602 := [unit-resolution #605 #604]: #318
 #10 := (:var 0 int)
 #12 := (uf_1 #10)
 #648 := (pattern #12)
@@ -9292,45 +9077,33 @@
 #131 := [mp #91 #130]: #126
 #172 := [mp~ #131 #155]: #126
 #660 := [mp #172 #659]: #655
-#605 := (not #655)
-#602 := (or #605 #273 #610)
+#337 := (not #655)
+#338 := (or #337 #302 #625)
 #315 := (>= #26 0::int)
-#332 := (or #610 #315)
-#606 := (or #605 #332)
-#599 := (iff #606 #602)
-#323 := (or #273 #610)
-#596 := (or #605 #323)
-#593 := (iff #596 #602)
-#598 := [rewrite]: #593
-#597 := (iff #606 #596)
-#318 := (iff #332 #323)
-#302 := 1::int
-#635 := (* 1::int #22)
-#636 := (>= #635 0::int)
-#333 := (or #610 #636)
-#603 := (iff #333 #323)
-#608 := (or #610 #273)
-#324 := (iff #608 #323)
-#325 := [rewrite]: #324
-#612 := (iff #333 #608)
-#615 := (iff #636 #273)
-#289 := (= #635 #22)
-#631 := [rewrite]: #289
-#277 := [monotonicity #631]: #615
-#322 := [monotonicity #277]: #612
-#604 := [trans #322 #325]: #603
-#607 := (iff #332 #333)
-#637 := (iff #315 #636)
-#638 := [rewrite]: #637
-#611 := [monotonicity #638]: #607
-#601 := [trans #611 #604]: #318
-#592 := [monotonicity #601]: #597
-#594 := [trans #592 #598]: #599
-#595 := [quant-inst]: #606
-#600 := [mp #595 #594]: #602
-#590 := [unit-resolution #600 #660 #589]: #273
-#278 := (not #273)
-#620 := (or #278 #421)
+#264 := (or #625 #315)
+#339 := (or #337 #264)
+#611 := (iff #339 #338)
+#627 := (or #302 #625)
+#609 := (or #337 #627)
+#333 := (iff #609 #338)
+#607 := [rewrite]: #333
+#610 := (iff #339 #609)
+#321 := (iff #264 #627)
+#265 := (or #625 #302)
+#613 := (iff #265 #627)
+#614 := [rewrite]: #613
+#626 := (iff #264 #265)
+#635 := (iff #315 #302)
+#636 := [rewrite]: #635
+#624 := [monotonicity #636]: #626
+#336 := [trans #624 #614]: #321
+#332 := [monotonicity #336]: #610
+#608 := [trans #332 #607]: #611
+#231 := [quant-inst]: #339
+#612 := [mp #231 #608]: #338
+#606 := [unit-resolution #612 #660 #602]: #302
+#637 := (not #302)
+#293 := (or #637 #421)
 #55 := (= #10 #13)
 #80 := (or #55 #74)
 #649 := (forall (vars (?x2 int)) (:pat #648) #80)
@@ -9380,50 +9153,41 @@
 #90 := [mp #54 #89]: #85
 #169 := [mp~ #90 #134]: #85
 #654 := [mp #169 #653]: #649
-#264 := (not #649)
-#265 := (or #264 #278 #421)
+#615 := (not #649)
+#277 := (or #615 #637 #421)
 #243 := (not #315)
 #317 := (= #26 #28)
 #296 := (or #317 #243)
-#626 := (or #264 #296)
-#337 := (iff #626 #265)
-#627 := (or #264 #620)
-#321 := (iff #627 #265)
-#336 := [rewrite]: #321
-#613 := (iff #626 #627)
-#623 := (iff #296 #620)
-#633 := (not #636)
-#288 := (or #421 #633)
-#622 := (iff #288 #620)
-#617 := (or #421 #278)
-#621 := (iff #617 #620)
-#616 := [rewrite]: #621
-#618 := (iff #288 #617)
-#279 := (iff #633 #278)
-#280 := [monotonicity #277]: #279
-#619 := [monotonicity #280]: #618
-#259 := [trans #619 #616]: #622
-#293 := (iff #296 #288)
-#639 := (iff #243 #633)
-#629 := [monotonicity #638]: #639
+#278 := (or #615 #296)
+#621 := (iff #278 #277)
+#280 := (or #615 #293)
+#619 := (iff #280 #277)
+#620 := [rewrite]: #619
+#617 := (iff #278 #280)
+#631 := (iff #296 #293)
+#639 := (or #421 #637)
+#630 := (iff #639 #293)
+#289 := [rewrite]: #630
+#629 := (iff #296 #639)
+#638 := (iff #243 #637)
+#633 := [monotonicity #636]: #638
 #628 := (iff #317 #421)
 #301 := [rewrite]: #628
-#630 := [monotonicity #301 #629]: #293
-#625 := [trans #630 #259]: #623
-#614 := [monotonicity #625]: #613
-#338 := [trans #614 #336]: #337
-#624 := [quant-inst]: #626
-#339 := [mp #624 #338]: #265
-#584 := [unit-resolution #339 #654]: #620
-#591 := [unit-resolution #584 #590]: #421
-#420 := (not #421)
-#422 := (or #420 #609)
-#423 := [th-lemma]: #422
-#576 := [unit-resolution #423 #591]: #609
-[th-lemma #152 #576 #139]: false
-unsat
-
-aiB004AWADNjynNrqCDsxw 9284
+#288 := [monotonicity #301 #633]: #629
+#273 := [trans #288 #289]: #631
+#618 := [monotonicity #273]: #617
+#616 := [trans #618 #620]: #621
+#279 := [quant-inst]: #278
+#622 := [mp #279 #616]: #277
+#595 := [unit-resolution #622 #654]: #293
+#596 := [unit-resolution #595 #606]: #421
+#597 := (not #421)
+#592 := (or #597 #623)
+#593 := [th-lemma]: #592
+#598 := [unit-resolution #593 #596]: #623
+[th-lemma #152 #598 #139]: false
+unsat
+M8P5WxpiY5AWxaJDBtXoLQ 367 0
 #2 := false
 #9 := 0::int
 decl uf_2 :: (-> T1 int)
@@ -9791,8 +9555,7 @@
 #456 := [th-lemma]: #455
 [unit-resolution #456 #464 #452]: false
 unsat
-
-twoPNF2RBLeff4yYfubpfg 7634
+Xs4JZCKb5egkcPabsrodXg 302 0
 #2 := false
 #9 := 0::int
 decl uf_2 :: (-> T1 int)
@@ -10095,8 +9858,7 @@
 #601 := [unit-resolution #615 #613]: #683
 [th-lemma #623 #188 #601 #628]: false
 unsat
-
-ZcLxnpFewGGQ0H47MfRXGw 12389
+clMAi2WqMi360EjFURRGLg 458 0
 #2 := false
 #9 := 0::int
 decl uf_2 :: (-> T1 int)
@@ -10110,8 +9872,8 @@
 #297 := (uf_2 #141)
 #357 := (= #297 0::int)
 #166 := (uf_1 0::int)
-#454 := (uf_2 #166)
-#515 := (= #454 0::int)
+#531 := (uf_2 #166)
+#537 := (= #531 0::int)
 #10 := (:var 0 int)
 #12 := (uf_1 #10)
 #672 := (pattern #12)
@@ -10168,40 +9930,40 @@
 #192 := [mp~ #95 #175]: #90
 #678 := [mp #192 #677]: #673
 #650 := (not #673)
-#468 := (or #650 #515)
-#528 := (>= 0::int 0::int)
-#508 := (not #528)
-#509 := (= 0::int #454)
-#490 := (or #509 #508)
-#469 := (or #650 #490)
-#471 := (iff #469 #468)
-#473 := (iff #468 #468)
-#474 := [rewrite]: #473
-#462 := (iff #490 #515)
-#495 := (or #515 false)
-#486 := (iff #495 #515)
-#507 := [rewrite]: #486
-#496 := (iff #490 #495)
-#492 := (iff #508 false)
+#528 := (or #650 #537)
+#529 := (>= 0::int 0::int)
+#530 := (not #529)
+#534 := (= 0::int #531)
+#535 := (or #534 #530)
+#508 := (or #650 #535)
+#509 := (iff #508 #528)
+#514 := (iff #528 #528)
+#515 := [rewrite]: #514
+#527 := (iff #535 #537)
+#520 := (or #537 false)
+#525 := (iff #520 #537)
+#526 := [rewrite]: #525
+#521 := (iff #535 #520)
+#519 := (iff #530 false)
 #1 := true
-#491 := (not true)
-#483 := (iff #491 false)
-#485 := [rewrite]: #483
-#450 := (iff #508 #491)
-#516 := (iff #528 true)
-#484 := [rewrite]: #516
-#481 := [monotonicity #484]: #450
-#494 := [trans #481 #485]: #492
-#514 := (iff #509 #515)
-#510 := [rewrite]: #514
-#506 := [monotonicity #510 #494]: #496
-#463 := [trans #506 #507]: #462
-#472 := [monotonicity #463]: #471
-#475 := [trans #472 #474]: #471
-#470 := [quant-inst]: #469
-#476 := [mp #470 #475]: #468
-#351 := [unit-resolution #476 #678]: #515
-#287 := (= #297 #454)
+#512 := (not true)
+#517 := (iff #512 false)
+#518 := [rewrite]: #517
+#513 := (iff #530 #512)
+#538 := (iff #529 true)
+#511 := [rewrite]: #538
+#406 := [monotonicity #511]: #513
+#524 := [trans #406 #518]: #519
+#536 := (iff #534 #537)
+#532 := [rewrite]: #536
+#522 := [monotonicity #532 #524]: #521
+#523 := [trans #522 #526]: #527
+#490 := [monotonicity #523]: #509
+#510 := [trans #490 #515]: #509
+#454 := [quant-inst]: #508
+#516 := [mp #454 #510]: #528
+#394 := [unit-resolution #516 #678]: #537
+#355 := (= #297 #531)
 #250 := (= #141 #166)
 #26 := 2::int
 #144 := (* 2::int #22)
@@ -10212,24 +9974,29 @@
 #161 := (uf_1 #156)
 #336 := (= #161 #166)
 #327 := (not #336)
-#568 := (uf_2 #161)
-#537 := (= #568 0::int)
-#354 := (= #568 #454)
-#352 := [hypothesis]: #336
-#344 := [monotonicity #352]: #354
-#355 := [trans #344 #351]: #537
-#368 := (not #537)
-#527 := (<= #568 0::int)
-#375 := (not #527)
-#566 := (>= #150 0::int)
-#447 := (>= #22 0::int)
-#421 := (= #22 0::int)
+#588 := (uf_2 #161)
+#555 := (= #588 0::int)
+#398 := (= #588 #531)
+#395 := [hypothesis]: #336
+#387 := [monotonicity #395]: #398
+#399 := [trans #387 #394]: #555
+#390 := (not #555)
+#547 := (<= #588 0::int)
+#403 := (not #547)
+#595 := (>= #150 0::int)
+#302 := -1::int
+#618 := (* -1::int #150)
+#624 := (+ #144 #618)
+#488 := (<= #624 0::int)
+#465 := (= #624 0::int)
+#609 := (>= #22 0::int)
+#442 := (= #22 0::int)
 #660 := (uf_1 #22)
-#451 := (uf_2 #660)
-#452 := (= #451 0::int)
-#603 := (not #447)
-#424 := [hypothesis]: #603
-#455 := (or #447 #452)
+#495 := (uf_2 #660)
+#496 := (= #495 0::int)
+#612 := (not #609)
+#451 := [hypothesis]: #612
+#506 := (or #496 #609)
 #18 := (= #13 0::int)
 #126 := (or #18 #78)
 #679 := (forall (vars (?x3 int)) (:pat #672) #126)
@@ -10287,23 +10054,15 @@
 #195 := [mp~ #136 #180]: #131
 #684 := [mp #195 #683]: #679
 #346 := (not #679)
-#458 := (or #346 #447 #452)
-#453 := (or #452 #447)
-#459 := (or #346 #453)
-#434 := (iff #459 #458)
-#443 := (or #346 #455)
-#432 := (iff #443 #458)
-#433 := [rewrite]: #432
-#461 := (iff #459 #443)
-#456 := (iff #453 #455)
-#457 := [rewrite]: #456
-#431 := [monotonicity #457]: #461
-#436 := [trans #431 #433]: #434
-#460 := [quant-inst]: #459
-#437 := [mp #460 #436]: #458
-#420 := [unit-resolution #437 #684]: #455
-#425 := [unit-resolution #420 #424]: #452
-#405 := (= #22 #451)
+#462 := (or #346 #496 #609)
+#463 := (or #346 #506)
+#469 := (iff #463 #462)
+#470 := [rewrite]: #469
+#468 := [quant-inst]: #463
+#471 := [mp #468 #470]: #462
+#452 := [unit-resolution #471 #684]: #506
+#453 := [unit-resolution #452 #451]: #496
+#456 := (= #22 #495)
 #661 := (= uf_3 #660)
 #4 := (:var 0 T1)
 #5 := (uf_2 #4)
@@ -10334,136 +10093,117 @@
 #663 := (not #665)
 #653 := (or #663 #661)
 #312 := [quant-inst]: #653
-#415 := [unit-resolution #312 #671]: #661
-#407 := [monotonicity #415]: #405
-#408 := [trans #407 #425]: #421
-#411 := (not #421)
-#412 := (or #411 #447)
-#416 := [th-lemma]: #412
-#409 := [unit-resolution #416 #424 #408]: false
-#417 := [lemma #409]: #447
-#302 := -1::int
-#618 := (* -1::int #150)
-#624 := (+ #144 #618)
-#595 := (<= #624 0::int)
-#465 := (= #624 0::int)
-#489 := (or #603 #465)
-#482 := (or #650 #603 #465)
+#455 := [unit-resolution #312 #671]: #661
+#457 := [monotonicity #455]: #456
+#458 := [trans #457 #453]: #442
+#459 := (not #442)
+#460 := (or #459 #609)
+#443 := [th-lemma]: #460
+#461 := [unit-resolution #443 #451 #458]: false
+#431 := [lemma #461]: #609
+#613 := (or #465 #612)
+#615 := (or #650 #465 #612)
 #616 := (>= #144 0::int)
 #617 := (not #616)
 #622 := (= #144 #150)
 #623 := (or #622 #617)
-#497 := (or #650 #623)
-#504 := (iff #497 #482)
-#500 := (or #650 #489)
-#502 := (iff #500 #482)
-#503 := [rewrite]: #502
-#493 := (iff #497 #500)
-#594 := (iff #623 #489)
-#609 := (* 1::int #22)
-#610 := (>= #609 0::int)
-#606 := (not #610)
-#614 := (or #465 #606)
-#498 := (iff #614 #489)
-#605 := (or #465 #603)
-#448 := (iff #605 #489)
-#596 := [rewrite]: #448
-#487 := (iff #614 #605)
-#604 := (iff #606 #603)
-#600 := (iff #610 #447)
-#444 := (= #609 #22)
-#446 := [rewrite]: #444
-#601 := [monotonicity #446]: #600
-#602 := [monotonicity #601]: #604
-#488 := [monotonicity #602]: #487
-#593 := [trans #488 #596]: #498
-#608 := (iff #623 #614)
-#607 := (iff #617 #606)
-#611 := (iff #616 #610)
-#612 := [rewrite]: #611
-#613 := [monotonicity #612]: #607
+#444 := (or #650 #623)
+#602 := (iff #444 #615)
+#447 := (or #650 #613)
+#603 := (iff #447 #615)
+#604 := [rewrite]: #603
+#600 := (iff #444 #447)
+#614 := (iff #623 #613)
+#606 := (iff #617 #612)
+#610 := (iff #616 #609)
+#611 := [rewrite]: #610
+#607 := [monotonicity #611]: #606
 #466 := (iff #622 #465)
 #467 := [rewrite]: #466
-#615 := [monotonicity #467 #613]: #608
-#597 := [trans #615 #593]: #594
-#501 := [monotonicity #597]: #493
-#505 := [trans #501 #503]: #504
-#499 := [quant-inst]: #497
-#598 := [mp #499 #505]: #482
-#404 := [unit-resolution #598 #678]: #489
-#386 := [unit-resolution #404 #417]: #465
-#388 := (not #465)
-#389 := (or #388 #595)
-#390 := [th-lemma]: #389
-#391 := [unit-resolution #390 #386]: #595
-#395 := (not #595)
-#413 := (or #566 #603 #395)
-#403 := [th-lemma]: #413
-#373 := [unit-resolution #403 #391 #417]: #566
-#553 := -3::int
-#551 := (* -1::int #568)
-#552 := (+ #150 #551)
-#535 := (<= #552 -3::int)
-#554 := (= #552 -3::int)
-#557 := (>= #150 -3::int)
+#608 := [monotonicity #467 #607]: #614
+#601 := [monotonicity #608]: #600
+#605 := [trans #601 #604]: #602
+#446 := [quant-inst]: #444
+#487 := [mp #446 #605]: #615
+#439 := [unit-resolution #487 #678]: #613
+#435 := [unit-resolution #439 #431]: #465
+#440 := (not #465)
+#419 := (or #440 #488)
+#422 := [th-lemma]: #419
+#426 := [unit-resolution #422 #435]: #488
+#430 := (not #488)
+#433 := (or #595 #612 #430)
+#438 := [th-lemma]: #433
+#402 := [unit-resolution #438 #431 #426]: #595
+#590 := -3::int
+#579 := (* -1::int #588)
+#589 := (+ #150 #579)
+#553 := (<= #589 -3::int)
+#591 := (= #589 -3::int)
+#581 := (>= #150 -3::int)
 #644 := (>= #22 -1::int)
-#392 := (or #603 #644)
-#393 := [th-lemma]: #392
-#394 := [unit-resolution #393 #417]: #644
+#428 := (or #612 #644)
+#429 := [th-lemma]: #428
+#427 := [unit-resolution #429 #431]: #644
 #646 := (not #644)
-#396 := (or #557 #646 #395)
-#397 := [th-lemma]: #396
-#398 := [unit-resolution #397 #391 #394]: #557
-#560 := (not #557)
-#539 := (or #554 #560)
-#543 := (or #650 #554 #560)
-#567 := (>= #156 0::int)
-#564 := (not #567)
-#548 := (= #156 #568)
-#549 := (or #548 #564)
-#544 := (or #650 #549)
-#530 := (iff #544 #543)
-#546 := (or #650 #539)
-#533 := (iff #546 #543)
-#529 := [rewrite]: #533
-#541 := (iff #544 #546)
-#540 := (iff #549 #539)
-#550 := (iff #564 #560)
-#558 := (iff #567 #557)
-#559 := [rewrite]: #558
-#561 := [monotonicity #559]: #550
-#555 := (iff #548 #554)
-#556 := [rewrite]: #555
-#542 := [monotonicity #556 #561]: #540
-#547 := [monotonicity #542]: #541
-#531 := [trans #547 #529]: #530
-#545 := [quant-inst]: #544
-#534 := [mp #545 #531]: #543
-#387 := [unit-resolution #534 #678]: #539
-#399 := [unit-resolution #387 #398]: #554
-#376 := (not #554)
-#378 := (or #376 #535)
-#379 := [th-lemma]: #378
-#380 := [unit-resolution #379 #399]: #535
-#365 := (not #535)
-#364 := (not #566)
-#366 := (or #375 #364 #365)
-#358 := [th-lemma]: #366
-#367 := [unit-resolution #358 #380 #373]: #375
-#359 := (or #368 #527)
-#369 := [th-lemma]: #359
-#350 := [unit-resolution #369 #367]: #368
-#321 := [unit-resolution #350 #355]: false
-#323 := [lemma #321]: #327
+#418 := (or #581 #646 #430)
+#421 := [th-lemma]: #418
+#423 := [unit-resolution #421 #426 #427]: #581
+#584 := (not #581)
+#573 := (or #584 #591)
+#562 := (or #650 #584 #591)
+#599 := (>= #156 0::int)
+#586 := (not #599)
+#580 := (= #156 #588)
+#577 := (or #580 #586)
+#563 := (or #650 #577)
+#549 := (iff #563 #562)
+#566 := (or #650 #573)
+#568 := (iff #566 #562)
+#548 := [rewrite]: #568
+#567 := (iff #563 #566)
+#571 := (iff #577 #573)
+#569 := (or #591 #584)
+#574 := (iff #569 #573)
+#575 := [rewrite]: #574
+#570 := (iff #577 #569)
+#578 := (iff #586 #584)
+#582 := (iff #599 #581)
+#583 := [rewrite]: #582
+#585 := [monotonicity #583]: #578
+#587 := (iff #580 #591)
+#592 := [rewrite]: #587
+#572 := [monotonicity #592 #585]: #570
+#576 := [trans #572 #575]: #571
+#564 := [monotonicity #576]: #567
+#551 := [trans #564 #548]: #549
+#565 := [quant-inst]: #563
+#552 := [mp #565 #551]: #562
+#424 := [unit-resolution #552 #678]: #573
+#420 := [unit-resolution #424 #423]: #591
+#425 := (not #591)
+#415 := (or #425 #553)
+#405 := [th-lemma]: #415
+#407 := [unit-resolution #405 #420]: #553
+#404 := (not #553)
+#401 := (not #595)
+#386 := (or #403 #401 #404)
+#388 := [th-lemma]: #386
+#389 := [unit-resolution #388 #407 #402]: #403
+#391 := (or #390 #547)
+#392 := [th-lemma]: #391
+#393 := [unit-resolution #392 #389]: #390
+#376 := [unit-resolution #393 #399]: false
+#378 := [lemma #376]: #327
 #249 := (= #141 #161)
 #334 := (not #249)
-#343 := (= #297 #568)
-#322 := [hypothesis]: #249
-#333 := [monotonicity #322]: #343
-#315 := (not #343)
-#414 := (+ #297 #551)
-#401 := (>= #414 0::int)
-#372 := (not #401)
+#396 := (= #297 #588)
+#385 := [hypothesis]: #249
+#370 := [monotonicity #385]: #396
+#380 := (not #396)
+#434 := (+ #297 #579)
+#280 := (>= #434 0::int)
+#414 := (not #280)
 #303 := (* -1::int #297)
 #304 := (+ #22 #303)
 #356 := (>= #304 -1::int)
@@ -10492,21 +10232,21 @@
 #256 := [trans #360 #362]: #363
 #637 := [quant-inst]: #651
 #633 := [mp #637 #256]: #648
-#381 := [unit-resolution #633 #678]: #649
-#382 := [unit-resolution #381 #394]: #641
-#383 := (not #641)
-#384 := (or #383 #356)
-#377 := [th-lemma]: #384
-#385 := [unit-resolution #377 #382]: #356
-#370 := [hypothesis]: #401
-#371 := [th-lemma #398 #370 #385 #380 #391]: false
-#374 := [lemma #371]: #372
-#328 := (or #315 #401)
-#329 := [th-lemma]: #328
-#332 := [unit-resolution #329 #374]: #315
-#316 := [unit-resolution #332 #333]: false
-#318 := [lemma #316]: #334
-#295 := (or #249 #250 #336)
+#408 := [unit-resolution #633 #678]: #649
+#411 := [unit-resolution #408 #427]: #641
+#412 := (not #641)
+#416 := (or #412 #356)
+#409 := [th-lemma]: #416
+#417 := [unit-resolution #409 #411]: #356
+#410 := [hypothesis]: #280
+#413 := [th-lemma #423 #410 #417 #407 #426]: false
+#400 := [lemma #413]: #414
+#381 := (or #380 #280)
+#382 := [th-lemma]: #381
+#377 := [unit-resolution #382 #400]: #380
+#371 := [unit-resolution #377 #370]: false
+#372 := [lemma #371]: #334
+#352 := (or #249 #250 #336)
 #335 := (not #250)
 #338 := (and #334 #335 #327)
 #339 := (not #338)
@@ -10554,31 +10294,30 @@
 #177 := [mp #137 #174]: #172
 #326 := (or #169 #339)
 #659 := [def-axiom]: #326
-#294 := [unit-resolution #659 #177]: #339
+#351 := [unit-resolution #659 #177]: #339
 #314 := (or #338 #249 #250 #336)
 #445 := [def-axiom]: #314
-#293 := [unit-resolution #445 #294]: #295
-#296 := [unit-resolution #293 #318 #323]: #250
-#290 := [monotonicity #296]: #287
-#285 := [trans #290 #351]: #357
-#310 := (not #357)
+#343 := [unit-resolution #445 #351]: #352
+#353 := [unit-resolution #343 #372 #378]: #250
+#321 := [monotonicity #353]: #355
+#323 := [trans #321 #394]: #357
+#368 := (not #357)
 #620 := (<= #297 0::int)
-#305 := (not #620)
+#364 := (not #620)
 #634 := (<= #304 -1::int)
-#319 := (or #383 #634)
-#298 := [th-lemma]: #319
-#300 := [unit-resolution #298 #382]: #634
-#306 := (not #634)
-#307 := (or #305 #603 #306)
-#308 := [th-lemma]: #307
-#309 := [unit-resolution #308 #300 #417]: #305
-#299 := (or #310 #620)
-#311 := [th-lemma]: #299
-#292 := [unit-resolution #311 #309]: #310
-[unit-resolution #292 #285]: false
-unsat
-
-ipe8HUL/t33OoeNl/z0smQ 4011
+#374 := (or #412 #634)
+#373 := [th-lemma]: #374
+#375 := [unit-resolution #373 #411]: #634
+#365 := (not #634)
+#366 := (or #364 #612 #365)
+#358 := [th-lemma]: #366
+#367 := [unit-resolution #358 #375 #431]: #364
+#359 := (or #368 #620)
+#369 := [th-lemma]: #359
+#350 := [unit-resolution #369 #367]: #368
+[unit-resolution #350 #323]: false
+unsat
+mu7O1os0t3tPqWZhwizjxw 161 0
 #2 := false
 #9 := 0::int
 decl uf_3 :: int
@@ -10740,8 +10479,7 @@
 #361 := [unit-resolution #639 #655]: #647
 [th-lemma #656 #361 #261]: false
 unsat
-
-9FrZeHP8ZKJM+JmhfAjimQ 14889
+08cmOtIT4NYs2PG/F3zeZw 557 0
 #2 := false
 #9 := 0::int
 decl uf_2 :: (-> T1 int)
@@ -10753,46 +10491,46 @@
 #38 := (* 4::int #37)
 #39 := (uf_1 #38)
 #40 := (uf_2 #39)
-#504 := (= #40 0::int)
-#995 := (not #504)
-#475 := (<= #40 0::int)
-#990 := (not #475)
+#527 := (= #40 0::int)
+#976 := (not #527)
+#502 := (<= #40 0::int)
+#971 := (not #502)
 #22 := 1::int
 #186 := (+ 1::int #40)
 #189 := (uf_1 #186)
-#467 := (uf_2 #189)
-#380 := (<= #467 1::int)
-#893 := (not #380)
+#506 := (uf_2 #189)
+#407 := (<= #506 1::int)
+#876 := (not #407)
 decl up_4 :: (-> T1 T1 bool)
 #4 := (:var 0 T1)
-#386 := (up_4 #4 #189)
-#374 := (pattern #386)
-#382 := (not #386)
-#385 := (= #4 #189)
+#408 := (up_4 #4 #189)
+#393 := (pattern #408)
+#413 := (= #4 #189)
+#414 := (not #408)
 #26 := (uf_1 1::int)
 #27 := (= #4 #26)
-#845 := (or #27 #385 #382)
-#848 := (forall (vars (?x5 T1)) (:pat #374) #845)
-#851 := (not #848)
-#854 := (or #380 #851)
-#857 := (not #854)
+#392 := (or #27 #414 #413)
+#397 := (forall (vars (?x5 T1)) (:pat #393) #392)
+#383 := (not #397)
+#382 := (or #383 #407)
+#375 := (not #382)
 decl up_3 :: (-> T1 bool)
 #192 := (up_3 #189)
-#840 := (not #192)
-#860 := (or #840 #857)
+#404 := (not #192)
+#841 := (or #404 #375)
 decl ?x5!0 :: (-> T1 T1)
-#392 := (?x5!0 #189)
-#397 := (up_4 #392 #189)
-#390 := (not #397)
-#396 := (= #26 #392)
-#395 := (= #189 #392)
-#866 := (or #395 #396 #390)
-#869 := (not #866)
-#872 := (or #192 #380 #869)
-#875 := (not #872)
-#863 := (not #860)
-#878 := (or #863 #875)
-#881 := (not #878)
+#422 := (?x5!0 #189)
+#434 := (= #189 #422)
+#417 := (up_4 #422 #189)
+#418 := (not #417)
+#415 := (= #26 #422)
+#847 := (or #415 #418 #434)
+#850 := (not #847)
+#853 := (or #192 #407 #850)
+#856 := (not #853)
+#844 := (not #841)
+#859 := (or #844 #856)
+#862 := (not #859)
 #5 := (uf_2 #4)
 #787 := (pattern #5)
 #21 := (up_3 #4)
@@ -10971,59 +10709,64 @@
 #314 := [mp #267 #313]: #311
 #839 := [mp #314 #838]: #836
 #589 := (not #836)
-#884 := (or #589 #881)
-#398 := (or #390 #396 #395)
-#383 := (not #398)
-#381 := (or #192 #380 #383)
-#384 := (not #381)
-#387 := (or #382 #27 #385)
-#376 := (forall (vars (?x5 T1)) (:pat #374) #387)
-#377 := (not #376)
-#375 := (or #380 #377)
-#378 := (not #375)
-#841 := (or #840 #378)
-#842 := (not #841)
-#843 := (or #842 #384)
-#844 := (not #843)
-#885 := (or #589 #844)
-#887 := (iff #885 #884)
-#889 := (iff #884 #884)
-#890 := [rewrite]: #889
-#882 := (iff #844 #881)
-#879 := (iff #843 #878)
-#876 := (iff #384 #875)
-#873 := (iff #381 #872)
-#870 := (iff #383 #869)
-#867 := (iff #398 #866)
-#868 := [rewrite]: #867
-#871 := [monotonicity #868]: #870
-#874 := [monotonicity #871]: #873
-#877 := [monotonicity #874]: #876
-#864 := (iff #842 #863)
-#861 := (iff #841 #860)
-#858 := (iff #378 #857)
-#855 := (iff #375 #854)
-#852 := (iff #377 #851)
-#849 := (iff #376 #848)
-#846 := (iff #387 #845)
-#847 := [rewrite]: #846
-#850 := [quant-intro #847]: #849
-#853 := [monotonicity #850]: #852
-#856 := [monotonicity #853]: #855
-#859 := [monotonicity #856]: #858
-#862 := [monotonicity #859]: #861
-#865 := [monotonicity #862]: #864
-#880 := [monotonicity #865 #877]: #879
-#883 := [monotonicity #880]: #882
-#888 := [monotonicity #883]: #887
-#891 := [trans #888 #890]: #887
-#886 := [quant-inst]: #885
-#892 := [mp #886 #891]: #884
-#966 := [unit-resolution #892 #839]: #881
-#924 := (or #878 #860)
-#925 := [def-axiom]: #924
-#967 := [unit-resolution #925 #966]: #860
-#970 := (or #863 #857)
+#865 := (or #589 #862)
+#416 := (or #418 #415 #434)
+#419 := (not #416)
+#409 := (or #192 #407 #419)
+#410 := (not #409)
+#389 := (or #414 #27 #413)
+#394 := (forall (vars (?x5 T1)) (:pat #393) #389)
+#399 := (not #394)
+#401 := (or #407 #399)
+#402 := (not #401)
+#400 := (or #404 #402)
+#405 := (not #400)
+#388 := (or #405 #410)
+#391 := (not #388)
+#866 := (or #589 #391)
+#868 := (iff #866 #865)
+#870 := (iff #865 #865)
+#871 := [rewrite]: #870
+#863 := (iff #391 #862)
+#860 := (iff #388 #859)
+#857 := (iff #410 #856)
+#854 := (iff #409 #853)
+#851 := (iff #419 #850)
+#848 := (iff #416 #847)
+#849 := [rewrite]: #848
+#852 := [monotonicity #849]: #851
+#855 := [monotonicity #852]: #854
+#858 := [monotonicity #855]: #857
+#845 := (iff #405 #844)
+#842 := (iff #400 #841)
+#378 := (iff #402 #375)
+#376 := (iff #401 #382)
+#384 := (or #407 #383)
+#387 := (iff #384 #382)
+#374 := [rewrite]: #387
+#385 := (iff #401 #384)
+#380 := (iff #399 #383)
+#390 := (iff #394 #397)
+#395 := (iff #389 #392)
+#396 := [rewrite]: #395
+#398 := [quant-intro #396]: #390
+#381 := [monotonicity #398]: #380
+#386 := [monotonicity #381]: #385
+#377 := [trans #386 #374]: #376
+#840 := [monotonicity #377]: #378
+#843 := [monotonicity #840]: #842
+#846 := [monotonicity #843]: #845
+#861 := [monotonicity #846 #858]: #860
+#864 := [monotonicity #861]: #863
+#869 := [monotonicity #864]: #868
+#872 := [trans #869 #871]: #868
+#867 := [quant-inst]: #866
+#873 := [mp #867 #872]: #865
+#947 := [unit-resolution #873 #839]: #862
+#905 := (or #859 #841)
+#906 := [def-axiom]: #905
+#948 := [unit-resolution #906 #947]: #841
+#951 := (or #844 #375)
 #41 := (+ #40 1::int)
 #42 := (uf_1 #41)
 #43 := (up_3 #42)
@@ -11035,30 +10778,30 @@
 #194 := [monotonicity #191]: #193
 #185 := [asserted]: #43
 #197 := [mp #185 #194]: #192
-#904 := (or #863 #840 #857)
-#905 := [def-axiom]: #904
-#971 := [unit-resolution #905 #197]: #970
-#972 := [unit-resolution #971 #967]: #857
-#894 := (or #854 #893)
-#895 := [def-axiom]: #894
-#973 := [unit-resolution #895 #972]: #893
+#885 := (or #844 #404 #375)
+#886 := [def-axiom]: #885
+#952 := [unit-resolution #886 #197]: #951
+#953 := [unit-resolution #952 #948]: #375
+#877 := (or #382 #876)
+#878 := [def-axiom]: #877
+#954 := [unit-resolution #878 #953]: #876
 #542 := -1::int
-#446 := (* -1::int #467)
-#447 := (+ #40 #446)
-#416 := (>= #447 -1::int)
-#438 := (= #447 -1::int)
-#453 := (>= #40 -1::int)
-#419 := (= #467 0::int)
-#978 := (not #419)
-#388 := (<= #467 0::int)
-#974 := (not #388)
-#975 := (or #974 #380)
-#976 := [th-lemma]: #975
-#977 := [unit-resolution #976 #973]: #974
-#979 := (or #978 #388)
-#980 := [th-lemma]: #979
-#981 := [unit-resolution #980 #977]: #978
-#409 := (or #419 #453)
+#508 := (* -1::int #506)
+#493 := (+ #40 #508)
+#438 := (>= #493 -1::int)
+#494 := (= #493 -1::int)
+#496 := (>= #40 -1::int)
+#451 := (= #506 0::int)
+#959 := (not #451)
+#432 := (<= #506 0::int)
+#955 := (not #432)
+#956 := (or #955 #407)
+#957 := [th-lemma]: #956
+#958 := [unit-resolution #957 #954]: #955
+#960 := (or #959 #432)
+#961 := [th-lemma]: #960
+#962 := [unit-resolution #961 #958]: #959
+#453 := (or #451 #496)
 #10 := (:var 0 int)
 #12 := (uf_1 #10)
 #795 := (pattern #12)
@@ -11121,28 +10864,28 @@
 #145 := [mp #105 #144]: #140
 #227 := [mp~ #145 #208]: #140
 #807 := [mp #227 #806]: #802
-#496 := (not #802)
-#408 := (or #496 #419 #453)
-#476 := (>= #186 0::int)
-#407 := (or #419 #476)
-#414 := (or #496 #407)
-#404 := (iff #414 #408)
-#393 := (or #496 #409)
-#401 := (iff #393 #408)
-#402 := [rewrite]: #401
-#394 := (iff #414 #393)
-#410 := (iff #407 #409)
-#454 := (iff #476 #453)
-#455 := [rewrite]: #454
-#413 := [monotonicity #455]: #410
-#399 := [monotonicity #413]: #394
-#400 := [trans #399 #402]: #404
-#389 := [quant-inst]: #414
-#405 := [mp #389 #400]: #408
-#982 := [unit-resolution #405 #807]: #409
-#983 := [unit-resolution #982 #981]: #453
-#445 := (not #453)
-#441 := (or #438 #445)
+#514 := (not #802)
+#445 := (or #514 #451 #496)
+#504 := (>= #186 0::int)
+#452 := (or #451 #504)
+#456 := (or #514 #452)
+#429 := (iff #456 #445)
+#441 := (or #514 #453)
+#423 := (iff #441 #445)
+#428 := [rewrite]: #423
+#442 := (iff #456 #441)
+#454 := (iff #452 #453)
+#498 := (iff #504 #496)
+#487 := [rewrite]: #498
+#455 := [monotonicity #487]: #454
+#421 := [monotonicity #455]: #442
+#430 := [trans #421 #428]: #429
+#439 := [quant-inst]: #456
+#431 := [mp #439 #430]: #445
+#963 := [unit-resolution #431 #807]: #453
+#964 := [unit-resolution #963 #962]: #496
+#488 := (not #496)
+#490 := (or #494 #488)
 #69 := (= #10 #13)
 #94 := (or #69 #88)
 #796 := (forall (vars (?x2 int)) (:pat #795) #94)
@@ -11192,48 +10935,48 @@
 #104 := [mp #68 #103]: #99
 #224 := [mp~ #104 #196]: #99
 #801 := [mp #224 #800]: #796
-#514 := (not #796)
-#423 := (or #514 #438 #445)
-#477 := (not #476)
-#478 := (= #186 #467)
-#444 := (or #478 #477)
-#428 := (or #514 #444)
-#434 := (iff #428 #423)
-#430 := (or #514 #441)
-#433 := (iff #430 #423)
-#422 := [rewrite]: #433
-#431 := (iff #428 #430)
-#442 := (iff #444 #441)
-#456 := (iff #477 #445)
-#439 := [monotonicity #455]: #456
-#451 := (iff #478 #438)
-#452 := [rewrite]: #451
-#421 := [monotonicity #452 #439]: #442
-#432 := [monotonicity #421]: #431
-#415 := [trans #432 #422]: #434
-#429 := [quant-inst]: #428
-#417 := [mp #429 #415]: #423
-#984 := [unit-resolution #417 #801]: #441
-#985 := [unit-resolution #984 #983]: #438
-#986 := (not #438)
-#987 := (or #986 #416)
-#988 := [th-lemma]: #987
-#989 := [unit-resolution #988 #985]: #416
-#991 := (not #416)
-#992 := (or #990 #380 #991)
-#993 := [th-lemma]: #992
-#994 := [unit-resolution #993 #989 #973]: #990
-#996 := (or #995 #475)
-#997 := [th-lemma]: #996
-#998 := [unit-resolution #997 #994]: #995
-#536 := (>= #37 0::int)
-#525 := (not #536)
+#530 := (not #796)
+#492 := (or #530 #494 #488)
+#505 := (not #504)
+#507 := (= #186 #506)
+#500 := (or #507 #505)
+#473 := (or #530 #500)
+#478 := (iff #473 #492)
+#475 := (or #530 #490)
+#477 := (iff #475 #492)
+#467 := [rewrite]: #477
+#466 := (iff #473 #475)
+#491 := (iff #500 #490)
+#489 := (iff #505 #488)
+#481 := [monotonicity #487]: #489
+#495 := (iff #507 #494)
+#497 := [rewrite]: #495
+#482 := [monotonicity #497 #481]: #491
+#476 := [monotonicity #482]: #466
+#444 := [trans #476 #467]: #478
+#474 := [quant-inst]: #473
+#446 := [mp #474 #444]: #492
+#965 := [unit-resolution #446 #801]: #490
+#966 := [unit-resolution #965 #964]: #494
+#967 := (not #494)
+#968 := (or #967 #438)
+#969 := [th-lemma]: #968
+#970 := [unit-resolution #969 #966]: #438
+#972 := (not #438)
+#973 := (or #971 #407 #972)
+#974 := [th-lemma]: #973
+#975 := [unit-resolution #974 #970 #954]: #971
+#977 := (or #976 #502)
+#978 := [th-lemma]: #977
+#979 := [unit-resolution #978 #975]: #976
+#553 := (>= #37 0::int)
+#546 := (not #553)
 #545 := (* -1::int #40)
 #549 := (+ #38 #545)
 #551 := (= #549 0::int)
-#1003 := (not #551)
-#503 := (>= #549 0::int)
-#999 := (not #503)
+#984 := (not #551)
+#524 := (>= #549 0::int)
+#980 := (not #524)
 #201 := (>= #37 1::int)
 #202 := (not #201)
 #44 := (<= 1::int #37)
@@ -11244,157 +10987,107 @@
 #204 := [monotonicity #200]: #203
 #195 := [asserted]: #45
 #205 := [mp #195 #204]: #202
-#1000 := (or #999 #201 #380 #991)
-#1001 := [th-lemma]: #1000
-#1002 := [unit-resolution #1001 #205 #989 #973]: #999
-#1004 := (or #1003 #503)
-#1005 := [th-lemma]: #1004
-#1006 := [unit-resolution #1005 #1002]: #1003
-#527 := (or #525 #551)
-#515 := (or #514 #525 #551)
+#981 := (or #980 #201 #407 #972)
+#982 := [th-lemma]: #981
+#983 := [unit-resolution #982 #205 #970 #954]: #980
+#985 := (or #984 #524)
+#986 := [th-lemma]: #985
+#987 := [unit-resolution #986 #983]: #984
+#548 := (or #551 #546)
+#531 := (or #530 #551 #546)
 #403 := (>= #38 0::int)
 #562 := (not #403)
 #558 := (= #38 #40)
 #563 := (or #558 #562)
-#516 := (or #514 #563)
+#534 := (or #530 #563)
+#537 := (iff #534 #531)
+#539 := (or #530 #548)
+#533 := (iff #539 #531)
+#536 := [rewrite]: #533
+#532 := (iff #534 #539)
+#538 := (iff #563 #548)
+#547 := (iff #562 #546)
+#541 := (iff #403 #553)
+#544 := [rewrite]: #541
+#543 := [monotonicity #544]: #547
+#552 := (iff #558 #551)
+#550 := [rewrite]: #552
+#528 := [monotonicity #550 #543]: #538
+#540 := [monotonicity #528]: #532
+#523 := [trans #540 #536]: #537
+#535 := [quant-inst]: #534
+#525 := [mp #535 #523]: #531
+#988 := [unit-resolution #525 #801]: #548
+#989 := [unit-resolution #988 #987]: #546
+#511 := (or #527 #553)
+#515 := (or #514 #527 #553)
+#509 := (or #527 #403)
+#516 := (or #514 #509)
 #522 := (iff #516 #515)
-#518 := (or #514 #527)
+#518 := (or #514 #511)
 #521 := (iff #518 #515)
 #510 := [rewrite]: #521
 #519 := (iff #516 #518)
-#512 := (iff #563 #527)
-#553 := (* 1::int #37)
-#541 := (>= #553 0::int)
-#547 := (not #541)
-#531 := (or #547 #551)
-#509 := (iff #531 #527)
-#526 := (iff #547 #525)
-#537 := (iff #541 #536)
-#540 := (= #553 #37)
-#533 := [rewrite]: #540
-#523 := [monotonicity #533]: #537
-#524 := [monotonicity #523]: #526
-#511 := [monotonicity #524]: #509
-#539 := (iff #563 #531)
-#538 := (or #551 #547)
-#534 := (iff #538 #531)
-#535 := [rewrite]: #534
-#528 := (iff #563 #538)
-#543 := (iff #562 #547)
-#544 := (iff #403 #541)
-#546 := [rewrite]: #544
-#548 := [monotonicity #546]: #543
-#552 := (iff #558 #551)
-#550 := [rewrite]: #552
-#530 := [monotonicity #550 #548]: #528
-#532 := [trans #530 #535]: #539
-#513 := [trans #532 #511]: #512
+#512 := (iff #509 #511)
+#513 := [monotonicity #544]: #512
 #520 := [monotonicity #513]: #519
 #499 := [trans #520 #510]: #522
 #517 := [quant-inst]: #516
 #501 := [mp #517 #499]: #515
-#1007 := [unit-resolution #501 #801]: #527
-#1008 := [unit-resolution #1007 #1006]: #525
-#508 := (or #504 #536)
-#498 := (or #496 #504 #536)
-#505 := (or #504 #403)
-#487 := (or #496 #505)
-#492 := (iff #487 #498)
-#489 := (or #496 #508)
-#491 := (iff #489 #498)
-#482 := [rewrite]: #491
-#481 := (iff #487 #489)
-#495 := (iff #505 #508)
-#506 := (or #504 #541)
-#493 := (iff #506 #508)
-#494 := [monotonicity #523]: #493
-#507 := (iff #505 #506)
-#500 := [monotonicity #546]: #507
-#497 := [trans #500 #494]: #495
-#490 := [monotonicity #497]: #481
-#473 := [trans #490 #482]: #492
-#488 := [quant-inst]: #487
-#474 := [mp #488 #473]: #498
-#1009 := [unit-resolution #474 #807]: #508
-[unit-resolution #1009 #1008 #998]: false
-unsat
-
-uq2MbDTgTG1nxWdwzZtWew 7
-unsat
-
-E5BydeDaPocMMvChMGY+og 7
-unsat
-
-p81EQzqwJrGunGO7GuNt3g 7
-unsat
-
-KpYfvnTcz2WncWNg3dJDCA 7
-unsat
-
-ybGRm230DLO0wH6aROKBBw 7
-unsat
-
-goFtZfJ6kkxA8sqBVpZutw 7
-unsat
-
-0+nmgsMqioeTuwam1ScP7g 7
-unsat
-
-nI63LP/VCL//bjsS1gNB2A 7
-unsat
-
-9+2QHvrRgbKz97Zg0kfybw 7
-unsat
-
-6kquszLXeBUhTwzaw6gd2Q 7
-unsat
-
-j5Z04lpza+N5I1cpno5mtw 7
-unsat
-
-mapbfUM6Ils30x5nEBJmaw 7
-unsat
-
-e8P++0FczY3zhNhEVclACw 7
-unsat
-
-yXMQNOyCylhI+EH8hNYxHA 7
-unsat
-
-GkYN9j7cjrR2KR/lb04/qw 7
-unsat
-
-PajzgNjLWHwVHpjoje+gnA 7
-unsat
-
-URpJYU7D8PO0VRnciRgE5A 7
-unsat
-
-D9ZGhymoV3L6zbWsJlwG2A 7
-unsat
-
-0QLuovrnnANWnCkUY3l10g 7
-unsat
-
-CYprps2G0Au5F3Z7n3KTRg 7
-unsat
-
-iyIMuJd6zijfEao8zKnx2w 7
-unsat
-
-49jzsAwAEfR6NSFBhBEisQ 7
-unsat
-
-T0j6xFgrghxif91jL+2yAw 7
-unsat
-
-md/M3rVve0+8sQ6oqIoL2w 7
-unsat
-
-pY7C8PCf5lVVaim6q7PJcQ 7
-unsat
-
-4zCFLQf4Jrov/gmEvsBm4Q 1036
+#990 := [unit-resolution #501 #807]: #511
+[unit-resolution #990 #989 #979]: false
+unsat
+8HdmSMHHP2B8XMFzjNuw5Q 1 0
+unsat
+O4aM0+/isn2q5CrIefZjzg 1 0
+unsat
+t/ni9djl2DqxH0iKupZSwg 1 0
+unsat
+RumBGekdxZQaBF1HNa3x9w 1 0
+unsat
+Q9d+IbQ8chjKld71X6/zqw 1 0
+unsat
+PhC8zQV8hnJ6E2YYjZPGjQ 1 0
+unsat
+mieI2RhSp3bYaojlWH1A4A 1 0
+unsat
+pRSV6nBLconzrQz2zUrJ6g 1 0
+unsat
+Js0JfdwDoKq3YuilPPgeZw 1 0
+unsat
+GRIqjLUJiqXbo+pXhAeKIw 1 0
+unsat
+Bg5scsmPFp82+7Y2ScL6Wg 1 0
+unsat
+XD6zX6850dLxyfZSfNv30A 1 0
+unsat
+BG/HwJYnumvDICXxtBu/tA 1 0
+unsat
+YMc4t19sUMWbUkx3woxCmQ 1 0
+unsat
+YyD9IF72pKXGGKZTO7FY5Q 1 0
+unsat
+zRPsIUi+TEoz5fPWP0H9bQ 1 0
+unsat
+8ipTE8BOXpvSo/U6D4p3lA 1 0
+unsat
+MSzQywedZPsOE0CDxrrO0g 1 0
+unsat
+SryZuXv48ItET8NPIv07pA 1 0
+unsat
+qOMUQN18hYFl/wWt54lvbA 1 0
+unsat
++njWXdn6fETK3/AjtiHjcA 1 0
+unsat
+5cQ7gJ33gzYTIIPA3hbBmQ 1 0
+unsat
+ZznT34cvumrP00mXZ3gcjw 1 0
+unsat
+//LQca1Et5RfhQJZA+CGCA 1 0
+unsat
+3ntxKz+kaQNfTrLzY9sVXw 1 0
+unsat
+4lL2Qo8ngE1EH1UdeN1Qng 43 0
 #2 := false
 #6 := 0::int
 decl uf_1 :: (-> bv[2] int)
@@ -11438,11 +11131,9 @@
 #287 := [th-lemma]: #627
 [unit-resolution #287 #47 #635]: false
 unsat
-
-czvSLyjMowmFNi82us4N2Q 7
-unsat
-
-aU+7kcyE8oAPbs5RjfuwIw 1160
++xe3O927LrflFUE6NDqRlA 1 0
+unsat
+JPoL7fPYhqhAkjUiVF+THQ 50 0
 #2 := false
 decl uf_6 :: T2
 #23 := uf_6
@@ -11493,8 +11184,7 @@
 #66 := [asserted]: #26
 [unit-resolution #66 #235]: false
 unsat
-
-dXfueqZAXkudfE6G0VKkwg 2559
+l23ZDmd0VbO/Q+uO5EtabA 105 0
 #2 := false
 decl uf_6 :: (-> T4 T2)
 decl uf_10 :: T4
@@ -11600,8 +11290,7 @@
 #110 := [asserted]: #46
 [unit-resolution #110 #238]: false
 unsat
-
-Dc/6bNJffjYYplvoitScJQ 4578
+GZjffeZPQnL3OyLCvxdCpg 181 0
 #2 := false
 decl uf_1 :: (-> T1 T2 T3)
 decl uf_3 :: T2
@@ -11783,8 +11472,7 @@
 #76 := [asserted]: #38
 [unit-resolution #76 #489]: false
 unsat
-
-jdmsd1j41Osn+WzTtqTUSQ 1352
+i6jCzzRosHYE0w7sF1Nraw 62 0
 #2 := false
 decl up_4 :: (-> T1 T2 bool)
 decl uf_3 :: T2
@@ -11847,8 +11535,7 @@
 #73 := [unit-resolution #71 #68]: #72
 [unit-resolution #73 #59 #61]: false
 unsat
-
-EA8ecQ7czWL46/C3k7D7tg 2697
+YZHSyhN2TGlpe+vpkzWrgQ 115 0
 #2 := false
 decl up_2 :: (-> T2 bool)
 decl uf_3 :: T2
@@ -11964,8 +11651,7 @@
 #560 := [mp #344 #559]: #557
 [unit-resolution #560 #576 #561]: false
 unsat
-
-mNfbN3NQCB4ik2xJmLE1UQ 11936
+TibRlXkU+X+1+zGPYTiT0g 464 0
 #2 := false
 decl uf_2 :: (-> T2 T3 T3)
 decl uf_4 :: T3
@@ -12430,8 +12116,7 @@
 #177 := [asserted]: #53
 [unit-resolution #177 #793]: false
 unsat
-
-Jtmz+P173L9nRQkQk52h+Q 420
+DJPKxi9AO25zGBcs5kxUrw 21 0
 #2 := false
 decl up_1 :: (-> T1 bool)
 #4 := (:var 0 T1)
@@ -12453,8 +12138,7 @@
 #25 := [asserted]: #9
 [mp #25 #34]: false
 unsat
-
-YG20f6Uf93koN6rVg/alpA 9742
+i5PnMbuM9mWv5LnVszz9+g 366 0
 #2 := false
 decl uf_1 :: (-> int T1)
 #37 := 6::int
@@ -12469,18 +12153,18 @@
 #35 := (uf_1 #34)
 #36 := (uf_3 #35)
 #39 := (= #36 #38)
-#484 := (uf_3 #38)
-#372 := (= #484 #38)
-#485 := (= #38 #484)
-#592 := (uf_2 #38)
+#476 := (uf_3 #38)
+#403 := (= #476 #38)
+#531 := (= #38 #476)
+#620 := (uf_2 #38)
 #142 := -10::int
-#496 := (+ -10::int #592)
-#497 := (uf_1 #496)
-#498 := (uf_3 #497)
-#499 := (= #484 #498)
+#513 := (+ -10::int #620)
+#472 := (uf_1 #513)
+#503 := (uf_3 #472)
+#505 := (= #476 #503)
 #22 := 10::int
-#500 := (>= #592 10::int)
-#501 := (ite #500 #499 #485)
+#507 := (>= #620 10::int)
+#514 := (ite #507 #505 #531)
 #4 := (:var 0 T1)
 #21 := (uf_3 #4)
 #707 := (pattern #21)
@@ -12554,12 +12238,12 @@
 #212 := [mp #207 #211]: #193
 #713 := [mp #212 #712]: #708
 #336 := (not #708)
-#463 := (or #336 #501)
-#464 := [quant-inst]: #463
-#444 := [unit-resolution #464 #713]: #501
-#473 := (not #500)
-#453 := (<= #592 6::int)
-#597 := (= #592 6::int)
+#518 := (or #336 #514)
+#528 := [quant-inst]: #518
+#477 := [unit-resolution #528 #713]: #514
+#529 := (not #507)
+#498 := (<= #620 6::int)
+#610 := (= #620 6::int)
 #10 := (:var 0 int)
 #12 := (uf_1 #10)
 #694 := (pattern #12)
@@ -12617,79 +12301,79 @@
 #201 := [mp~ #99 #183]: #94
 #700 := [mp #201 #699]: #695
 #673 := (not #695)
-#576 := (or #673 #597)
-#607 := (>= 6::int 0::int)
-#591 := (not #607)
-#594 := (= 6::int #592)
-#595 := (or #594 #591)
-#577 := (or #673 #595)
-#579 := (iff #577 #576)
-#581 := (iff #576 #576)
-#582 := [rewrite]: #581
-#574 := (iff #595 #597)
-#586 := (or #597 false)
-#571 := (iff #586 #597)
-#573 := [rewrite]: #571
-#590 := (iff #595 #586)
-#588 := (iff #591 false)
+#591 := (or #673 #610)
+#526 := (>= 6::int 0::int)
+#527 := (not #526)
+#617 := (= 6::int #620)
+#621 := (or #617 #527)
+#592 := (or #673 #621)
+#595 := (iff #592 #591)
+#597 := (iff #591 #591)
+#593 := [rewrite]: #597
+#600 := (iff #621 #610)
+#614 := (or #610 false)
+#605 := (iff #614 #610)
+#606 := [rewrite]: #605
+#603 := (iff #621 #614)
+#613 := (iff #527 false)
 #1 := true
 #663 := (not true)
 #666 := (iff #663 false)
 #667 := [rewrite]: #666
-#585 := (iff #591 #663)
-#598 := (iff #607 true)
-#584 := [rewrite]: #598
-#587 := [monotonicity #584]: #585
-#589 := [trans #587 #667]: #588
-#596 := (iff #594 #597)
-#593 := [rewrite]: #596
-#570 := [monotonicity #593 #589]: #590
-#575 := [trans #570 #573]: #574
-#580 := [monotonicity #575]: #579
-#572 := [trans #580 #582]: #579
-#578 := [quant-inst]: #577
-#583 := [mp #578 #572]: #576
-#448 := [unit-resolution #583 #700]: #597
-#445 := (not #597)
-#446 := (or #445 #453)
-#442 := [th-lemma]: #446
-#447 := [unit-resolution #442 #448]: #453
-#437 := (not #453)
-#427 := (or #437 #473)
-#429 := [th-lemma]: #427
-#430 := [unit-resolution #429 #447]: #473
-#471 := (not #501)
-#477 := (or #471 #500 #485)
-#478 := [def-axiom]: #477
-#433 := [unit-resolution #478 #430 #444]: #485
-#373 := [symm #433]: #372
-#374 := (= #36 #484)
+#611 := (iff #527 #663)
+#599 := (iff #526 true)
+#601 := [rewrite]: #599
+#612 := [monotonicity #601]: #611
+#609 := [trans #612 #667]: #613
+#608 := (iff #617 #610)
+#602 := [rewrite]: #608
+#604 := [monotonicity #602 #609]: #603
+#607 := [trans #604 #606]: #600
+#596 := [monotonicity #607]: #595
+#598 := [trans #596 #593]: #595
+#594 := [quant-inst]: #592
+#584 := [mp #594 #598]: #591
+#478 := [unit-resolution #584 #700]: #610
+#453 := (not #610)
+#454 := (or #453 #498)
+#455 := [th-lemma]: #454
+#456 := [unit-resolution #455 #478]: #498
+#458 := (not #498)
+#459 := (or #458 #529)
+#460 := [th-lemma]: #459
+#302 := [unit-resolution #460 #456]: #529
+#508 := (not #514)
+#490 := (or #508 #507 #531)
+#491 := [def-axiom]: #490
+#461 := [unit-resolution #491 #302 #477]: #531
+#404 := [symm #461]: #403
+#405 := (= #36 #476)
 #649 := (uf_2 #35)
-#554 := (+ -10::int #649)
-#549 := (uf_1 #554)
-#545 := (uf_3 #549)
-#381 := (= #545 #484)
-#395 := (= #549 #38)
-#394 := (= #554 6::int)
+#582 := (+ -10::int #649)
+#553 := (uf_1 #582)
+#556 := (uf_3 #553)
+#401 := (= #556 #476)
+#417 := (= #553 #38)
+#415 := (= #582 6::int)
 #335 := (uf_2 #31)
 #647 := -1::int
-#459 := (* -1::int #335)
-#460 := (+ #33 #459)
-#302 := (<= #460 0::int)
-#458 := (= #33 #335)
-#426 := (= #32 #31)
-#555 := (= #31 #32)
-#551 := (+ -10::int #335)
-#552 := (uf_1 #551)
-#553 := (uf_3 #552)
-#556 := (= #32 #553)
-#557 := (>= #335 10::int)
-#558 := (ite #557 #556 #555)
-#560 := (or #336 #558)
-#533 := [quant-inst]: #560
-#434 := [unit-resolution #533 #713]: #558
-#535 := (not #557)
-#531 := (<= #335 4::int)
+#502 := (* -1::int #335)
+#463 := (+ #33 #502)
+#464 := (<= #463 0::int)
+#486 := (= #33 #335)
+#445 := (= #32 #31)
+#574 := (= #31 #32)
+#575 := (+ -10::int #335)
+#576 := (uf_1 #575)
+#577 := (uf_3 #576)
+#578 := (= #32 #577)
+#579 := (>= #335 10::int)
+#580 := (ite #579 #578 #574)
+#572 := (or #336 #580)
+#583 := [quant-inst]: #572
+#457 := [unit-resolution #583 #713]: #580
+#562 := (not #579)
+#554 := (<= #335 4::int)
 #324 := (= #335 4::int)
 #659 := (or #673 #324)
 #678 := (>= 4::int 0::int)
@@ -12719,125 +12403,109 @@
 #277 := [trans #383 #385]: #382
 #367 := [quant-inst]: #660
 #655 := [mp #367 #277]: #659
-#438 := [unit-resolution #655 #700]: #324
-#431 := (not #324)
-#439 := (or #431 #531)
-#432 := [th-lemma]: #439
-#435 := [unit-resolution #432 #438]: #531
-#436 := (not #531)
-#422 := (or #436 #535)
-#424 := [th-lemma]: #422
-#425 := [unit-resolution #424 #435]: #535
-#534 := (not #558)
-#540 := (or #534 #557 #555)
-#541 := [def-axiom]: #540
-#423 := [unit-resolution #541 #425 #434]: #555
-#408 := [symm #423]: #426
-#410 := [monotonicity #408]: #458
-#411 := (not #458)
-#412 := (or #411 #302)
-#413 := [th-lemma]: #412
-#414 := [unit-resolution #413 #410]: #302
-#461 := (>= #460 0::int)
-#415 := (or #411 #461)
-#416 := [th-lemma]: #415
-#417 := [unit-resolution #416 #410]: #461
-#512 := (>= #335 4::int)
-#418 := (or #431 #512)
-#419 := [th-lemma]: #418
-#420 := [unit-resolution #419 #438]: #512
+#462 := [unit-resolution #655 #700]: #324
+#441 := (not #324)
+#444 := (or #441 #554)
+#448 := [th-lemma]: #444
+#450 := [unit-resolution #448 #462]: #554
+#451 := (not #554)
+#449 := (or #451 #562)
+#452 := [th-lemma]: #449
+#440 := [unit-resolution #452 #450]: #562
+#561 := (not #580)
+#566 := (or #561 #579 #574)
+#567 := [def-axiom]: #566
+#443 := [unit-resolution #567 #440 #457]: #574
+#446 := [symm #443]: #445
+#442 := [monotonicity #446]: #486
+#447 := (not #486)
+#437 := (or #447 #464)
+#427 := [th-lemma]: #437
+#429 := [unit-resolution #427 #442]: #464
+#471 := (>= #463 0::int)
+#430 := (or #447 #471)
+#433 := [th-lemma]: #430
+#434 := [unit-resolution #433 #442]: #471
+#560 := (>= #335 4::int)
+#438 := (or #441 #560)
+#431 := [th-lemma]: #438
+#439 := [unit-resolution #431 #462]: #560
 #651 := (* -1::int #649)
 #648 := (+ #34 #651)
-#521 := (<= #648 0::int)
+#625 := (<= #648 0::int)
 #652 := (= #648 0::int)
-#630 := (>= #33 0::int)
-#421 := (not #461)
-#409 := (not #512)
-#398 := (or #630 #409 #421)
-#400 := [th-lemma]: #398
-#401 := [unit-resolution #400 #420 #417]: #630
-#468 := (not #630)
-#623 := (or #468 #652)
-#509 := (or #673 #468 #652)
+#643 := (>= #33 0::int)
+#435 := (not #471)
+#432 := (not #560)
+#436 := (or #643 #432 #435)
+#422 := [th-lemma]: #436
+#424 := [unit-resolution #422 #439 #434]: #643
+#644 := (not #643)
+#489 := (or #644 #652)
+#628 := (or #673 #644 #652)
 #370 := (>= #34 0::int)
 #371 := (not #370)
 #650 := (= #34 #649)
 #364 := (or #650 #371)
-#510 := (or #673 #364)
-#619 := (iff #510 #509)
-#470 := (or #673 #623)
-#615 := (iff #470 #509)
-#616 := [rewrite]: #615
-#618 := (iff #510 #470)
-#624 := (iff #364 #623)
-#643 := 1::int
-#638 := (* 1::int #33)
-#639 := (>= #638 0::int)
-#640 := (not #639)
-#632 := (or #640 #652)
-#625 := (iff #632 #623)
-#469 := (iff #640 #468)
-#637 := (iff #639 #630)
-#635 := (= #638 #33)
-#636 := [rewrite]: #635
-#466 := [monotonicity #636]: #637
-#622 := [monotonicity #466]: #469
-#626 := [monotonicity #622]: #625
-#628 := (iff #364 #632)
-#488 := (or #652 #640)
-#633 := (iff #488 #632)
-#634 := [rewrite]: #633
-#489 := (iff #364 #488)
-#646 := (iff #371 #640)
-#644 := (iff #370 #639)
-#645 := [rewrite]: #644
-#487 := [monotonicity #645]: #646
+#629 := (or #673 #364)
+#469 := (iff #629 #628)
+#636 := (or #673 #489)
+#466 := (iff #636 #628)
+#468 := [rewrite]: #466
+#630 := (iff #629 #636)
+#633 := (iff #364 #489)
+#646 := (or #652 #644)
+#631 := (iff #646 #489)
+#632 := [rewrite]: #631
+#487 := (iff #364 #646)
+#645 := (iff #371 #644)
+#638 := (iff #370 #643)
+#639 := [rewrite]: #638
+#640 := [monotonicity #639]: #645
 #641 := (iff #650 #652)
 #642 := [rewrite]: #641
-#631 := [monotonicity #642 #487]: #489
-#629 := [trans #631 #634]: #628
-#627 := [trans #629 #626]: #624
-#520 := [monotonicity #627]: #618
-#504 := [trans #520 #616]: #619
-#511 := [quant-inst]: #510
-#519 := [mp #511 #504]: #509
-#402 := [unit-resolution #519 #700]: #623
-#403 := [unit-resolution #402 #401]: #652
-#404 := (not #652)
-#405 := (or #404 #521)
-#406 := [th-lemma]: #405
-#399 := [unit-resolution #406 #403]: #521
-#522 := (>= #648 0::int)
-#407 := (or #404 #522)
-#392 := [th-lemma]: #407
-#393 := [unit-resolution #392 #403]: #522
-#396 := [th-lemma #393 #399 #420 #435 #417 #414]: #394
-#397 := [monotonicity #396]: #395
-#391 := [monotonicity #397]: #381
-#550 := (= #36 #545)
-#559 := (= #35 #36)
-#530 := (>= #649 10::int)
-#476 := (ite #530 #550 #559)
-#536 := (or #336 #476)
-#537 := [quant-inst]: #536
-#386 := [unit-resolution #537 #713]: #476
-#387 := (not #521)
-#388 := (or #530 #387 #409 #421)
-#380 := [th-lemma]: #388
-#389 := [unit-resolution #380 #399 #420 #417]: #530
-#538 := (not #530)
-#532 := (not #476)
-#506 := (or #532 #538 #550)
-#513 := [def-axiom]: #506
-#390 := [unit-resolution #513 #389 #386]: #550
-#365 := [trans #390 #391]: #374
-#375 := [trans #365 #373]: #39
+#488 := [monotonicity #642 #640]: #487
+#634 := [trans #488 #632]: #633
+#637 := [monotonicity #634]: #630
+#622 := [trans #637 #468]: #469
+#635 := [quant-inst]: #629
+#623 := [mp #635 #622]: #628
+#425 := [unit-resolution #623 #700]: #489
+#423 := [unit-resolution #425 #424]: #652
+#426 := (not #652)
+#408 := (or #426 #625)
+#410 := [th-lemma]: #408
+#411 := [unit-resolution #410 #423]: #625
+#626 := (>= #648 0::int)
+#412 := (or #426 #626)
+#413 := [th-lemma]: #412
+#414 := [unit-resolution #413 #423]: #626
+#416 := [th-lemma #414 #411 #439 #450 #434 #429]: #415
+#418 := [monotonicity #416]: #417
+#402 := [monotonicity #418]: #401
+#557 := (= #36 #556)
+#581 := (= #35 #36)
+#558 := (>= #649 10::int)
+#559 := (ite #558 #557 #581)
+#533 := (or #336 #559)
+#534 := [quant-inst]: #533
+#419 := [unit-resolution #534 #713]: #559
+#420 := (not #625)
+#409 := (or #558 #420 #432 #435)
+#421 := [th-lemma]: #409
+#398 := [unit-resolution #421 #411 #439 #434]: #558
+#428 := (not #558)
+#535 := (not #559)
+#539 := (or #535 #428 #557)
+#540 := [def-axiom]: #539
+#400 := [unit-resolution #540 #398 #419]: #557
+#406 := [trans #400 #402]: #405
+#399 := [trans #406 #404]: #39
 #40 := (not #39)
 #182 := [asserted]: #40
-[unit-resolution #182 #375]: false
-unsat
-
-/fwo5o8vhLVHyS4oYEs4QA 10902
+[unit-resolution #182 #399]: false
+unsat
+K2SXMHU6QCZJ8TRs6zjKRg 408 0
 #2 := false
 #22 := 0::int
 #8 := 2::int
@@ -12913,18 +12581,18 @@
 #604 := [trans #593 #597]: #562
 #603 := [quant-inst]: #596
 #606 := [mp #603 #604]: #628
-#516 := [unit-resolution #606 #817]: #566
-#498 := (not #566)
-#432 := (or #498 #608)
-#411 := [th-lemma]: #432
-#413 := [unit-resolution #411 #516]: #608
+#528 := [unit-resolution #606 #817]: #566
+#521 := (not #566)
+#464 := (or #521 #608)
+#456 := [th-lemma]: #464
+#465 := [unit-resolution #456 #528]: #608
 decl uf_10 :: int
 #52 := uf_10
 #57 := (mod uf_10 2::int)
 #243 := (* -1::int #57)
 #244 := (+ #56 #243)
 #447 := (>= #244 0::int)
-#372 := (not #447)
+#387 := (not #447)
 #245 := (= #244 0::int)
 #248 := (not #245)
 #218 := (* -1::int #55)
@@ -12944,9 +12612,9 @@
 #662 := (not #672)
 #1 := true
 #81 := [true-axiom]: true
-#514 := (or false #662)
-#515 := [th-lemma]: #514
-#513 := [unit-resolution #515 #81]: #662
+#520 := (or false #662)
+#523 := [th-lemma]: #520
+#524 := [unit-resolution #523 #81]: #662
 #701 := (* -1::int #613)
 #204 := -2::int
 #691 := (* -2::int #222)
@@ -12959,58 +12627,82 @@
 #546 := [th-lemma]: #545
 #548 := [unit-resolution #546 #81]: #692
 #549 := (not #692)
-#482 := (or #549 #694)
-#483 := [th-lemma]: #482
-#484 := [unit-resolution #483 #548]: #694
+#497 := (or #549 #694)
+#482 := [th-lemma]: #497
+#483 := [unit-resolution #482 #548]: #694
 #536 := (not #448)
-#382 := [hypothesis]: #536
+#395 := [hypothesis]: #536
 #555 := (* -1::int uf_10)
 #573 := (+ #51 #555)
 #543 := (<= #573 0::int)
 #53 := (= #51 uf_10)
 #256 := (not #253)
 #259 := (or #248 #256)
+#502 := 1::int
 #731 := (div uf_10 2::int)
-#723 := (* -2::int #731)
-#521 := (* -2::int #624)
-#529 := (+ #521 #723)
+#515 := (* -1::int #731)
+#513 := (+ #640 #515)
 #618 := (div #51 2::int)
-#583 := (* -2::int #618)
-#522 := (+ #583 #529)
-#528 := (* -2::int #57)
-#525 := (+ #528 #522)
-#524 := (* 2::int #56)
-#526 := (+ #524 #525)
-#523 := (* 2::int uf_10)
-#512 := (+ #523 #526)
-#520 := (>= #512 2::int)
+#514 := (* -1::int #618)
+#516 := (+ #514 #513)
+#498 := (+ #243 #516)
+#500 := (+ #56 #498)
+#501 := (+ uf_10 #500)
+#503 := (>= #501 1::int)
+#486 := (not #503)
+#361 := (<= #244 0::int)
 #453 := (not #259)
-#519 := [hypothesis]: #453
+#517 := [hypothesis]: #453
 #440 := (or #259 #245)
 #451 := [def-axiom]: #440
-#470 := [unit-resolution #451 #519]: #245
-#465 := (or #248 #447)
-#466 := [th-lemma]: #465
-#457 := [unit-resolution #466 #470]: #447
-#544 := (>= #573 0::int)
-#441 := (not #544)
+#519 := [unit-resolution #451 #517]: #245
+#478 := (or #248 #361)
+#470 := [th-lemma]: #478
+#479 := [unit-resolution #470 #519]: #361
+#449 := (>= #252 0::int)
 #452 := (or #259 #253)
 #380 := [def-axiom]: #452
-#481 := [unit-resolution #380 #519]: #253
-#467 := (or #256 #448)
-#434 := [th-lemma]: #467
-#436 := [unit-resolution #434 #481]: #448
-#532 := (or #543 #536)
-#695 := (>= #699 0::int)
-#550 := (or #549 #695)
-#393 := [th-lemma]: #550
-#551 := [unit-resolution #393 #548]: #695
-#547 := (not #543)
-#552 := [hypothesis]: #547
+#480 := [unit-resolution #380 #517]: #253
+#471 := (or #256 #449)
+#481 := [th-lemma]: #471
+#462 := [unit-resolution #481 #480]: #449
+#487 := (not #361)
+#485 := (not #449)
+#476 := (or #486 #485 #487)
+#607 := (<= #620 0::int)
+#529 := (or #521 #607)
+#522 := [th-lemma]: #529
+#525 := [unit-resolution #522 #528]: #607
+#723 := (* -2::int #731)
+#724 := (+ #243 #723)
+#718 := (+ uf_10 #724)
+#720 := (<= #718 0::int)
+#722 := (= #718 0::int)
+#526 := (or false #722)
+#512 := [th-lemma]: #526
+#504 := [unit-resolution #512 #81]: #722
+#505 := (not #722)
+#506 := (or #505 #720)
+#507 := [th-lemma]: #506
+#508 := [unit-resolution #507 #504]: #720
+#509 := [hypothesis]: #361
+#583 := (* -2::int #618)
+#584 := (+ #583 #640)
+#585 := (+ #51 #584)
+#587 := (<= #585 0::int)
+#582 := (= #585 0::int)
+#510 := (or false #582)
+#499 := [th-lemma]: #510
+#511 := [unit-resolution #499 #81]: #582
+#488 := (not #582)
+#490 := (or #488 #587)
+#491 := [th-lemma]: #490
+#492 := [unit-resolution #491 #511]: #587
+#493 := [hypothesis]: #503
 #649 := (* -2::int #60)
 #644 := (+ #218 #649)
 #650 := (+ #51 #644)
-#631 := (<= #650 0::int)
+#636 := (>= #650 0::int)
 #623 := (= #650 0::int)
 #43 := (uf_7 uf_2 #35)
 #44 := (uf_6 #34 #43)
@@ -13071,6 +12763,33 @@
 #630 := [quant-inst]: #629
 #531 := [unit-resolution #630 #824]: #623
 #534 := (not #623)
+#494 := (or #534 #636)
+#495 := [th-lemma]: #494
+#496 := [unit-resolution #495 #531]: #636
+#489 := [hypothesis]: #449
+#484 := [th-lemma #483 #489 #496 #493 #492 #509 #508 #525 #524]: false
+#477 := [lemma #484]: #476
+#463 := [unit-resolution #477 #462 #479]: #486
+#727 := (>= #718 0::int)
+#466 := (or #505 #727)
+#457 := [th-lemma]: #466
+#467 := [unit-resolution #457 #504]: #727
+#434 := (or #248 #447)
+#436 := [th-lemma]: #434
+#437 := [unit-resolution #436 #519]: #447
+#544 := (>= #573 0::int)
+#445 := (not #544)
+#428 := (or #256 #448)
+#441 := [th-lemma]: #428
+#442 := [unit-resolution #441 #480]: #448
+#532 := (or #543 #536)
+#695 := (>= #699 0::int)
+#550 := (or #549 #695)
+#393 := [th-lemma]: #550
+#551 := [unit-resolution #393 #548]: #695
+#547 := (not #543)
+#552 := [hypothesis]: #547
+#631 := (<= #650 0::int)
 #538 := (or #534 #631)
 #540 := [th-lemma]: #538
 #541 := [unit-resolution #540 #531]: #631
@@ -13081,8 +12800,8 @@
 #533 := [unit-resolution #530 #81]: #666
 #535 := [th-lemma #533 #539 #541 #552 #551]: false
 #537 := [lemma #535]: #532
-#437 := [unit-resolution #537 #436]: #543
-#444 := (or #547 #441)
+#443 := [unit-resolution #537 #442]: #543
+#429 := (or #547 #445)
 #764 := (not #53)
 #771 := (or #764 #259)
 #262 := (iff #53 #259)
@@ -13135,121 +12854,67 @@
 #438 := (or #764 #259 #433)
 #439 := [def-axiom]: #438
 #772 := [unit-resolution #439 #267]: #771
-#428 := [unit-resolution #772 #519]: #764
-#442 := (or #53 #547 #441)
-#443 := [th-lemma]: #442
-#445 := [unit-resolution #443 #428]: #444
-#435 := [unit-resolution #445 #437]: #441
-#584 := (+ #583 #640)
-#585 := (+ #51 #584)
+#444 := [unit-resolution #772 #517]: #764
+#435 := (or #53 #547 #445)
+#446 := [th-lemma]: #435
+#431 := [unit-resolution #446 #444]: #429
+#432 := [unit-resolution #431 #443]: #445
 #588 := (>= #585 0::int)
-#582 := (= #585 0::int)
-#499 := (or false #582)
-#511 := [th-lemma]: #499
-#488 := [unit-resolution #511 #81]: #582
-#490 := (not #582)
-#446 := (or #490 #588)
-#429 := [th-lemma]: #446
-#431 := [unit-resolution #429 #488]: #588
-#724 := (+ #243 #723)
-#718 := (+ uf_10 #724)
-#727 := (>= #718 0::int)
-#722 := (= #718 0::int)
-#503 := (or false #722)
-#504 := [th-lemma]: #503
-#505 := [unit-resolution #504 #81]: #722
-#506 := (not #722)
-#418 := (or #506 #727)
-#419 := [th-lemma]: #418
-#420 := [unit-resolution #419 #505]: #727
-#421 := [th-lemma #420 #413 #431 #435 #457]: #520
-#485 := (not #520)
-#361 := (<= #244 0::int)
-#479 := (or #248 #361)
-#480 := [th-lemma]: #479
-#471 := [unit-resolution #480 #470]: #361
-#449 := (>= #252 0::int)
-#462 := (or #256 #449)
-#463 := [th-lemma]: #462
-#464 := [unit-resolution #463 #481]: #449
-#476 := (not #361)
-#487 := (not #449)
-#477 := (or #485 #487 #476)
-#607 := (<= #620 0::int)
-#500 := (or #498 #607)
-#501 := [th-lemma]: #500
-#502 := [unit-resolution #501 #516]: #607
-#720 := (<= #718 0::int)
-#507 := (or #506 #720)
-#508 := [th-lemma]: #507
-#509 := [unit-resolution #508 #505]: #720
-#510 := [hypothesis]: #361
-#587 := (<= #585 0::int)
-#491 := (or #490 #587)
-#492 := [th-lemma]: #491
-#493 := [unit-resolution #492 #488]: #587
-#494 := [hypothesis]: #520
-#636 := (>= #650 0::int)
-#495 := (or #534 #636)
-#496 := [th-lemma]: #495
-#489 := [unit-resolution #496 #531]: #636
-#497 := [hypothesis]: #449
-#486 := [th-lemma #484 #497 #489 #494 #493 #510 #509 #502 #513]: false
-#478 := [lemma #486]: #477
-#456 := [unit-resolution #478 #464 #471]: #485
-#422 := [unit-resolution #456 #421]: false
-#423 := [lemma #422]: #259
+#411 := (or #488 #588)
+#413 := [th-lemma]: #411
+#418 := [unit-resolution #413 #511]: #588
+#419 := [th-lemma #418 #432 #437 #467 #465 #463]: false
+#420 := [lemma #419]: #259
 #427 := (or #53 #453)
 #768 := (or #53 #453 #433)
 #770 := [def-axiom]: #768
 #557 := [unit-resolution #770 #267]: #427
-#399 := [unit-resolution #557 #423]: #53
-#385 := (or #764 #543)
-#386 := [th-lemma]: #385
-#387 := [unit-resolution #386 #399]: #543
-#379 := [th-lemma #489 #387 #382 #484 #513]: false
-#388 := [lemma #379]: #448
-#381 := (or #253 #536)
-#397 := [hypothesis]: #487
-#400 := (or #764 #544)
-#403 := [th-lemma]: #400
-#398 := [unit-resolution #403 #399]: #544
-#404 := [th-lemma #398 #397 #533 #551 #541]: false
-#378 := [lemma #404]: #449
-#395 := (or #253 #536 #487)
-#377 := [th-lemma]: #395
-#658 := [unit-resolution #377 #378]: #381
-#653 := [unit-resolution #658 #388]: #253
+#406 := [unit-resolution #557 #420]: #53
+#377 := (or #764 #543)
+#381 := [th-lemma]: #377
+#382 := [unit-resolution #381 #406]: #543
+#385 := [th-lemma #496 #382 #395 #483 #524]: false
+#386 := [lemma #385]: #448
+#390 := (or #253 #536)
+#408 := [hypothesis]: #485
+#409 := (or #764 #544)
+#397 := [th-lemma]: #409
+#399 := [unit-resolution #397 #406]: #544
+#400 := [th-lemma #399 #408 #533 #551 #541]: false
+#403 := [lemma #400]: #449
+#392 := (or #253 #536 #485)
+#394 := [th-lemma]: #392
+#657 := [unit-resolution #394 #403]: #390
+#658 := [unit-resolution #657 #386]: #253
 #450 := (or #453 #248 #256)
 #454 := [def-axiom]: #450
-#664 := [unit-resolution #454 #423]: #259
-#665 := [unit-resolution #664 #653]: #248
-#373 := (or #245 #372)
-#708 := (+ #57 #640)
-#705 := (>= #708 0::int)
-#560 := (= #57 #624)
-#408 := (= #624 #57)
-#406 := [monotonicity #399]: #408
-#409 := [symm #406]: #560
-#706 := (not #560)
-#655 := (or #706 #705)
-#569 := [th-lemma]: #655
-#570 := [unit-resolution #569 #409]: #705
-#383 := [hypothesis]: #476
-#384 := [th-lemma #502 #383 #570]: false
-#389 := [lemma #384]: #361
-#369 := (or #245 #476 #372)
-#370 := [th-lemma]: #369
-#374 := [unit-resolution #370 #389]: #373
-#375 := [unit-resolution #374 #665]: #372
-#610 := (<= #708 0::int)
-#371 := (or #706 #610)
-#376 := [th-lemma]: #371
-#363 := [unit-resolution #376 #409]: #610
-[th-lemma #363 #375 #413]: false
-unsat
-
-s8LL71+1HTN+eIuEYeKhUw 1251
+#762 := [unit-resolution #454 #420]: #259
+#664 := [unit-resolution #762 #658]: #248
+#372 := (or #245 #387)
+#560 := (+ #57 #640)
+#610 := (>= #560 0::int)
+#742 := (= #57 #624)
+#424 := (= #624 #57)
+#405 := [monotonicity #406]: #424
+#407 := [symm #405]: #742
+#705 := (not #742)
+#706 := (or #705 #610)
+#568 := [th-lemma]: #706
+#569 := [unit-resolution #568 #407]: #610
+#398 := [hypothesis]: #487
+#404 := [th-lemma #525 #398 #569]: false
+#378 := [lemma #404]: #361
+#379 := (or #245 #487 #387)
+#388 := [th-lemma]: #379
+#369 := [unit-resolution #388 #378]: #372
+#370 := [unit-resolution #369 #664]: #387
+#708 := (<= #560 0::int)
+#373 := (or #705 #708)
+#374 := [th-lemma]: #373
+#375 := [unit-resolution #374 #407]: #708
+[th-lemma #375 #370 #465]: false
+unsat
+1DhSL9G2fGRGmuI8IaMNOA 50 0
 #2 := false
 decl up_35 :: (-> int bool)
 #112 := 1::int
@@ -13300,8 +12965,7 @@
 #504 := [quant-inst]: #503
 [unit-resolution #504 #916 #297]: false
 unsat
-
-MZYsU5krlrOK4MkB71Ikeg 12985
+dyXROdcPFSd36N3K7dpmDw 506 0
 #2 := false
 decl uf_17 :: (-> T8 T3)
 decl uf_18 :: (-> T1 T8)
@@ -13808,4 +13472,3 @@
 #325 := [asserted]: #108
 [unit-resolution #325 #554]: false
 unsat
-
--- a/src/HOL/SMT/SMT_Base.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/SMT/SMT_Base.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -8,6 +8,7 @@
 imports Real "~~/src/HOL/Word/Word"
   "~~/src/HOL/Decision_Procs/Dense_Linear_Order"
 uses
+  "~~/src/Tools/Cache_IO/cache_io.ML"
   ("Tools/smt_normalize.ML")
   ("Tools/smt_monomorph.ML")
   ("Tools/smt_translate.ML")
--- a/src/HOL/SMT/Tools/smt_solver.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/SMT/Tools/smt_solver.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -20,7 +20,7 @@
   type solver_config = {
     command: {env_var: string, remote_name: string option},
     arguments: string list,
-    interface: interface,
+    interface: string list -> interface,
     reconstruct: proof_data -> thm }
 
   (*options*)
@@ -28,8 +28,10 @@
   val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b
   val trace: bool Config.T
   val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit
+
+  (*certificates*)
   val record: bool Config.T
-  val certificates: string Config.T
+  val select_certificates: string -> Context.generic -> Context.generic
 
   (*solvers*)
   type solver = Proof.context -> thm list -> thm
@@ -71,7 +73,7 @@
 type solver_config = {
   command: {env_var: string, remote_name: string option},
   arguments: string list,
-  interface: interface,
+  interface: string list -> interface,
   reconstruct: proof_data -> thm }
 
 
@@ -88,30 +90,29 @@
 fun trace_msg ctxt f x =
   if Config.get ctxt trace then tracing (f x) else ()
 
+
+(* SMT certificates *)
+
 val (record, setup_record) = Attrib.config_bool "smt_record" true
-val no_certificates = ""
-val (certificates, setup_certificates) =
-  Attrib.config_string "smt_certificates" no_certificates
 
+structure Certificates = Generic_Data
+(
+  type T = Cache_IO.cache option
+  val empty = NONE
+  val extend = I
+  fun merge (s, _) = s
+)
+
+fun select_certificates name = Certificates.put (
+  if name = "" then NONE
+  else SOME (Cache_IO.make (Path.explode name)))
 
 
 (* interface to external solvers *)
 
 local
 
-fun with_files ctxt f =
-  let
-    val paths as (problem_path, proof_path) =
-      "smt-" ^ serial_string ()
-      |> (fn n => (n, n ^ ".proof"))
-      |> pairself (File.tmp_path o Path.explode)
-
-    val y = Exn.capture f (problem_path, proof_path)
-
-    val _ = pairself (try File.rm) paths
-  in Exn.release y end
-
-fun invoke ctxt output f (paths as (problem_path, proof_path)) =
+fun invoke ctxt output f problem_path =
   let
     fun pretty tag ls = Pretty.string_of (Pretty.big_list tag
       (map Pretty.str ls))
@@ -120,11 +121,10 @@
     val _ = trace_msg ctxt (pretty "SMT problem:" o split_lines o File.read)
       problem_path
 
-    val (s, _) = with_timeout ctxt f paths
-    val _ = trace_msg ctxt (pretty "SMT solver:") (split_lines s)
+    val (res, err) = with_timeout ctxt f problem_path
+    val _ = trace_msg ctxt (pretty "SMT solver:") err
 
-    fun lines_of path = the_default [] (try (File.fold_lines cons path) [])
-    val ls = rev (dropwhile (equal "") (lines_of proof_path))
+    val ls = rev (dropwhile (equal "") (rev res))
     val _ = trace_msg ctxt (pretty "SMT result:") ls
   in (x, ls) end
 
@@ -135,47 +135,53 @@
     val remote_url = getenv "REMOTE_SMT_URL"
   in
     if local_solver <> ""
-    then (["local", local_solver],
-      "Invoking local SMT solver " ^ quote local_solver ^ " ...")
-    else if remote_solver <> "" andalso remote_url <> ""
-    then (["remote", remote_solver],
-      "Invoking remote SMT solver " ^ quote remote_solver ^ " at " ^
-      quote remote_url ^ " ...")
+    then 
+     (tracing ("Invoking local SMT solver " ^ quote local_solver ^ " ...");
+      [local_solver])
+    else if remote_solver <> ""
+    then
+     (tracing ("Invoking remote SMT solver " ^ quote remote_solver ^ " at " ^
+        quote remote_url ^ " ...");
+      [getenv "REMOTE_SMT", remote_solver])
     else error ("Undefined Isabelle environment variable: " ^ quote env_var)
   end
 
-fun run ctxt cmd args (problem_path, proof_path) =
-  let
-    val certs = Config.get ctxt certificates
-    val certs' = 
-      if certs = no_certificates then "-"
-      else File.shell_path (Path.explode certs)
-    val (solver, msg) =
-      if certs = no_certificates orelse Config.get ctxt record
-      then choose cmd
-      else (["certificate"], "Using certificate from " ^ quote certs' ^ " ...")
-    val _ = tracing msg
+fun make_cmd solver args problem_path proof_path = space_implode " " (
+  map File.shell_quote (solver @ args) @
+  [File.shell_path problem_path, "2>&1", ">", File.shell_path proof_path])
+
+fun no_cmd _ _ = error ("Bad certificates cache: missing certificate")
+
+fun run ctxt cmd args problem_path =
+  let val certs = Certificates.get (Context.Proof ctxt)
   in
-    bash_output (space_implode " " (
-      File.shell_path (Path.explode (getenv "RUN_SMT_SOLVER")) :: certs' ::
-      map File.shell_quote (solver @ args) @
-      map File.shell_path [problem_path, proof_path]) ^ " 2>&1")
+    if is_none certs 
+    then Cache_IO.run' (make_cmd (choose cmd) args) problem_path
+    else if Config.get ctxt record
+    then Cache_IO.cached' (the certs) (make_cmd (choose cmd) args) problem_path
+    else
+     (tracing ("Using cached certificate from " ^
+        File.shell_path (Cache_IO.cache_path_of (the certs)) ^ " ...");
+      Cache_IO.cached' (the certs) no_cmd problem_path)
   end
 
 in
 
 fun run_solver ctxt cmd args output =
-  with_files ctxt (invoke ctxt output (run ctxt cmd args))
+  Cache_IO.with_tmp_file "smt-" (invoke ctxt output (run ctxt cmd args))
 
 end
 
 fun make_proof_data ctxt ((recon, thms), ls) =
   {context=ctxt, output=ls, recon=recon, assms=SOME thms}
 
-fun gen_solver solver ctxt prems =
+fun gen_solver name solver ctxt prems =
   let
     val {command, arguments, interface, reconstruct} = solver ctxt
-    val {normalize=nc, translate=tc} = interface
+    val comments = ("solver: " ^ name) ::
+      ("timeout: " ^ string_of_int (Config.get ctxt timeout)) ::
+      "arguments:" :: arguments
+    val {normalize=nc, translate=tc} = interface comments
     val thy = ProofContext.theory_of ctxt
   in
     SMT_Normalize.normalize nc ctxt prems
@@ -218,17 +224,19 @@
 
 val solver_name_of = Selected_Solver.get
 
-fun select_solver name gen =
-  if is_none (lookup_solver (Context.theory_of gen) name)
+fun select_solver name context =
+  if is_none (lookup_solver (Context.theory_of context) name)
   then error ("SMT solver not registered: " ^ quote name)
-  else Selected_Solver.map (K name) gen
+  else Selected_Solver.map (K name) context
 
-fun raw_solver_of gen =
-  (case lookup_solver (Context.theory_of gen) (solver_name_of gen) of
+fun raw_solver_of context name =
+  (case lookup_solver (Context.theory_of context) name of
     NONE => error "No SMT solver selected"
   | SOME (s, _) => s)
 
-val solver_of = gen_solver o raw_solver_of
+fun solver_of context =
+  let val name = solver_name_of context
+  in gen_solver name (raw_solver_of context name) end
 
 
 (* SMT tactic *)
@@ -278,7 +286,10 @@
   setup_timeout #>
   setup_trace #>
   setup_record #>
-  setup_certificates #>
+  Attrib.setup (Binding.name "smt_certificates")
+    (Scan.lift (OuterParse.$$$ "=" |-- Args.name) >>
+      (Thm.declaration_attribute o K o select_certificates))
+    "SMT certificates" #>
   Method.setup (Binding.name "smt") smt_method
     "Applies an SMT solver to the current goal."
 
--- a/src/HOL/SMT/Tools/smtlib_interface.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/SMT/Tools/smtlib_interface.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -8,9 +8,9 @@
 sig
   val basic_builtins: SMT_Translate.builtins
   val default_attributes: string list
-  val gen_interface: SMT_Translate.builtins -> string list ->
+  val gen_interface: SMT_Translate.builtins -> string list -> string list ->
     SMT_Solver.interface
-  val interface: SMT_Solver.interface
+  val interface: string list -> SMT_Solver.interface
 end
 
 structure SMTLIB_Interface: SMTLIB_INTERFACE =
@@ -118,12 +118,13 @@
           | wr_pat (T.SNoPat ts) = wrp "nopat" ts
       in par (wr_quant q #> fold wr_var vs #> wre b #> fold wr_pat ps) end)
 
-fun serialize attributes ({typs, funs, preds} : T.sign) ts stream =
+fun serialize attributes comments ({typs, funs, preds} : T.sign) ts stream =
   let
     fun wr_decl (n, Ts) = wr_line (sep (par (wr n #> fold wr1 Ts)))
   in
     stream
     |> wr_line (wr "(benchmark Isabelle")
+    |> wr_line (wr ":status unknown")
     |> fold (wr_line o wr) attributes
     |> length typs > 0 ?
          wr_line (wr ":extrasorts" #> par (fold wr1 typs))
@@ -138,6 +139,7 @@
     |> fold (fn t => wr ":assumption" #> wr_line (wr_expr false [] t)) ts
     |> wr_line (wr ":formula true")
     |> wr_line (wr ")")
+    |> fold (fn comment => wr_line (wr ("; " ^ comment))) comments
     |> K ()
   end
 
@@ -149,9 +151,9 @@
   builtin_num = builtin_num,
   builtin_fun = (fn true => builtin_funcs | false => builtin_preds) }
 
-val default_attributes = [":logic AUFLIRA", ":status unknown"]
+val default_attributes = [":logic AUFLIRA"]
 
-fun gen_interface builtins attributes = {
+fun gen_interface builtins attributes comments = {
   normalize = [
     SMT_Normalize.RewriteTrivialLets,
     SMT_Normalize.RewriteNegativeNumerals,
@@ -170,8 +172,9 @@
       term_marker = term_marker,
       formula_marker = formula_marker },
     builtins = builtins,
-    serialize = serialize attributes }}
+    serialize = serialize attributes comments }}
 
-val interface = gen_interface basic_builtins default_attributes
+fun interface comments =
+  gen_interface basic_builtins default_attributes comments
 
 end
--- a/src/HOL/SMT/Tools/z3_interface.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/SMT/Tools/z3_interface.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -6,7 +6,7 @@
 
 signature Z3_INTERFACE =
 sig
-  val interface: SMT_Solver.interface
+  val interface: string list -> SMT_Solver.interface
 end
 
 structure Z3_Interface: Z3_INTERFACE =
--- a/src/HOL/SMT/lib/scripts/run_smt_solver	Wed Feb 17 09:48:53 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,93 +0,0 @@
-#!/usr/bin/env perl
-#
-# Author: Sascha Boehme, TU Muenchen
-#
-# Cache for prover results, based on discriminating problems using MD5.
-
-use strict;
-use warnings;
-use Digest::MD5;
-
-
-# arguments
-
-my $certs_file = shift @ARGV;
-my $location = shift @ARGV;
-my $result_file = pop @ARGV;
-my $problem_file = $ARGV[-1];
-
-my $use_certs = not ($certs_file eq "-");
-
-
-# create MD5 problem digest
-
-my $problem_digest = "";
-if ($use_certs) {
-  my $md5 = Digest::MD5->new;
-  open FILE, "<$problem_file" or
-    die "ERROR: Failed to open '$problem_file' ($!)";
-  $md5->addfile(*FILE);
-  close FILE;
-  $problem_digest = $md5->b64digest;
-}
-
-
-# lookup problem digest among existing certificates and
-# return a cached result (if there is a certificate for the problem)
-
-if ($use_certs and -e $certs_file) {
-  my $cached = 0;
-  open CERTS, "<$certs_file" or die "ERROR: Failed to open '$certs_file' ($!)";
-  while (<CERTS>) {
-    if (m/(\S+) (\d+)/) {
-      if ($1 eq $problem_digest) {
-        my $start = tell CERTS;
-        open FILE, ">$result_file" or
-          die "ERROR: Failed to open '$result_file ($!)";
-        while ((tell CERTS) - $start < $2) {
-          my $line = <CERTS>;
-          print FILE $line;
-        }
-        close FILE;
-        $cached = 1;
-        last;
-      }
-      else { seek CERTS, $2, 1; }
-    }
-    else { die "ERROR: Invalid file format in '$certs_file'"; }
-  }
-  close CERTS;
-  if ($cached) { exit 0; }
-}
-
-
-# invoke (remote or local) prover
-
-my $result;
-
-if ($location eq "remote") {
-  $result = system "$ENV{'REMOTE_SMT'} @ARGV >$result_file 2>&1";
-}
-elsif ($location eq "local") {
-  $result = system "@ARGV >$result_file 2>&1";
-}
-else { die "ERROR: No SMT solver invoked"; }
-
-
-# cache result
-
-if ($use_certs) {
-  open CERTS, ">>$certs_file" or
-    die "ERROR: Failed to open '$certs_file' ($!)";
-  print CERTS
-    ("$problem_digest " . ((-s $result_file) + (length "\n")) . "\n");
-
-  open FILE, "<$result_file" or
-    die "ERROR: Failed to open '$result_file' ($!)";
-  foreach (<FILE>) { print CERTS $_; }
-  close FILE; 
-
-  print CERTS "\n";
-  close CERTS;
-}
-
--- a/src/HOL/Tools/Datatype/datatype.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -651,7 +651,7 @@
 
     val (tyvars, _, _, _)::_ = dts;
     val (new_dts, types_syntax) = ListPair.unzip (map (fn (tvs, tname, mx, _) =>
-      let val full_tname = Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname)
+      let val full_tname = Sign.full_name tmp_thy tname
       in
         (case duplicates (op =) tvs of
           [] =>
@@ -675,10 +675,10 @@
             val _ =
               (case subtract (op =) tvs (fold (curry OldTerm.add_typ_tfree_names) cargs' []) of
                 [] => ()
-              | vs => error ("Extra type variables on rhs: " ^ commas vs))
-          in (constrs @ [(Sign.full_name_path tmp_thy tname'
-                  (Binding.map_name (Syntax.const_name mx') cname),
-                   map (dtyp_of_typ new_dts) cargs')],
+              | vs => error ("Extra type variables on rhs: " ^ commas vs));
+            val c = Sign.full_name_path tmp_thy tname' cname;
+          in
+            (constrs @ [(c, map (dtyp_of_typ new_dts) cargs')],
               constr_syntax' @ [(cname, mx')], sorts'')
           end handle ERROR msg => cat_error msg
            ("The error above occured in constructor " ^ quote (Binding.str_of cname) ^
@@ -686,14 +686,12 @@
 
         val (constrs', constr_syntax', sorts') =
           fold prep_constr constrs ([], [], sorts)
-
       in
         case duplicates (op =) (map fst constrs') of
-           [] =>
-             (dts' @ [(i, (Sign.full_name tmp_thy (Binding.map_name (Syntax.type_name mx) tname),
-                map DtTFree tvs, constrs'))],
+          [] =>
+            (dts' @ [(i, (Sign.full_name tmp_thy tname, map DtTFree tvs, constrs'))],
               constr_syntax @ [constr_syntax'], sorts', i + 1)
-         | dups => error ("Duplicate constructors " ^ commas dups ^
+        | dups => error ("Duplicate constructors " ^ commas dups ^
              " in datatype " ^ quote (Binding.str_of tname))
       end;
 
--- a/src/HOL/Tools/Datatype/datatype_case.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Tools/Datatype/datatype_case.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -146,7 +146,7 @@
                         (replicate (length rstp) "x")
                     in
                       [((prfx, gvars @ map Free (xs ~~ Ts)),
-                        (Const (@{const_name HOL.undefined}, res_ty), (~1, false)))]
+                        (Const (@{const_syntax undefined}, res_ty), (~1, false)))]
                     end
                   else in_group
               in
@@ -315,7 +315,7 @@
       fun prep_pat ((c as Const (@{syntax_const "_constrain"}, _)) $ t $ tT) used =
             let val (t', used') = prep_pat t used
             in (c $ t' $ tT, used') end
-        | prep_pat (Const (@{const_name dummy_pattern}, T)) used =
+        | prep_pat (Const (@{const_syntax dummy_pattern}, T)) used =
             let val x = Name.variant used "x"
             in (Free (x, T), x :: used) end
         | prep_pat (Const (s, T)) used =
@@ -381,7 +381,7 @@
         fun count_cases (_, _, true) = I
           | count_cases (c, (_, body), false) =
               AList.map_default op aconv (body, []) (cons c);
-        val is_undefined = name_of #> equal (SOME @{const_name HOL.undefined});
+        val is_undefined = name_of #> equal (SOME @{const_syntax undefined});
         fun mk_case (c, (xs, body), _) = (list_comb (c, xs), body)
       in case ty_info tab cname of
           SOME {constructors, case_name} =>
@@ -399,7 +399,7 @@
                   (fold_rev count_cases cases []);
                 val R = type_of t;
                 val dummy =
-                  if d then Const (@{const_name dummy_pattern}, R)
+                  if d then Const (@{const_syntax dummy_pattern}, R)
                   else Free (Name.variant used "x", R)
               in
                 SOME (x, map mk_case (case find_first (is_undefined o fst) cases' of
--- a/src/HOL/Tools/record.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Tools/record.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -9,6 +9,18 @@
 
 signature BASIC_RECORD =
 sig
+  type record_info =
+   {args: (string * sort) list,
+    parent: (typ list * string) option,
+    fields: (string * typ) list,
+    extension: (string * typ list),
+    ext_induct: thm, ext_inject: thm, ext_surjective: thm, ext_split: thm, ext_def: thm,
+    select_convs: thm list, update_convs: thm list, select_defs: thm list, update_defs: thm list,
+    fold_congs: thm list, unfold_congs: thm list, splits: thm list, defs: thm list,
+    surjective: thm, equality: thm, induct_scheme: thm, induct: thm, cases_scheme: thm,
+    cases: thm, simps: thm list, iffs: thm list}
+  val get_record: theory -> string -> record_info option
+  val the_record: theory -> string -> record_info
   val record_simproc: simproc
   val record_eq_simproc: simproc
   val record_upd_simproc: simproc
@@ -42,10 +54,10 @@
   val print_records: theory -> unit
   val read_typ: Proof.context -> string -> (string * sort) list -> typ * (string * sort) list
   val cert_typ: Proof.context -> typ -> (string * sort) list -> typ * (string * sort) list
-  val add_record: bool -> string list * string -> string option -> (string * string * mixfix) list
-    -> theory -> theory
-  val add_record_i: bool -> string list * string -> (typ list * string) option
-    -> (string * typ * mixfix) list -> theory -> theory
+  val add_record: bool -> string list * binding -> (typ list * string) option ->
+    (binding * typ * mixfix) list -> theory -> theory
+  val add_record_cmd: bool -> string list * binding -> string option ->
+    (binding * string * mixfix) list -> theory -> theory
   val setup: theory -> theory
 end;
 
@@ -71,8 +83,7 @@
 
 fun named_cterm_instantiate values thm =
   let
-    fun match name ((name', _), _) = name = name'
-      | match name _ = false;
+    fun match name ((name', _), _) = name = name';
     fun getvar name =
       (case find_first (match name) (Term.add_vars (prop_of thm) []) of
         SOME var => cterm_of (theory_of_thm thm) (Var var)
@@ -93,8 +104,8 @@
   let
     fun get_thms thy name =
       let
-        val SOME {Rep_inject = rep_inject, Abs_name = absN, abs_type = absT,
-          Abs_inverse = abs_inverse, ...} = Typedef.get_info thy name;
+        val {Rep_inject = rep_inject, Abs_name = absN, abs_type = absT,
+          Abs_inverse = abs_inverse, ...} = Typedef.the_info thy name;
         val rewrite_rule =
           MetaSimplifier.rewrite_rule [@{thm iso_tuple_UNIV_I}, @{thm iso_tuple_True_simp}];
       in
@@ -142,7 +153,8 @@
     val ([isom_def], cdef_thy) =
       typ_thy
       |> Sign.add_consts_i [Syntax.no_syn (isom_bind, isomT)]
-      |> PureThy.add_defs false [Thm.no_attributes (apfst Binding.name isom_spec)];
+      |> PureThy.add_defs false
+        [Thm.no_attributes (apfst (Binding.conceal o Binding.name) isom_spec)];
 
     val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro));
     val cons = Const (@{const_name iso_tuple_cons}, isomT --> leftT --> rightT --> absT);
@@ -156,7 +168,8 @@
     ((isom, cons $ isom), thm_thy)
   end;
 
-val iso_tuple_intros_tac = resolve_from_net_tac iso_tuple_intros THEN'
+val iso_tuple_intros_tac =
+  resolve_from_net_tac iso_tuple_intros THEN'
   CSUBGOAL (fn (cgoal, i) =>
     let
       val thy = Thm.theory_of_cterm cgoal;
@@ -189,7 +202,7 @@
 val meta_allE = @{thm Pure.meta_allE};
 val prop_subst = @{thm prop_subst};
 val K_record_comp = @{thm K_record_comp};
-val K_comp_convs = [@{thm o_apply}, K_record_comp]
+val K_comp_convs = [@{thm o_apply}, K_record_comp];
 val o_assoc = @{thm o_assoc};
 val id_apply = @{thm id_apply};
 val id_o_apps = [@{thm id_apply}, @{thm id_o}, @{thm o_id}];
@@ -197,21 +210,21 @@
 
 val refl_conj_eq = @{thm refl_conj_eq};
 
-val surject_assistI = @{thm "iso_tuple_surjective_proof_assistI"};
-val surject_assist_idE = @{thm "iso_tuple_surjective_proof_assist_idE"};
-
-val updacc_accessor_eqE = @{thm "update_accessor_accessor_eqE"};
-val updacc_updator_eqE = @{thm "update_accessor_updator_eqE"};
-val updacc_eq_idI = @{thm "iso_tuple_update_accessor_eq_assist_idI"};
-val updacc_eq_triv = @{thm "iso_tuple_update_accessor_eq_assist_triv"};
-
-val updacc_foldE = @{thm "update_accessor_congruence_foldE"};
-val updacc_unfoldE = @{thm "update_accessor_congruence_unfoldE"};
-val updacc_noopE = @{thm "update_accessor_noopE"};
-val updacc_noop_compE = @{thm "update_accessor_noop_compE"};
-val updacc_cong_idI = @{thm "update_accessor_cong_assist_idI"};
-val updacc_cong_triv = @{thm "update_accessor_cong_assist_triv"};
-val updacc_cong_from_eq = @{thm "iso_tuple_update_accessor_cong_from_eq"};
+val surject_assistI = @{thm iso_tuple_surjective_proof_assistI};
+val surject_assist_idE = @{thm iso_tuple_surjective_proof_assist_idE};
+
+val updacc_accessor_eqE = @{thm update_accessor_accessor_eqE};
+val updacc_updator_eqE = @{thm update_accessor_updator_eqE};
+val updacc_eq_idI = @{thm iso_tuple_update_accessor_eq_assist_idI};
+val updacc_eq_triv = @{thm iso_tuple_update_accessor_eq_assist_triv};
+
+val updacc_foldE = @{thm update_accessor_congruence_foldE};
+val updacc_unfoldE = @{thm update_accessor_congruence_unfoldE};
+val updacc_noopE = @{thm update_accessor_noopE};
+val updacc_noop_compE = @{thm update_accessor_noop_compE};
+val updacc_cong_idI = @{thm update_accessor_cong_assist_idI};
+val updacc_cong_triv = @{thm update_accessor_cong_assist_triv};
+val updacc_cong_from_eq = @{thm iso_tuple_update_accessor_cong_from_eq};
 
 val o_eq_dest = @{thm o_eq_dest};
 val o_eq_id_dest = @{thm o_eq_id_dest};
@@ -257,11 +270,9 @@
 val Trueprop = HOLogic.mk_Trueprop;
 fun All xs t = Term.list_all_free (xs, t);
 
-infix 9 $$;
 infix 0 :== ===;
 infixr 0 ==>;
 
-val op $$ = Term.list_comb;
 val op :== = Primitive_Defs.mk_defpair;
 val op === = Trueprop o HOLogic.mk_eq;
 val op ==> = Logic.mk_implies;
@@ -337,24 +348,55 @@
   parent: (typ list * string) option,
   fields: (string * typ) list,
   extension: (string * typ list),
+
+  ext_induct: thm,
+  ext_inject: thm,
+  ext_surjective: thm,
+  ext_split: thm,
+  ext_def: thm,
+
+  select_convs: thm list,
+  update_convs: thm list,
+  select_defs: thm list,
+  update_defs: thm list,
+  fold_congs: thm list,
+  unfold_congs: thm list,
+  splits: thm list,
+  defs: thm list,
+
+  surjective: thm,
+  equality: thm,
+  induct_scheme: thm,
   induct: thm,
-  extdef: thm};
-
-fun make_record_info args parent fields extension induct extdef =
+  cases_scheme: thm,
+  cases: thm,
+
+  simps: thm list,
+  iffs: thm list};
+
+fun make_record_info args parent fields extension
+    ext_induct ext_inject ext_surjective ext_split ext_def
+    select_convs update_convs select_defs update_defs fold_congs unfold_congs splits defs
+    surjective equality induct_scheme induct cases_scheme cases
+    simps iffs : record_info =
  {args = args, parent = parent, fields = fields, extension = extension,
-  induct = induct, extdef = extdef}: record_info;
-
+  ext_induct = ext_induct, ext_inject = ext_inject, ext_surjective = ext_surjective,
+  ext_split = ext_split, ext_def = ext_def, select_convs = select_convs,
+  update_convs = update_convs, select_defs = select_defs, update_defs = update_defs,
+  fold_congs = fold_congs, unfold_congs = unfold_congs, splits = splits, defs = defs,
+  surjective = surjective, equality = equality, induct_scheme = induct_scheme,
+  induct = induct, cases_scheme = cases_scheme, cases = cases, simps = simps, iffs = iffs};
 
 type parent_info =
  {name: string,
   fields: (string * typ) list,
   extension: (string * typ list),
-  induct: thm,
-  extdef: thm};
-
-fun make_parent_info name fields extension induct extdef =
+  induct_scheme: thm,
+  ext_def: thm};
+
+fun make_parent_info name fields extension ext_def induct_scheme : parent_info =
  {name = name, fields = fields, extension = extension,
-  induct = induct, extdef = extdef}: parent_info;
+  ext_def = ext_def, induct_scheme = induct_scheme};
 
 
 (* theory data *)
@@ -371,7 +413,7 @@
   equalities: thm Symtab.table,
   extinjects: thm list,
   extsplit: thm Symtab.table,  (*maps extension name to split rule*)
-  splits: (thm * thm * thm * thm) Symtab.table,  (*!!, !, EX - split-equalities, induct rule*)
+  splits: (thm * thm * thm * thm) Symtab.table,  (*!!, ALL, EX - split-equalities, induct rule*)
   extfields: (string * typ) list Symtab.table,  (*maps extension to its fields*)
   fieldext: (string * typ list) Symtab.table};  (*maps field to its extension*)
 
@@ -381,7 +423,7 @@
   equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits,
   extfields = extfields, fieldext = fieldext }: record_data;
 
-structure RecordsData = Theory_Data
+structure Records_Data = Theory_Data
 (
   type T = record_data;
   val empty =
@@ -434,7 +476,7 @@
 
 fun print_records thy =
   let
-    val {records = recs, ...} = RecordsData.get thy;
+    val {records = recs, ...} = Records_Data.get thy;
     val prt_typ = Syntax.pretty_typ_global thy;
 
     fun pretty_parent NONE = []
@@ -454,20 +496,25 @@
 
 (* access 'records' *)
 
-val get_record = Symtab.lookup o #records o RecordsData.get;
+val get_record = Symtab.lookup o #records o Records_Data.get;
+
+fun the_record thy name =
+  (case get_record thy name of
+    SOME info => info
+  | NONE => error ("Unknown record type " ^ quote name));
 
 fun put_record name info thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
+      Records_Data.get thy;
     val data = make_record_data (Symtab.update (name, info) records)
       sel_upd equalities extinjects extsplit splits extfields fieldext;
-  in RecordsData.put data thy end;
+  in Records_Data.put data thy end;
 
 
 (* access 'sel_upd' *)
 
-val get_sel_upd = #sel_upd o RecordsData.get;
+val get_sel_upd = #sel_upd o Records_Data.get;
 
 val is_selector = Symtab.defined o #selectors o get_sel_upd;
 val get_updates = Symtab.lookup o #updates o get_sel_upd;
@@ -492,7 +539,7 @@
     val upds = map (suffix updateN) all ~~ all;
 
     val {records, sel_upd = {selectors, updates, simpset, defset, foldcong, unfoldcong},
-      equalities, extinjects, extsplit, splits, extfields, fieldext} = RecordsData.get thy;
+      equalities, extinjects, extsplit, splits, extfields, fieldext} = Records_Data.get thy;
     val data = make_record_data records
       {selectors = fold Symtab.update_new sels selectors,
         updates = fold Symtab.update_new upds updates,
@@ -501,7 +548,7 @@
         foldcong = foldcong addcongs folds,
         unfoldcong = unfoldcong addcongs unfolds}
        equalities extinjects extsplit splits extfields fieldext;
-  in RecordsData.put data thy end;
+  in Records_Data.put data thy end;
 
 
 (* access 'equalities' *)
@@ -509,12 +556,12 @@
 fun add_record_equalities name thm thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
+      Records_Data.get thy;
     val data = make_record_data records sel_upd
       (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext;
-  in RecordsData.put data thy end;
-
-val get_equalities = Symtab.lookup o #equalities o RecordsData.get;
+  in Records_Data.put data thy end;
+
+val get_equalities = Symtab.lookup o #equalities o Records_Data.get;
 
 
 (* access 'extinjects' *)
@@ -522,13 +569,13 @@
 fun add_extinjects thm thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
+      Records_Data.get thy;
     val data =
       make_record_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects)
         extsplit splits extfields fieldext;
-  in RecordsData.put data thy end;
-
-val get_extinjects = rev o #extinjects o RecordsData.get;
+  in Records_Data.put data thy end;
+
+val get_extinjects = rev o #extinjects o Records_Data.get;
 
 
 (* access 'extsplit' *)
@@ -536,11 +583,11 @@
 fun add_extsplit name thm thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
-    val data = make_record_data records sel_upd
-      equalities extinjects (Symtab.update_new (name, thm) extsplit) splits
-      extfields fieldext;
-  in RecordsData.put data thy end;
+      Records_Data.get thy;
+    val data =
+      make_record_data records sel_upd equalities extinjects
+        (Symtab.update_new (name, thm) extsplit) splits extfields fieldext;
+  in Records_Data.put data thy end;
 
 
 (* access 'splits' *)
@@ -548,19 +595,19 @@
 fun add_record_splits name thmP thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
-    val data = make_record_data records sel_upd
-      equalities extinjects extsplit (Symtab.update_new (name, thmP) splits)
-      extfields fieldext;
-  in RecordsData.put data thy end;
-
-val get_splits = Symtab.lookup o #splits o RecordsData.get;
+      Records_Data.get thy;
+    val data =
+      make_record_data records sel_upd equalities extinjects extsplit
+        (Symtab.update_new (name, thmP) splits) extfields fieldext;
+  in Records_Data.put data thy end;
+
+val get_splits = Symtab.lookup o #splits o Records_Data.get;
 
 
 (* parent/extension of named record *)
 
-val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o RecordsData.get);
-val get_extension = Option.map #extension oo (Symtab.lookup o #records o RecordsData.get);
+val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Records_Data.get);
+val get_extension = Option.map #extension oo (Symtab.lookup o #records o Records_Data.get);
 
 
 (* access 'extfields' *)
@@ -568,14 +615,13 @@
 fun add_extfields name fields thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
+      Records_Data.get thy;
     val data =
-      make_record_data records sel_upd
-        equalities extinjects extsplit splits
+      make_record_data records sel_upd equalities extinjects extsplit splits
         (Symtab.update_new (name, fields) extfields) fieldext;
-  in RecordsData.put data thy end;
-
-val get_extfields = Symtab.lookup o #extfields o RecordsData.get;
+  in Records_Data.put data thy end;
+
+val get_extfields = Symtab.lookup o #extfields o Records_Data.get;
 
 fun get_extT_fields thy T =
   let
@@ -585,21 +631,21 @@
       in Long_Name.implode (rev (nm :: rst)) end;
     val midx = maxidx_of_typs (moreT :: Ts);
     val varifyT = varifyT midx;
-    val {records, extfields, ...} = RecordsData.get thy;
-    val (flds, (more, _)) = split_last (Symtab.lookup_list extfields name);
+    val {records, extfields, ...} = Records_Data.get thy;
+    val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name);
     val args = map varifyT (snd (#extension (the (Symtab.lookup records recname))));
 
-    val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) (Vartab.empty);
-    val flds' = map (apsnd ((Envir.norm_type subst) o varifyT)) flds;
-  in (flds', (more, moreT)) end;
+    val subst = fold (Sign.typ_match thy) (but_last args ~~ but_last Ts) Vartab.empty;
+    val fields' = map (apsnd (Envir.norm_type subst o varifyT)) fields;
+  in (fields', (more, moreT)) end;
 
 fun get_recT_fields thy T =
   let
-    val (root_flds, (root_more, root_moreT)) = get_extT_fields thy T;
-    val (rest_flds, rest_more) =
+    val (root_fields, (root_more, root_moreT)) = get_extT_fields thy T;
+    val (rest_fields, rest_more) =
       if is_recT root_moreT then get_recT_fields thy root_moreT
       else ([], (root_more, root_moreT));
-  in (root_flds @ rest_flds, rest_more) end;
+  in (root_fields @ rest_fields, rest_more) end;
 
 
 (* access 'fieldext' *)
@@ -607,15 +653,15 @@
 fun add_fieldext extname_types fields thy =
   let
     val {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} =
-      RecordsData.get thy;
+      Records_Data.get thy;
     val fieldext' =
       fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext;
     val data =
       make_record_data records sel_upd equalities extinjects
         extsplit splits extfields fieldext';
-  in RecordsData.put data thy end;
-
-val get_fieldext = Symtab.lookup o #fieldext o RecordsData.get;
+  in Records_Data.put data thy end;
+
+val get_fieldext = Symtab.lookup o #fieldext o Records_Data.get;
 
 
 (* parent records *)
@@ -625,7 +671,7 @@
       let
         fun err msg = error (msg ^ " parent record " ^ quote name);
 
-        val {args, parent, fields, extension, induct, extdef} =
+        val {args, parent, fields, extension, induct_scheme, ext_def, ...} =
           (case get_record thy name of SOME info => info | NONE => err "Unknown");
         val _ = if length types <> length args then err "Bad number of arguments for" else ();
 
@@ -641,7 +687,7 @@
         val extension' = apsnd (map subst) extension;
       in
         add_parents thy parent'
-          (make_parent_info name fields' extension' induct extdef :: parents)
+          (make_parent_info name fields' extension' ext_def induct_scheme :: parents)
       end;
 
 
@@ -652,7 +698,8 @@
 
 fun decode_type thy t =
   let
-    fun get_sort xs n = AList.lookup (op =) xs (n: indexname) |> the_default (Sign.defaultS thy);
+    fun get_sort env xi =
+      the_default (Sign.defaultS thy) (AList.lookup (op =) env (xi: indexname));
     val map_sort = Sign.intern_sort thy;
   in
     Syntax.typ_of_term (get_sort (Syntax.term_sorts map_sort t)) map_sort t
@@ -662,149 +709,138 @@
 
 (* parse translations *)
 
-fun gen_field_tr mark sfx (t as Const (c, _) $ Const (name, _) $ arg) =
-      if c = mark then Syntax.const (suffix sfx name) $ Abs ("_", dummyT, arg)
-      else raise TERM ("gen_field_tr: " ^ mark, [t])
-  | gen_field_tr mark _ t = raise TERM ("gen_field_tr: " ^ mark, [t]);
-
-fun gen_fields_tr sep mark sfx (tm as Const (c, _) $ t $ u) =
-      if c = sep then gen_field_tr mark sfx t :: gen_fields_tr sep mark sfx u
-      else [gen_field_tr mark sfx tm]
-  | gen_fields_tr _ mark sfx tm = [gen_field_tr mark sfx tm];
-
-
-fun record_update_tr [t, u] =
-      Library.foldr (op $) (rev (gen_fields_tr "_updates" "_update" updateN u), t)
-  | record_update_tr ts = raise TERM ("record_update_tr", ts);
-
-fun update_name_tr (Free (x, T) :: ts) = Free (suffix updateN x, T) $$ ts
-  | update_name_tr (Const (x, T) :: ts) = Const (suffix updateN x, T) $$ ts
-  | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
-      (c $ update_name_tr [t] $ (Syntax.const "fun" $ ty $ Syntax.const "dummy")) $$ ts
-  | update_name_tr ts = raise TERM ("update_name_tr", ts);
-
-fun dest_ext_field mark (t as (Const (c, _) $ Const (name, _) $ arg)) =
-      if c = mark then (name, arg)
-      else raise TERM ("dest_ext_field: " ^ mark, [t])
-  | dest_ext_field _ t = raise TERM ("dest_ext_field", [t]);
-
-fun dest_ext_fields sep mark (trm as (Const (c, _) $ t $ u)) =
-      if c = sep then dest_ext_field mark t :: dest_ext_fields sep mark u
-      else [dest_ext_field mark trm]
-  | dest_ext_fields _ mark t = [dest_ext_field mark t];
-
-fun gen_ext_fields_tr sep mark sfx more ctxt t =
+local
+
+fun field_type_tr ((Const (@{syntax_const "_field_type"}, _) $ Const (name, _) $ arg)) =
+      (name, arg)
+  | field_type_tr t = raise TERM ("field_type_tr", [t]);
+
+fun field_types_tr (Const (@{syntax_const "_field_types"}, _) $ t $ u) =
+      field_type_tr t :: field_types_tr u
+  | field_types_tr t = [field_type_tr t];
+
+fun record_field_types_tr more ctxt t =
   let
     val thy = ProofContext.theory_of ctxt;
-    val msg = "error in record input: ";
-
-    val fieldargs = dest_ext_fields sep mark t;
-    fun splitargs (field :: fields) ((name, arg) :: fargs) =
-          if can (unsuffix name) field
-          then
-            let val (args, rest) = splitargs fields fargs
+    fun err msg = raise TERM ("Error in record-type input: " ^ msg, [t]);
+
+    fun split_args (field :: fields) ((name, arg) :: fargs) =
+          if can (unsuffix name) field then
+            let val (args, rest) = split_args fields fargs
             in (arg :: args, rest) end
-          else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t])
-      | splitargs [] (fargs as (_ :: _)) = ([], fargs)
-      | splitargs (_ :: _) [] = raise TERM (msg ^ "expecting more fields", [t])
-      | splitargs _ _ = ([], []);
-
-    fun mk_ext (fargs as (name, _) :: _) =
-          (case get_fieldext thy (Sign.intern_const thy name) of
-            SOME (ext, _) =>
-              (case get_extfields thy ext of
-                SOME flds =>
-                  let
-                    val (args, rest) = splitargs (map fst (but_last flds)) fargs;
-                    val more' = mk_ext rest;
-                  in list_comb (Syntax.const (suffix sfx ext), args @ [more']) end
-              | NONE => raise TERM (msg ^ "no fields defined for " ^ ext, [t]))
-          | NONE => raise TERM (msg ^ name ^" is no proper field", [t]))
-      | mk_ext [] = more;
-  in mk_ext fieldargs end;
-
-fun gen_ext_type_tr sep mark sfx more ctxt t =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val msg = "error in record-type input: ";
-
-    val fieldargs = dest_ext_fields sep mark t;
-    fun splitargs (field :: fields) ((name, arg) :: fargs) =
-          if can (unsuffix name) field then
-            let val (args, rest) = splitargs fields fargs
-            in (arg :: args, rest) end
-          else raise TERM (msg ^ "expecting field " ^ field ^ " but got " ^ name, [t])
-      | splitargs [] (fargs as (_ :: _)) = ([], fargs)
-      | splitargs (_ :: _) [] = raise TERM (msg ^ "expecting more fields", [t])
-      | splitargs _ _ = ([], []);
+          else err ("expecting field " ^ field ^ " but got " ^ name)
+      | split_args [] (fargs as (_ :: _)) = ([], fargs)
+      | split_args (_ :: _) [] = err "expecting more fields"
+      | split_args _ _ = ([], []);
 
     fun mk_ext (fargs as (name, _) :: _) =
           (case get_fieldext thy (Sign.intern_const thy name) of
             SOME (ext, alphas) =>
               (case get_extfields thy ext of
-                SOME flds =>
-                 (let
-                    val flds' = but_last flds;
-                    val types = map snd flds';
-                    val (args, rest) = splitargs (map fst flds') fargs;
+                SOME fields =>
+                  let
+                    val fields' = but_last fields;
+                    val types = map snd fields';
+                    val (args, rest) = split_args (map fst fields') fargs;
                     val argtypes = map (Sign.certify_typ thy o decode_type thy) args;
                     val midx = fold Term.maxidx_typ argtypes 0;
                     val varifyT = varifyT midx;
                     val vartypes = map varifyT types;
 
-                    val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes) Vartab.empty;
+                    val subst = fold (Sign.typ_match thy) (vartypes ~~ argtypes) Vartab.empty
+                      handle Type.TYPE_MATCH => err "type is no proper record (extension)";
                     val alphas' =
                       map (Syntax.term_of_typ (! Syntax.show_sorts) o Envir.norm_type subst o varifyT)
                         (but_last alphas);
 
                     val more' = mk_ext rest;
                   in
-                    list_comb (Syntax.const (suffix sfx ext), alphas' @ [more'])
-                  end handle Type.TYPE_MATCH =>
-                    raise TERM (msg ^ "type is no proper record (extension)", [t]))
-              | NONE => raise TERM (msg ^ "no fields defined for " ^ ext, [t]))
-          | NONE => raise TERM (msg ^ name ^" is no proper field", [t]))
+                    (* FIXME authentic syntax *)
+                    list_comb (Syntax.const (suffix ext_typeN ext), alphas' @ [more'])
+                  end
+              | NONE => err ("no fields defined for " ^ ext))
+          | NONE => err (name ^ " is no proper field"))
       | mk_ext [] = more;
-
-  in mk_ext fieldargs end;
-
-fun gen_adv_record_tr sep mark sfx unit ctxt [t] =
-      gen_ext_fields_tr sep mark sfx unit ctxt t
-  | gen_adv_record_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
-
-fun gen_adv_record_scheme_tr sep mark sfx ctxt [t, more] =
-      gen_ext_fields_tr sep mark sfx more ctxt t
-  | gen_adv_record_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
-
-fun gen_adv_record_type_tr sep mark sfx unit ctxt [t] =
-      gen_ext_type_tr sep mark sfx unit ctxt t
-  | gen_adv_record_type_tr _ _ _ _ _ ts = raise TERM ("gen_record_tr", ts);
-
-fun gen_adv_record_type_scheme_tr sep mark sfx ctxt [t, more] =
-      gen_ext_type_tr sep mark sfx more ctxt t
-  | gen_adv_record_type_scheme_tr _ _ _ _ ts = raise TERM ("gen_record_scheme_tr", ts);
-
-val adv_record_tr = gen_adv_record_tr "_fields" "_field" extN HOLogic.unit;
-
-val adv_record_scheme_tr = gen_adv_record_scheme_tr "_fields" "_field" extN;
-
-val adv_record_type_tr =
-  gen_adv_record_type_tr "_field_types" "_field_type" ext_typeN
-    (Syntax.term_of_typ false (HOLogic.unitT));
-
-val adv_record_type_scheme_tr =
-  gen_adv_record_type_scheme_tr "_field_types" "_field_type" ext_typeN;
-
+  in
+    mk_ext (field_types_tr t)
+  end;
+
+(* FIXME @{type_syntax} *)
+fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const "Product_Type.unit") ctxt t
+  | record_type_tr _ ts = raise TERM ("record_type_tr", ts);
+
+fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t
+  | record_type_scheme_tr _ ts = raise TERM ("record_type_scheme_tr", ts);
+
+
+fun field_tr ((Const (@{syntax_const "_field"}, _) $ Const (name, _) $ arg)) = (name, arg)
+  | field_tr t = raise TERM ("field_tr", [t]);
+
+fun fields_tr (Const (@{syntax_const "_fields"}, _) $ t $ u) = field_tr t :: fields_tr u
+  | fields_tr t = [field_tr t];
+
+fun record_fields_tr more ctxt t =
+  let
+    val thy = ProofContext.theory_of ctxt;
+    fun err msg = raise TERM ("Error in record input: " ^ msg, [t]);
+
+    fun split_args (field :: fields) ((name, arg) :: fargs) =
+          if can (unsuffix name) field
+          then
+            let val (args, rest) = split_args fields fargs
+            in (arg :: args, rest) end
+          else err ("expecting field " ^ field ^ " but got " ^ name)
+      | split_args [] (fargs as (_ :: _)) = ([], fargs)
+      | split_args (_ :: _) [] = err "expecting more fields"
+      | split_args _ _ = ([], []);
+
+    fun mk_ext (fargs as (name, _) :: _) =
+          (case get_fieldext thy (Sign.intern_const thy name) of
+            SOME (ext, _) =>
+              (case get_extfields thy ext of
+                SOME fields =>
+                  let
+                    val (args, rest) = split_args (map fst (but_last fields)) fargs;
+                    val more' = mk_ext rest;
+                  in
+                    (* FIXME authentic syntax *)
+                    list_comb (Syntax.const (suffix extN ext), args @ [more'])
+                  end
+              | NONE => err ("no fields defined for " ^ ext))
+          | NONE => err (name ^ " is no proper field"))
+      | mk_ext [] = more;
+  in mk_ext (fields_tr t) end;
+
+fun record_tr ctxt [t] = record_fields_tr (Syntax.const @{const_syntax Unity}) ctxt t
+  | record_tr _ ts = raise TERM ("record_tr", ts);
+
+fun record_scheme_tr ctxt [t, more] = record_fields_tr more ctxt t
+  | record_scheme_tr _ ts = raise TERM ("record_scheme_tr", ts);
+
+
+fun field_update_tr (Const (@{syntax_const "_field_update"}, _) $ Const (name, _) $ arg) =
+      Syntax.const (suffix updateN name) $ Abs ("_", dummyT, arg)
+  | field_update_tr t = raise TERM ("field_update_tr", [t]);
+
+fun field_updates_tr (Const (@{syntax_const "_field_updates"}, _) $ t $ u) =
+      field_update_tr t :: field_updates_tr u
+  | field_updates_tr t = [field_update_tr t];
+
+fun record_update_tr [t, u] = fold (curry op $) (field_updates_tr u) t
+  | record_update_tr ts = raise TERM ("record_update_tr", ts);
+
+in
 
 val parse_translation =
- [("_record_update", record_update_tr),
-  ("_update_name", update_name_tr)];
-
-val adv_parse_translation =
- [("_record", adv_record_tr),
-  ("_record_scheme", adv_record_scheme_tr),
-  ("_record_type", adv_record_type_tr),
-  ("_record_type_scheme", adv_record_type_scheme_tr)];
+ [(@{syntax_const "_record_update"}, record_update_tr)];
+
+val advanced_parse_translation =
+ [(@{syntax_const "_record"}, record_tr),
+  (@{syntax_const "_record_scheme"}, record_scheme_tr),
+  (@{syntax_const "_record_type"}, record_type_tr),
+  (@{syntax_const "_record_type_scheme"}, record_type_scheme_tr)];
+
+end;
 
 
 (* print translations *)
@@ -812,7 +848,9 @@
 val print_record_type_abbr = Unsynchronized.ref true;
 val print_record_type_as_fields = Unsynchronized.ref true;
 
-fun gen_field_upds_tr' mark sfx (tm as Const (name_field, _) $ k $ u) =
+local
+
+fun field_updates_tr' (tm as Const (c, _) $ k $ u) =
       let
         val t =
           (case k of
@@ -821,86 +859,147 @@
           | Abs (_, _, t) =>
               if null (loose_bnos t) then t else raise Match
           | _ => raise Match);
-
-          (* FIXME ? *)
-          (* (case k of (Const ("K_record", _) $ t) => t
-                   | Abs (x, _, Const ("K_record", _) $ t $ Bound 0) => t
-                   | _ => raise Match)*)
       in
-        (case try (unsuffix sfx) name_field of
+        (case try (unsuffix updateN) c of
           SOME name =>
-            apfst (cons (Syntax.const mark $ Syntax.free name $ t)) (gen_field_upds_tr' mark sfx u)
+            (* FIXME check wrt. record data (??) *)
+            (* FIXME early extern (!??) *)
+            apfst (cons (Syntax.const @{syntax_const "_field_update"} $ Syntax.free name $ t))
+              (field_updates_tr' u)
         | NONE => ([], tm))
       end
-  | gen_field_upds_tr' _ _ tm = ([], tm);
+  | field_updates_tr' tm = ([], tm);
 
 fun record_update_tr' tm =
-  let val (ts, u) = gen_field_upds_tr' "_update" updateN tm in
-    if null ts then raise Match
-    else
-      Syntax.const "_record_update" $ u $
-        foldr1 (fn (v, w) => Syntax.const "_updates" $ v $ w) (rev ts)
-  end;
-
-fun gen_field_tr' sfx tr' name =
-  let val name_sfx = suffix sfx name
-  in (name_sfx, fn [t, u] => tr' (Syntax.const name_sfx $ t $ u) | _ => raise Match) end;
-
-fun record_tr' sep mark record record_scheme unit ctxt t =
+  (case field_updates_tr' tm of
+    ([], _) => raise Match
+  | (ts, u) =>
+      Syntax.const @{syntax_const "_record_update"} $ u $
+        foldr1 (fn (v, w) => Syntax.const @{syntax_const "_field_updates"} $ v $ w) (rev ts));
+
+in
+
+fun field_update_tr' name =
+  let
+    (* FIXME authentic syntax *)
+    val update_name = suffix updateN name;
+    fun tr' [t, u] = record_update_tr' (Syntax.const update_name $ t $ u)
+      | tr' _ = raise Match;
+  in (update_name, tr') end;
+
+end;
+
+
+local
+
+(* FIXME early extern (!??) *)
+(* FIXME Syntax.free (??) *)
+fun field_tr' (c, t) = Syntax.const @{syntax_const "_field"} $ Syntax.const c $ t;
+
+fun fields_tr' (t, u) = Syntax.const @{syntax_const "_fields"} $ t $ u;
+
+fun record_tr' ctxt t =
   let
     val thy = ProofContext.theory_of ctxt;
 
-    fun field_lst t =
+    fun strip_fields t =
       (case strip_comb t of
         (Const (ext, _), args as (_ :: _)) =>
-          (case try (unsuffix extN) (Sign.intern_const thy ext) of
+          (case try (unsuffix extN) (Sign.intern_const thy ext) of  (* FIXME authentic syntax *)
             SOME ext' =>
               (case get_extfields thy ext' of
-                SOME flds =>
+                SOME fields =>
                  (let
-                    val f :: fs = but_last (map fst flds);
-                    val flds' = Sign.extern_const thy f :: map Long_Name.base_name fs;
+                    val f :: fs = but_last (map fst fields);
+                    val fields' = Sign.extern_const thy f :: map Long_Name.base_name fs;
                     val (args', more) = split_last args;
-                  in (flds' ~~ args') @ field_lst more end
+                  in (fields' ~~ args') @ strip_fields more end
                   handle Library.UnequalLengths => [("", t)])
               | NONE => [("", t)])
           | NONE => [("", t)])
        | _ => [("", t)]);
 
-    val (flds, (_, more)) = split_last (field_lst t);
-    val _ = if null flds then raise Match else ();
-    val flds' = map (fn (n, t) => Syntax.const mark $ Syntax.const n $ t) flds;
-    val flds'' = foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds';
+    val (fields, (_, more)) = split_last (strip_fields t);
+    val _ = null fields andalso raise Match;
+    val u = foldr1 fields_tr' (map field_tr' fields);
   in
-    if unit more
-    then Syntax.const record $ flds''
-    else Syntax.const record_scheme $ flds'' $ more
+    case more of
+      Const (@{const_syntax Unity}, _) => Syntax.const @{syntax_const "_record"} $ u
+    | _ => Syntax.const @{syntax_const "_record_scheme"} $ u $ more
   end;
 
-fun gen_record_tr' name =
+in
+
+fun record_ext_tr' name =
+  let
+    val ext_name = suffix extN name;
+    fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts));
+  in (ext_name, tr') end;
+
+end;
+
+
+local
+
+(* FIXME early extern (!??) *)
+(* FIXME Syntax.free (??) *)
+fun field_type_tr' (c, t) = Syntax.const @{syntax_const "_field_type"} $ Syntax.const c $ t;
+
+fun field_types_tr' (t, u) = Syntax.const @{syntax_const "_field_types"} $ t $ u;
+
+fun record_type_tr' ctxt t =
   let
-    val name_sfx = suffix extN name;
-    val unit = (fn Const (@{const_syntax "Product_Type.Unity"}, _) => true | _ => false);
-    fun tr' ctxt ts =
-      record_tr' "_fields" "_field" "_record" "_record_scheme" unit ctxt
-        (list_comb (Syntax.const name_sfx, ts));
-  in (name_sfx, tr') end;
-
-fun print_translation names =
-  map (gen_field_tr' updateN record_update_tr') names;
-
-
-(* record_type_abbr_tr' *)
+    val thy = ProofContext.theory_of ctxt;
+
+    val T = decode_type thy t;
+    val varifyT = varifyT (Term.maxidx_of_typ T);
+
+    val term_of_type = Syntax.term_of_typ (! Syntax.show_sorts) o Sign.extern_typ thy;
+
+    fun strip_fields T =
+      (case T of
+        Type (ext, args) =>
+          (case try (unsuffix ext_typeN) ext of
+            SOME ext' =>
+              (case get_extfields thy ext' of
+                SOME fields =>
+                  (case get_fieldext thy (fst (hd fields)) of
+                    SOME (_, alphas) =>
+                     (let
+                        val f :: fs = but_last fields;
+                        val fields' =
+                          apfst (Sign.extern_const thy) f :: map (apfst Long_Name.base_name) fs;
+                        val (args', more) = split_last args;
+                        val alphavars = map varifyT (but_last alphas);
+                        val subst = fold (Sign.typ_match thy) (alphavars ~~ args') Vartab.empty;
+                        val fields'' = (map o apsnd) (Envir.norm_type subst o varifyT) fields';
+                      in fields'' @ strip_fields more end
+                      handle Type.TYPE_MATCH => [("", T)]
+                        | Library.UnequalLengths => [("", T)])
+                  | NONE => [("", T)])
+              | NONE => [("", T)])
+          | NONE => [("", T)])
+      | _ => [("", T)]);
+
+    val (fields, (_, moreT)) = split_last (strip_fields T);
+    val _ = null fields andalso raise Match;
+    val u = foldr1 field_types_tr' (map (field_type_tr' o apsnd term_of_type) fields);
+  in
+    if not (! print_record_type_as_fields) orelse null fields then raise Match
+    else if moreT = HOLogic.unitT then Syntax.const @{syntax_const "_record_type"} $ u
+    else Syntax.const @{syntax_const "_record_type_scheme"} $ u $ term_of_type moreT
+  end;
 
 (*try to reconstruct the record name type abbreviation from
   the (nested) extension types*)
-fun record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt tm =
+fun record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt tm =
   let
     val thy = ProofContext.theory_of ctxt;
 
     (*tm is term representation of a (nested) field type. We first reconstruct the
       type from tm so that we can continue on the type level rather then the term level*)
 
+    (* FIXME !??? *)
     (*WORKAROUND:
       If a record type occurs in an error message of type inference there
       may be some internal frees donoted by ??:
@@ -929,80 +1028,35 @@
     if ! print_record_type_abbr then
       (case last_extT T of
         SOME (name, _) =>
-          if name = lastExt then
-           (let val subst = match schemeT T in
+          if name = last_ext then
+            let val subst = match schemeT T in
               if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree (zeta, Sign.defaultS thy))))
               then mk_type_abbr subst abbr alphas
               else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta])
-            end handle Type.TYPE_MATCH => default_tr' ctxt tm)
+            end handle Type.TYPE_MATCH => record_type_tr' ctxt tm
           else raise Match (*give print translation of specialised record a chance*)
       | _ => raise Match)
-    else default_tr' ctxt tm
+    else record_type_tr' ctxt tm
   end;
 
-fun record_type_tr' sep mark record record_scheme ctxt t =
+in
+
+fun record_ext_type_tr' name =
   let
-    val thy = ProofContext.theory_of ctxt;
-
-    val T = decode_type thy t;
-    val varifyT = varifyT (Term.maxidx_of_typ T);
-
-    fun term_of_type T = Syntax.term_of_typ (! Syntax.show_sorts) (Sign.extern_typ thy T);
-
-    fun field_lst T =
-      (case T of
-        Type (ext, args) =>
-          (case try (unsuffix ext_typeN) ext of
-            SOME ext' =>
-              (case get_extfields thy ext' of
-                SOME flds =>
-                  (case get_fieldext thy (fst (hd flds)) of
-                    SOME (_, alphas) =>
-                     (let
-                        val f :: fs = but_last flds;
-                        val flds' = apfst (Sign.extern_const thy) f ::
-                          map (apfst Long_Name.base_name) fs;
-                        val (args', more) = split_last args;
-                        val alphavars = map varifyT (but_last alphas);
-                        val subst = fold2 (curry (Sign.typ_match thy)) alphavars args' Vartab.empty;
-                        val flds'' = (map o apsnd) (Envir.norm_type subst o varifyT) flds';
-                      in flds'' @ field_lst more end
-                      handle Type.TYPE_MATCH => [("", T)] | Library.UnequalLengths => [("", T)])
-                  | NONE => [("", T)])
-              | NONE => [("", T)])
-          | NONE => [("", T)])
-      | _ => [("", T)]);
-
-    val (flds, (_, moreT)) = split_last (field_lst T);
-    val flds' = map (fn (n, T) => Syntax.const mark $ Syntax.const n $ term_of_type T) flds;
-    val flds'' =
-      foldr1 (fn (x, y) => Syntax.const sep $ x $ y) flds'
-        handle Empty => raise Match;
-  in
-    if not (! print_record_type_as_fields) orelse null flds then raise Match
-    else if moreT = HOLogic.unitT then Syntax.const record $ flds''
-    else Syntax.const record_scheme $ flds'' $ term_of_type moreT
-  end;
-
-
-fun gen_record_type_tr' name =
+    val ext_type_name = suffix ext_typeN name;
+    fun tr' ctxt ts =
+      record_type_tr' ctxt (list_comb (Syntax.const ext_type_name, ts));
+  in (ext_type_name, tr') end;
+
+fun record_ext_type_abbr_tr' abbr alphas zeta last_ext schemeT name =
   let
-    val name_sfx = suffix ext_typeN name;
+    val ext_type_name = suffix ext_typeN name;
     fun tr' ctxt ts =
-      record_type_tr' "_field_types" "_field_type" "_record_type" "_record_type_scheme"
-        ctxt (list_comb (Syntax.const name_sfx, ts))
-  in (name_sfx, tr') end;
-
-
-fun gen_record_type_abbr_tr' abbr alphas zeta lastExt schemeT name =
-  let
-    val name_sfx = suffix ext_typeN name;
-    val default_tr' =
-      record_type_tr' "_field_types" "_field_type" "_record_type" "_record_type_scheme";
-    fun tr' ctxt ts =
-      record_type_abbr_tr' default_tr' abbr alphas zeta lastExt schemeT ctxt
-        (list_comb (Syntax.const name_sfx, ts));
-  in (name_sfx, tr') end;
+      record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt
+        (list_comb (Syntax.const ext_type_name, ts));
+  in (ext_type_name, tr') end;
+
+end;
 
 
 
@@ -1040,11 +1094,11 @@
     val B = range_type X;
     val C = range_type (fastype_of f);
     val T = (B --> C) --> (A --> B) --> A --> C;
-  in Const ("Fun.comp", T) $ f $ g end;
+  in Const (@{const_name Fun.comp}, T) $ f $ g end;
 
 fun mk_comp_id f =
   let val T = range_type (fastype_of f)
-  in mk_comp (Const ("Fun.id", T --> T)) f end;
+  in mk_comp (Const (@{const_name Fun.id}, T --> T)) f end;
 
 fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t
   | get_upd_funs _ = [];
@@ -1055,6 +1109,7 @@
     val upd_funs = sort_distinct TermOrd.fast_term_ord (get_upd_funs body);
     fun get_simp upd =
       let
+        (* FIXME fresh "f" (!?) *)
         val T = domain_type (fastype_of upd);
         val lhs = mk_comp acc (upd $ Free ("f", T));
         val rhs =
@@ -1077,6 +1132,7 @@
 
 fun get_updupd_simp thy defset u u' comp =
   let
+    (* FIXME fresh "f" (!?) *)
     val f = Free ("f", domain_type (fastype_of u));
     val f' = Free ("f'", domain_type (fastype_of u'));
     val lhs = mk_comp (u $ f) (u' $ f');
@@ -1171,7 +1227,7 @@
             ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) =>
           if is_selector thy s andalso is_some (get_updates thy u) then
             let
-              val {sel_upd = {updates, ...}, extfields, ...} = RecordsData.get thy;
+              val {sel_upd = {updates, ...}, extfields, ...} = Records_Data.get thy;
 
               fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) =
                     (case Symtab.lookup updates u of
@@ -1306,7 +1362,8 @@
                   K_skeleton (Long_Name.base_name s) (domain_type T) (Bound (length vars)) f;
                 val (isnoop, skelf') = is_upd_noop s f term;
                 val funT = domain_type T;
-                fun mk_comp_local (f, f') = Const ("Fun.comp", funT --> funT --> funT) $ f $ f';
+                fun mk_comp_local (f, f') =
+                  Const (@{const_name Fun.comp}, funT --> funT --> funT) $ f $ f';
               in
                 if isnoop then
                   (upd $ skelf' $ lhs, rhs, vars,
@@ -1359,7 +1416,7 @@
 val record_eq_simproc =
   Simplifier.simproc @{theory HOL} "record_eq_simp" ["r = s"]
     (fn thy => fn _ => fn t =>
-      (case t of Const ("op =", Type (_, [T, _])) $ _ $ _ =>
+      (case t of Const (@{const_name "op ="}, Type (_, [T, _])) $ _ $ _ =>
         (case rec_id ~1 T of
           "" => NONE
         | name =>
@@ -1381,7 +1438,10 @@
     (fn thy => fn _ => fn t =>
       (case t of
         Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ =>
-          if quantifier = "All" orelse quantifier = "all" orelse quantifier = "Ex" then
+          if quantifier = @{const_name all} orelse
+            quantifier = @{const_name All} orelse
+            quantifier = @{const_name Ex}
+          then
             (case rec_id ~1 T of
               "" => NONE
             | _ =>
@@ -1392,9 +1452,9 @@
                     | SOME (all_thm, All_thm, Ex_thm, _) =>
                         SOME
                           (case quantifier of
-                            "all" => all_thm
-                          | "All" => All_thm RS eq_reflection
-                          | "Ex" => Ex_thm RS eq_reflection
+                            @{const_name all} => all_thm
+                          | @{const_name All} => All_thm RS eq_reflection
+                          | @{const_name Ex} => Ex_thm RS eq_reflection
                           | _ => error "record_split_simproc"))
                   else NONE
                 end)
@@ -1419,22 +1479,23 @@
                 else raise TERM ("", [x]);
               val sel' = Const (sel, Tsel) $ Bound 0;
               val (l, r) = if lr then (sel', x') else (x', sel');
-            in Const ("op =", Teq) $ l $ r end
+            in Const (@{const_name "op ="}, Teq) $ l $ r end
           else raise TERM ("", [Const (sel, Tsel)]);
 
-        fun dest_sel_eq (Const ("op =", Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) =
+        fun dest_sel_eq (Const (@{const_name "op ="}, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) =
               (true, Teq, (sel, Tsel), X)
-          | dest_sel_eq (Const ("op =", Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) =
+          | dest_sel_eq (Const (@{const_name "op ="}, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) =
               (false, Teq, (sel, Tsel), X)
           | dest_sel_eq _ = raise TERM ("", []);
       in
         (case t of
-          Const ("Ex", Tex) $ Abs (s, T, t) =>
+          Const (@{const_name Ex}, Tex) $ Abs (s, T, t) =>
            (let
-              val eq = mkeq (dest_sel_eq t) 0;
-              val prop =
-                list_all ([("r", T)],
-                  Logic.mk_equals (Const ("Ex", Tex) $ Abs (s, T, eq), HOLogic.true_const));
+             val eq = mkeq (dest_sel_eq t) 0;
+             val prop =
+               list_all ([("r", T)],
+                 Logic.mk_equals
+                  (Const (@{const_name Ex}, Tex) $ Abs (s, T, eq), HOLogic.true_const));
             in SOME (prove prop) end
             handle TERM _ => NONE)
         | _ => NONE)
@@ -1459,7 +1520,8 @@
 
     val has_rec = exists_Const
       (fn (s, Type (_, [Type (_, [T, _]), _])) =>
-          (s = "all" orelse s = "All" orelse s = "Ex") andalso is_recT T
+          (s = @{const_name all} orelse s = @{const_name All} orelse s = @{const_name Ex}) andalso
+          is_recT T
         | _ => false);
 
     fun mk_split_free_tac free induct_thm i =
@@ -1501,20 +1563,20 @@
 
 (* record_split_tac *)
 
-(*Split all records in the goal, which are quantified by ! or !!.*)
+(*Split all records in the goal, which are quantified by !! or ALL.*)
 val record_split_tac = CSUBGOAL (fn (cgoal, i) =>
   let
     val goal = term_of cgoal;
 
     val has_rec = exists_Const
       (fn (s, Type (_, [Type (_, [T, _]), _])) =>
-          (s = "all" orelse s = "All") andalso is_recT T
+          (s = @{const_name all} orelse s = @{const_name All}) andalso is_recT T
         | _ => false);
 
     fun is_all t =
       (case t of
         Const (quantifier, _) $ _ =>
-          if quantifier = "All" orelse quantifier = "all" then ~1 else 0
+          if quantifier = @{const_name all} orelse quantifier = @{const_name All} then ~1 else 0
       | _ => 0);
   in
     if has_rec goal then
@@ -1759,16 +1821,16 @@
       end;
     val induct = timeit_msg "record extension induct proof:" induct_prf;
 
-    val ([inject', induct', surjective', split_meta'], thm_thy) =
+    val ([induct', inject', surjective', split_meta'], thm_thy) =
       defs_thy
       |> PureThy.add_thms (map (Thm.no_attributes o apfst Binding.name)
-           [("ext_inject", inject),
-            ("ext_induct", induct),
+           [("ext_induct", induct),
+            ("ext_inject", inject),
             ("ext_surjective", surject),
             ("ext_split", split_meta)])
       ||> Code.add_default_eqn ext_def;
 
-  in (thm_thy, extT, induct', inject', split_meta', ext_def) end;
+  in ((extT, induct', inject', surjective', split_meta', ext_def), thm_thy) end;
 
 fun chunks [] [] = []
   | chunks [] xs = [xs]
@@ -1814,17 +1876,19 @@
 
 (* record_definition *)
 
-fun record_definition (args, bname) parent (parents: parent_info list) raw_fields thy =
+fun record_definition (args, b) parent (parents: parent_info list) raw_fields thy =
   let
     val external_names = Name_Space.external_names (Sign.naming_of thy);
 
     val alphas = map fst args;
-    val name = Sign.full_bname thy bname;
-    val full = Sign.full_bname_path thy bname;
+
+    val base_name = Binding.name_of b;   (* FIXME include qualifier etc. (!?) *)
+    val name = Sign.full_name thy b;
+    val full = Sign.full_name_path thy base_name;
     val base = Long_Name.base_name;
 
-    val (bfields, field_syntax) =
-      split_list (map (fn (x, T, mx) => ((x, T), mx)) raw_fields);
+    val bfields = map (fn (x, T, _) => (x, T)) raw_fields;
+    val field_syntax = map #3 raw_fields;
 
     val parent_fields = maps #fields parents;
     val parent_chunks = map (length o #fields) parents;
@@ -1837,14 +1901,14 @@
 
     val fields = map (apfst full) bfields;
     val names = map fst fields;
-    val extN = full bname;
+    val extN = full b;
     val types = map snd fields;
     val alphas_fields = fold Term.add_tfree_namesT types [];
     val alphas_ext = inter (op =) alphas_fields alphas;
     val len = length fields;
     val variants =
       Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants)
-        (map fst bfields);
+        (map (Binding.name_of o fst) bfields);
     val vars = ListPair.map Free (variants, types);
     val named_vars = names ~~ vars;
     val idxms = 0 upto len;
@@ -1859,8 +1923,8 @@
     val zeta = Name.variant alphas "'z";
     val moreT = TFree (zeta, HOLogic.typeS);
     val more = Free (moreN, moreT);
-    val full_moreN = full moreN;
-    val bfields_more = bfields @ [(moreN, moreT)];
+    val full_moreN = full (Binding.name moreN);
+    val bfields_more = bfields @ [(Binding.name moreN, moreT)];
     val fields_more = fields @ [(full_moreN, moreT)];
     val named_vars_more = named_vars @ [(full_moreN, more)];
     val all_vars_more = all_vars @ [more];
@@ -1869,9 +1933,9 @@
 
     (* 1st stage: extension_thy *)
 
-    val (extension_thy, extT, ext_induct, ext_inject, ext_split, ext_def) =
+    val ((extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), extension_thy) =
       thy
-      |> Sign.add_path bname
+      |> Sign.add_path base_name
       |> extension_definition extN fields alphas_ext zeta moreT more vars;
 
     val _ = timing_msg "record preparing definitions";
@@ -1915,47 +1979,46 @@
 
 
     (* prepare print translation functions *)
-
-    val field_tr's =
-      print_translation (distinct (op =) (maps external_names (full_moreN :: names)));
-
-    val adv_ext_tr's =
-      let val trnames = external_names extN
-      in map (gen_record_tr') trnames end;
-
-    val adv_record_type_abbr_tr's =
+    (* FIXME authentic syntax *)
+
+    val field_update_tr's =
+      map field_update_tr' (distinct (op =) (maps external_names (full_moreN :: names)));
+
+    val record_ext_tr's = map record_ext_tr' (external_names extN);
+
+    val record_ext_type_abbr_tr's =
       let
         val trnames = external_names (hd extension_names);
-        val lastExt = unsuffix ext_typeN (fst extension);
-      in map (gen_record_type_abbr_tr' name alphas zeta lastExt rec_schemeT0) trnames end;
-
-    val adv_record_type_tr's =
+        val last_ext = unsuffix ext_typeN (fst extension);
+      in map (record_ext_type_abbr_tr' name alphas zeta last_ext rec_schemeT0) trnames end;
+
+    val record_ext_type_tr's =
       let
+        (*avoid conflict with record_type_abbr_tr's*)
         val trnames = if parent_len > 0 then external_names extN else [];
-        (*avoid conflict with adv_record_type_abbr_tr's*)
-      in map (gen_record_type_tr') trnames end;
+      in map record_ext_type_tr' trnames end;
 
 
     (* prepare declarations *)
 
-    val sel_decls = map (mk_selC rec_schemeT0) bfields_more;
-    val upd_decls = map (mk_updC updateN rec_schemeT0) bfields_more;
+    val sel_decls = map (mk_selC rec_schemeT0 o apfst Binding.name_of) bfields_more;
+    val upd_decls = map (mk_updC updateN rec_schemeT0 o apfst Binding.name_of) bfields_more;
     val make_decl = (makeN, all_types ---> recT0);
     val fields_decl = (fields_selN, types ---> Type extension);
     val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0);
     val truncate_decl = (truncateN, rec_schemeT0 --> recT0);
 
+
     (* prepare definitions *)
 
     (*record (scheme) type abbreviation*)
     val recordT_specs =
-      [(Binding.name (suffix schemeN bname), alphas @ [zeta], rec_schemeT0, Syntax.NoSyn),
-        (Binding.name bname, alphas, recT0, Syntax.NoSyn)];
-
-    val ext_defs = ext_def :: map #extdef parents;
+      [(Binding.suffix_name schemeN b, alphas @ [zeta], rec_schemeT0, NoSyn),
+        (b, alphas, recT0, NoSyn)];
+
+    val ext_defs = ext_def :: map #ext_def parents;
 
     (*Theorems from the iso_tuple intros.
-      This is complex enough to deserve a full comment.
       By unfolding ext_defs from r_rec0 we create a tree of constructor
       calls (many of them Pair, but others as well). The introduction
       rules for update_accessor_eq_assist can unify two different ways
@@ -2007,45 +2070,49 @@
     (*updates*)
     fun mk_upd_spec ((c, T), thm) =
       let
-        val (upd $ _ $ arg) =
-          (fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Envir.beta_eta_contract o concl_of) thm;
+        val upd $ _ $ arg =
+          fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (concl_of thm))));
         val _ =
-          if (arg aconv r_rec0) then ()
+          if arg aconv r_rec0 then ()
           else raise TERM ("mk_sel_spec: different arg", [arg]);
       in Const (mk_updC updateN rec_schemeT0 (c, T)) :== upd end;
     val upd_specs = map mk_upd_spec (fields_more ~~ lastN updator_thms);
 
     (*derived operations*)
-    val make_spec = Const (full makeN, all_types ---> recT0) $$ all_vars :==
-      mk_rec (all_vars @ [HOLogic.unit]) 0;
-    val fields_spec = Const (full fields_selN, types ---> Type extension) $$ vars :==
-      mk_rec (all_vars @ [HOLogic.unit]) parent_len;
+    val make_spec =
+      list_comb (Const (full (Binding.name makeN), all_types ---> recT0), all_vars) :==
+        mk_rec (all_vars @ [HOLogic.unit]) 0;
+    val fields_spec =
+      list_comb (Const (full (Binding.name fields_selN), types ---> Type extension), vars) :==
+        mk_rec (all_vars @ [HOLogic.unit]) parent_len;
     val extend_spec =
-      Const (full extendN, recT0-->moreT-->rec_schemeT0) $ r_unit0 $ more :==
-      mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0;
-    val truncate_spec = Const (full truncateN, rec_schemeT0 --> recT0) $ r0 :==
-      mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0;
+      Const (full (Binding.name extendN), recT0 --> moreT --> rec_schemeT0) $ r_unit0 $ more :==
+        mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0;
+    val truncate_spec =
+      Const (full (Binding.name truncateN), rec_schemeT0 --> recT0) $ r0 :==
+        mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0;
 
 
     (* 2st stage: defs_thy *)
 
     fun mk_defs () =
       extension_thy
-      |> Sign.add_trfuns ([], [], field_tr's, [])
+      |> Sign.add_trfuns ([], [], field_update_tr's, [])
       |> Sign.add_advanced_trfuns
-          ([], [], adv_ext_tr's @ adv_record_type_tr's @ adv_record_type_abbr_tr's, [])
+        ([], [], record_ext_tr's @ record_ext_type_tr's @ record_ext_type_abbr_tr's, [])
       |> Sign.parent_path
       |> Sign.add_tyabbrs_i recordT_specs
-      |> Sign.add_path bname
+      |> Sign.add_path base_name
       |> Sign.add_consts_i
-          (map2 (fn (x, T) => fn mx => (Binding.name x, T, mx))
-            sel_decls (field_syntax @ [Syntax.NoSyn]))
-      |> (Sign.add_consts_i o map (fn (x, T) => (Binding.name x, T, Syntax.NoSyn)))
+          (map2 (fn (x, T) => fn mx => (Binding.name x, T, mx)) sel_decls (field_syntax @ [NoSyn]))
+      |> (Sign.add_consts_i o map (fn (x, T) => (Binding.name x, T, NoSyn)))
           (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl])
-      |> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) sel_specs)
-      ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name)) upd_specs)
-      ||>> ((PureThy.add_defs false o map (Thm.no_attributes o apfst Binding.name))
-             [make_spec, fields_spec, extend_spec, truncate_spec])
+      |> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
+        sel_specs
+      ||>> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
+        upd_specs
+      ||>> (PureThy.add_defs false o map (Thm.no_attributes o apfst (Binding.conceal o Binding.name)))
+        [make_spec, fields_spec, extend_spec, truncate_spec]
       |->
         (fn defs as ((sel_defs, upd_defs), derived_defs) =>
           fold Code.add_default_eqn sel_defs
@@ -2156,13 +2223,13 @@
     val (fold_congs, unfold_congs) =
       timeit_msg "record upd fold/unfold congs:" get_upd_acc_congs;
 
-    val parent_induct = if null parents then [] else [#induct (hd (rev parents))];
+    val parent_induct = Option.map #induct_scheme (try List.last parents);
 
     fun induct_scheme_prf () =
       prove_standard [] induct_scheme_prop
         (fn _ =>
           EVERY
-           [if null parent_induct then all_tac else try_param_tac rN (hd parent_induct) 1,
+           [case parent_induct of NONE => all_tac | SOME ind => try_param_tac rN ind 1,
             try_param_tac rN ext_induct 1,
             asm_simp_tac HOL_basic_ss 1]);
     val induct_scheme = timeit_msg "record induct_scheme proof:" induct_scheme_prf;
@@ -2285,7 +2352,7 @@
 
     val ((([sel_convs', upd_convs', sel_defs', upd_defs',
             fold_congs', unfold_congs',
-          [split_meta', split_object', split_ex'], derived_defs'],
+          splits' as [split_meta', split_object', split_ex'], derived_defs'],
           [surjective', equality']),
           [induct_scheme', induct', cases_scheme', cases']), thms_thy) =
       defs_thy
@@ -2311,12 +2378,22 @@
     val sel_upd_defs = sel_defs' @ upd_defs';
     val iffs = [ext_inject]
     val depth = parent_len + 1;
-    val final_thy =
+
+    val ([simps', iffs'], thms_thy') =
       thms_thy
-      |> (snd oo PureThy.add_thmss)
+      |> PureThy.add_thmss
           [((Binding.name "simps", sel_upd_simps), [Simplifier.simp_add]),
-           ((Binding.name "iffs", iffs), [iff_add])]
-      |> put_record name (make_record_info args parent fields extension induct_scheme' ext_def)
+           ((Binding.name "iffs", iffs), [iff_add])];
+
+    val info =
+      make_record_info args parent fields extension
+        ext_induct ext_inject ext_surjective ext_split ext_def
+        sel_convs' upd_convs' sel_defs' upd_defs' fold_congs' unfold_congs' splits' derived_defs'
+        surjective' equality' induct_scheme' induct' cases_scheme' cases' simps' iffs';
+
+    val final_thy =
+      thms_thy'
+      |> put_record name info
       |> put_sel_upd names full_moreN depth sel_upd_simps sel_upd_defs (fold_congs', unfold_congs')
       |> add_record_equalities extension_id equality'
       |> add_extinjects ext_inject
@@ -2333,10 +2410,13 @@
 
 (*We do all preparations and error checks here, deferring the real
   work to record_definition.*)
-fun gen_add_record prep_typ prep_raw_parent quiet_mode (params, bname) raw_parent raw_fields thy =
+fun gen_add_record prep_typ prep_raw_parent quiet_mode
+    (params, b) raw_parent raw_fields thy =
   let
     val _ = Theory.requires thy "Record" "record definitions";
-    val _ = if quiet_mode then () else writeln ("Defining record " ^ quote bname ^ " ...");
+    val _ =
+      if quiet_mode then ()
+      else writeln ("Defining record " ^ quote (Binding.name_of b) ^ " ...");
 
     val ctxt = ProofContext.init thy;
 
@@ -2357,10 +2437,12 @@
 
     (* fields *)
 
-    fun prep_field (c, raw_T, mx) env =
-      let val (T, env') = prep_typ ctxt raw_T env handle ERROR msg =>
-        cat_error msg ("The error(s) above occured in record field " ^ quote c)
-      in ((c, T, mx), env') end;
+    fun prep_field (x, raw_T, mx) env =
+      let
+        val (T, env') =
+          prep_typ ctxt raw_T env handle ERROR msg =>
+            cat_error msg ("The error(s) above occured in record field " ^ quote (Binding.str_of x));
+      in ((x, T, mx), env') end;
 
     val (bfields, envir) = fold_map prep_field raw_fields init_env;
     val envir_names = map fst envir;
@@ -2374,7 +2456,7 @@
 
     (* errors *)
 
-    val name = Sign.full_bname thy bname;
+    val name = Sign.full_name thy b;
     val err_dup_record =
       if is_none (get_record thy name) then []
       else ["Duplicate definition of record " ^ quote name];
@@ -2392,12 +2474,12 @@
     val err_no_fields = if null bfields then ["No fields present"] else [];
 
     val err_dup_fields =
-      (case duplicates (op =) (map #1 bfields) of
+      (case duplicates Binding.eq_name (map #1 bfields) of
         [] => []
-      | dups => ["Duplicate field(s) " ^ commas_quote dups]);
+      | dups => ["Duplicate field(s) " ^ commas_quote (map Binding.str_of dups)]);
 
     val err_bad_fields =
-      if forall (not_equal moreN o #1) bfields then []
+      if forall (not_equal moreN o Binding.name_of o #1) bfields then []
       else ["Illegal field name " ^ quote moreN];
 
     val err_dup_sorts =
@@ -2411,19 +2493,19 @@
 
     val _ = if null errs then () else error (cat_lines errs);
   in
-    thy |> record_definition (args, bname) parent parents bfields
+    thy |> record_definition (args, b) parent parents bfields
   end
-  handle ERROR msg => cat_error msg ("Failed to define record " ^ quote bname);
-
-val add_record = gen_add_record read_typ read_raw_parent;
-val add_record_i = gen_add_record cert_typ (K I);
+  handle ERROR msg => cat_error msg ("Failed to define record " ^ quote (Binding.str_of b));
+
+val add_record = gen_add_record cert_typ (K I);
+val add_record_cmd = gen_add_record read_typ read_raw_parent;
 
 
 (* setup theory *)
 
 val setup =
   Sign.add_trfuns ([], parse_translation, [], []) #>
-  Sign.add_advanced_trfuns ([], adv_parse_translation, [], []) #>
+  Sign.add_advanced_trfuns ([], advanced_parse_translation, [], []) #>
   Simplifier.map_simpset (fn ss =>
     ss addsimprocs [record_simproc, record_upd_simproc, record_eq_simproc]);
 
@@ -2432,13 +2514,11 @@
 
 local structure P = OuterParse and K = OuterKeyword in
 
-val record_decl =
-  P.type_args -- P.name --
-    (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const);
-
 val _ =
   OuterSyntax.command "record" "define extensible record" K.thy_decl
-    (record_decl >> (fn (x, (y, z)) => Toplevel.theory (add_record false x y z)));
+    (P.type_args -- P.binding --
+      (P.$$$ "=" |-- Scan.option (P.typ --| P.$$$ "+") -- Scan.repeat1 P.const_binding)
+    >> (fn (x, (y, z)) => Toplevel.theory (add_record_cmd false x y z)));
 
 end;
 
--- a/src/HOL/Tools/typedef.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOL/Tools/typedef.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -19,6 +19,7 @@
   val typedef_cmd: (bool * binding) * (binding * string list * mixfix) * string
     * (binding * binding) option -> theory -> Proof.state
   val get_info: theory -> string -> info option
+  val the_info: theory -> string -> info
   val interpretation: (string -> theory -> theory) -> theory -> theory
   val setup: theory -> theory
 end;
@@ -45,6 +46,12 @@
 );
 
 val get_info = Symtab.lookup o TypedefData.get;
+
+fun the_info thy name =
+  (case get_info thy name of
+    SOME info => info
+  | NONE => error ("Unknown typedef " ^ quote name));
+
 fun put_info name info = TypedefData.map (Symtab.update (name, info));
 
 
@@ -55,7 +62,7 @@
 structure Typedef_Interpretation = Interpretation(type T = string val eq = op =);
 val interpretation = Typedef_Interpretation.interpretation;
 
-fun prepare_typedef prep_term def name (t, vs, mx) raw_set opt_morphs thy =
+fun prepare_typedef prep_term def name (tname, vs, mx) raw_set opt_morphs thy =
   let
     val _ = Theory.requires thy "Typedef" "typedefs";
     val ctxt = ProofContext.init thy;
@@ -79,7 +86,6 @@
       |> filter (member (op =) rhs_tfrees andf (not o member (op =) rhs_tfreesT))
       |> map TFree;
 
-    val tname = Binding.map_name (Syntax.type_name mx) t;
     val full_tname = full tname;
     val newT = Type (full_tname, map TFree lhs_tfrees);
 
@@ -125,7 +131,7 @@
           in Drule.export_without_context (Drule.equal_elim_rule2 OF [goal_eq, th]) end;
 
     fun typedef_result inhabited =
-      ObjectLogic.typedecl (t, vs, mx)
+      ObjectLogic.typedecl (tname, vs, mx)
       #> snd
       #> Sign.add_consts_i
         [(Rep_name, newT --> oldT, NoSyn),
@@ -252,8 +258,7 @@
     Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
 
 fun mk_typedef ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
-  typedef_cmd ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name),
-    (t, vs, mx), A, morphs);
+  typedef_cmd ((def, the_default t opt_name), (t, vs, mx), A, morphs);
 
 val _ =
   OuterSyntax.command "typedef" "HOL type definition (requires non-emptiness proof)"
--- a/src/HOLCF/Tools/Domain/domain_extender.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_extender.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -160,7 +160,7 @@
       | typid (TFree (id,_)   ) = hd (strip (tl (Symbol.explode id)))
       | typid (TVar ((id,_),_)) = hd (tl (Symbol.explode id));
     fun one_con (con,args,mx) =
-        ((Syntax.const_name mx (Binding.name_of con)),
+        (Binding.name_of con,  (* FIXME preverse binding (!?) *)
          ListPair.map (fn ((lazy,sel,tp),vn) =>
            mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp),
                    Option.map Binding.name_of sel,vn))
@@ -235,7 +235,7 @@
       | typid (TFree (id,_)   ) = hd (strip (tl (Symbol.explode id)))
       | typid (TVar ((id,_),_)) = hd (tl (Symbol.explode id));
     fun one_con (con,args,mx) =
-        ((Syntax.const_name mx (Binding.name_of con)),
+        (Binding.name_of con,   (* FIXME preverse binding (!?) *)
          ListPair.map (fn ((lazy,sel,tp),vn) =>
            mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp),
                    Option.map Binding.name_of sel,vn))
--- a/src/HOLCF/Tools/Domain/domain_syntax.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOLCF/Tools/Domain/domain_syntax.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -28,7 +28,7 @@
 open Domain_Library;
 infixr 5 -->; infixr 6 ->>;
 
-fun calc_syntax
+fun calc_syntax  (* FIXME authentic syntax *)
     (definitional : bool)
     (dtypeprod : typ)
     ((dname : string, typevars : typ list), 
@@ -115,7 +115,7 @@
 
     local open Syntax in
     local
-      fun c_ast con mx = Constant (Syntax.const_name mx (Binding.name_of con));
+      fun c_ast con mx = Constant (Binding.name_of con);   (* FIXME proper const syntax *)
       fun expvar n     = Variable ("e"^(string_of_int n));
       fun argvar n m _ = Variable ("a"^(string_of_int n)^"_"^
                                    (string_of_int m));
--- a/src/HOLCF/Tools/cont_consts.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOLCF/Tools/cont_consts.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -40,9 +40,9 @@
         fold (fn arg => fn t => Syntax.mk_appl (Constant "Rep_CFun") [t, Variable arg])
           vnames (Constant name1))] @
     (case mx of
-      InfixName _ => [extra_parse_rule]
-    | InfixlName _ => [extra_parse_rule]
-    | InfixrName _ => [extra_parse_rule]
+      Infix _ => [extra_parse_rule]
+    | Infixl _ => [extra_parse_rule]
+    | Infixr _ => [extra_parse_rule]
     | _ => [])
   end;
 
@@ -53,19 +53,8 @@
    declaration with the original name, type ...=>..., and the original mixfix
    is generated and connected to the other declaration via some translation.
 *)
-fun const_binding mx = Binding.name o Syntax.const_name mx o Binding.name_of;
-
-fun fix_mixfix (syn                 , T, mx as Infix           p ) =
-               (const_binding mx syn, T,       InfixName (Binding.name_of syn, p))
-  | fix_mixfix (syn                 , T, mx as Infixl           p ) =
-               (const_binding mx syn, T,       InfixlName (Binding.name_of syn, p))
-  | fix_mixfix (syn                 , T, mx as Infixr           p ) =
-               (const_binding mx syn, T,       InfixrName (Binding.name_of syn, p))
-  | fix_mixfix decl = decl;
-
-fun transform decl =
+fun transform (c, T, mx) =
     let
-        val (c, T, mx) = fix_mixfix decl;
         val c1 = Binding.name_of c;
         val c2 = "_cont_" ^ c1;
         val n  = Syntax.mixfix_args mx
@@ -78,9 +67,9 @@
 
 fun is_contconst (_,_,NoSyn   ) = false
 |   is_contconst (_,_,Binder _) = false
-|   is_contconst (c,T,mx      ) = cfun_arity T >= Syntax.mixfix_args mx
-                         handle ERROR msg => cat_error msg ("in mixfix annotation for " ^
-                                               quote (Syntax.const_name mx (Binding.name_of c)));
+|   is_contconst (c,T,mx      ) =
+      cfun_arity T >= Syntax.mixfix_args mx
+        handle ERROR msg => cat_error msg ("in mixfix annotation for " ^ quote (Binding.str_of c));
 
 
 (* add_consts(_i) *)
@@ -94,6 +83,7 @@
     thy
     |> Sign.add_consts_i
       (normal_decls @ map first transformed_decls @ map second transformed_decls)
+    (* FIXME authentic syntax *)
     |> Sign.add_trrules_i (maps third transformed_decls)
   end;
 
--- a/src/HOLCF/Tools/pcpodef.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOLCF/Tools/pcpodef.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -153,7 +153,7 @@
 fun declare_type_name a =
   Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS)));
 
-fun prepare prep_term name (t, vs, mx) raw_set opt_morphs thy =
+fun prepare prep_term name (tname, vs, mx) raw_set opt_morphs thy =
   let
     val _ = Theory.requires thy "Pcpodef" "pcpodefs";
     val ctxt = ProofContext.init thy;
@@ -168,7 +168,6 @@
     (*lhs*)
     val defS = Sign.defaultS thy;
     val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
-    val tname = Binding.map_name (Syntax.type_name mx) t;
     val full_tname = Sign.full_name thy tname;
     val newT = Type (full_tname, map TFree lhs_tfrees);
 
@@ -346,7 +345,7 @@
 
 fun mk_pcpodef_proof pcpo ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
   (if pcpo then pcpodef_proof_cmd else cpodef_proof_cmd)
-    ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name), (t, vs, mx), A, morphs);
+    ((def, the_default t opt_name), (t, vs, mx), A, morphs);
 
 val _ =
   OuterSyntax.command "pcpodef" "HOLCF type definition (requires admissibility proof)" K.thy_goal
--- a/src/HOLCF/Tools/repdef.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/HOLCF/Tools/repdef.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -59,7 +59,7 @@
       (prep_term: Proof.context -> 'a -> term)
       (def: bool)
       (name: binding)
-      (typ as (t, vs, mx) : binding * string list * mixfix)
+      (typ as (tname, vs, mx) : binding * string list * mixfix)
       (raw_defl: 'a)
       (opt_morphs: (binding * binding) option)
       (thy: theory)
@@ -79,7 +79,6 @@
     val defS = Sign.defaultS thy;
     val lhs_tfrees = map (fn v => (v, the_default defS (AList.lookup (op =) rhs_tfrees v))) vs;
     val lhs_sorts = map snd lhs_tfrees;
-    val tname = Binding.map_name (Syntax.type_name mx) t;
     val full_tname = Sign.full_name thy tname;
     val newT = Type (full_tname, map TFree lhs_tfrees);
 
@@ -172,8 +171,7 @@
     Scan.option (P.$$$ "morphisms" |-- P.!!! (P.binding -- P.binding));
 
 fun mk_repdef ((((((def, opt_name), (vs, t)), mx), A), morphs)) =
-  repdef_cmd
-    ((def, the_default (Binding.map_name (Syntax.type_name mx) t) opt_name), (t, vs, mx), A, morphs);
+  repdef_cmd ((def, the_default t opt_name), (t, vs, mx), A, morphs);
 
 val _ =
   OuterSyntax.command "repdef" "HOLCF definition of representable domains" K.thy_decl
--- a/src/LCF/LCF.thy	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/LCF/LCF.thy	Wed Feb 17 09:51:46 2010 +0100
@@ -19,8 +19,8 @@
 
 typedecl tr
 typedecl void
-typedecl ('a,'b) "*"    (infixl 6)
-typedecl ('a,'b) "+"    (infixl 5)
+typedecl ('a,'b) "*"    (infixl "*" 6)
+typedecl ('a,'b) "+"    (infixl "+" 5)
 
 arities
   "fun" :: (cpo, cpo) cpo
--- a/src/Pure/Isar/expression.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Isar/expression.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -606,7 +606,7 @@
 
 fun aprop_tr' n c = (Syntax.constN ^ c, fn ctxt => fn args =>
   if length args = n then
-    Syntax.const "_aprop" $
+    Syntax.const "_aprop" $   (* FIXME avoid old-style early externing (!??) *)
       Term.list_comb (Syntax.free (Consts.extern (ProofContext.consts_of ctxt) c), args)
   else raise Match);
 
--- a/src/Pure/Isar/isar_cmd.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Isar/isar_cmd.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -332,16 +332,14 @@
 val print_abbrevs = Toplevel.unknown_context o
   Toplevel.keep (ProofContext.print_abbrevs o Toplevel.context_of);
 
-val print_facts = Toplevel.unknown_context o Toplevel.keep (fn state =>
-  ProofContext.setmp_verbose_CRITICAL
-    ProofContext.print_lthms (Toplevel.context_of state));
+val print_facts = Toplevel.unknown_context o
+  Toplevel.keep (ProofContext.print_lthms o Toplevel.context_of);
 
-val print_configs = Toplevel.unknown_context o Toplevel.keep (fn state =>
-  Attrib.print_configs (Toplevel.context_of state));
+val print_configs = Toplevel.unknown_context o
+  Toplevel.keep (Attrib.print_configs o Toplevel.context_of);
 
-val print_theorems_proof = Toplevel.keep (fn state =>
-  ProofContext.setmp_verbose_CRITICAL
-    ProofContext.print_lthms (Proof.context_of (Toplevel.proof_of state)));
+val print_theorems_proof =
+  Toplevel.keep (ProofContext.print_lthms o Proof.context_of o Toplevel.proof_of);
 
 fun print_theorems_theory verbose = Toplevel.keep (fn state =>
   Toplevel.theory_of state |>
@@ -430,11 +428,11 @@
 
 (* print proof context contents *)
 
-val print_binds = Toplevel.unknown_context o Toplevel.keep (fn state =>
-  ProofContext.setmp_verbose_CRITICAL ProofContext.print_binds (Toplevel.context_of state));
+val print_binds = Toplevel.unknown_context o
+  Toplevel.keep (ProofContext.print_binds o Toplevel.context_of);
 
-val print_cases = Toplevel.unknown_context o Toplevel.keep (fn state =>
-  ProofContext.setmp_verbose_CRITICAL ProofContext.print_cases (Toplevel.context_of state));
+val print_cases = Toplevel.unknown_context o
+  Toplevel.keep (ProofContext.print_cases o Toplevel.context_of);
 
 
 (* print theorems, terms, types etc. *)
--- a/src/Pure/Isar/object_logic.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Isar/object_logic.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -84,10 +84,9 @@
 
 (* typedecl *)
 
-fun typedecl (a, vs, mx) thy =
+fun typedecl (b, vs, mx) thy =
   let
     val base_sort = get_base_sort thy;
-    val b = Binding.map_name (Syntax.type_name mx) a;
     val _ = has_duplicates (op =) vs andalso
       error ("Duplicate parameters in type declaration " ^ quote (Binding.str_of b));
     val name = Sign.full_name thy b;
@@ -95,7 +94,7 @@
     val T = Type (name, map (fn v => TFree (v, [])) vs);
   in
     thy
-    |> Sign.add_types [(a, n, mx)]
+    |> Sign.add_types [(b, n, mx)]
     |> (case base_sort of NONE => I | SOME S => AxClass.axiomatize_arity (name, replicate n S, S))
     |> pair T
   end;
@@ -106,7 +105,7 @@
 local
 
 fun gen_add_judgment add_consts (b, T, mx) thy =
-  let val c = Sign.full_name thy (Binding.map_name (Syntax.const_name mx) b) in
+  let val c = Sign.full_name thy b in
     thy
     |> add_consts [(b, T, mx)]
     |> (fn thy' => Theory.add_deps c (c, Sign.the_const_type thy' c) [] thy')
--- a/src/Pure/Isar/outer_parse.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Isar/outer_parse.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -266,9 +266,9 @@
   !!! (Scan.optional ($$$ "[" |-- !!! (list nat --| $$$ "]")) [] --
     Scan.optional nat Syntax.max_pri) >> (Mixfix o triple2);
 
-val infx = $$$ "infix" |-- !!! (nat >> Infix || string -- nat >> InfixName);
-val infxl = $$$ "infixl" |-- !!! (nat >> Infixl || string -- nat >> InfixlName);
-val infxr = $$$ "infixr" |-- !!! (nat >> Infixr || string -- nat >> InfixrName);
+val infx = $$$ "infix" |-- !!! (string -- nat >> Infix);
+val infxl = $$$ "infixl" |-- !!! (string -- nat >> Infixl);
+val infxr = $$$ "infixr" |-- !!! (string -- nat >> Infixr);
 
 val binder = $$$ "binder" |--
   !!! (string -- ($$$ "[" |-- nat --| $$$ "]" -- nat || nat >> (fn n => (n, n))))
--- a/src/Pure/Isar/proof_context.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Isar/proof_context.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -30,7 +30,6 @@
   val consts_of: Proof.context -> Consts.T
   val const_syntax_name: Proof.context -> string -> string
   val the_const_constraint: Proof.context -> string -> typ
-  val mk_const: Proof.context -> string * typ list -> term
   val set_syntax_mode: Syntax.mode -> Proof.context -> Proof.context
   val restore_syntax_mode: Proof.context -> Proof.context -> Proof.context
   val facts_of: Proof.context -> Facts.T
@@ -123,14 +122,13 @@
   val add_const_constraint: string * typ option -> Proof.context -> Proof.context
   val add_abbrev: string -> binding * term -> Proof.context -> (term * term) * Proof.context
   val revert_abbrev: string -> string -> Proof.context -> Proof.context
-  val verbose: bool Unsynchronized.ref
-  val setmp_verbose_CRITICAL: ('a -> 'b) -> 'a -> 'b
   val print_syntax: Proof.context -> unit
   val print_abbrevs: Proof.context -> unit
   val print_binds: Proof.context -> unit
   val print_lthms: Proof.context -> unit
   val print_cases: Proof.context -> unit
   val debug: bool Unsynchronized.ref
+  val verbose: bool Unsynchronized.ref
   val prems_limit: int Unsynchronized.ref
   val pretty_ctxt: Proof.context -> Pretty.T list
   val pretty_context: Proof.context -> Pretty.T list
@@ -240,8 +238,6 @@
 val const_syntax_name = Consts.syntax_name o consts_of;
 val the_const_constraint = Consts.the_constraint o consts_of;
 
-fun mk_const ctxt (c, Ts) = Const (c, Consts.instance (consts_of ctxt) (c, Ts));
-
 val facts_of = #facts o rep_context;
 val cases_of = #cases o rep_context;
 
@@ -966,19 +962,18 @@
 local
 
 fun prep_vars prep_typ internal =
-  fold_map (fn (raw_b, raw_T, raw_mx) => fn ctxt =>
+  fold_map (fn (b, raw_T, mx) => fn ctxt =>
     let
-      val raw_x = Name.of_binding raw_b;
-      val (x, mx) = Syntax.const_mixfix raw_x raw_mx;
+      val x = Name.of_binding b;
       val _ = Syntax.is_identifier (no_skolem internal x) orelse
-        error ("Illegal variable name: " ^ quote x);
+        error ("Illegal variable name: " ^ quote (Binding.str_of b));
 
       fun cond_tvars T =
         if internal then T
         else Type.no_tvars T handle TYPE (msg, _, _) => error msg;
       val opt_T = Option.map (cond_tvars o cert_typ ctxt o prep_typ ctxt) raw_T;
-      val var = (Binding.map_name (K x) raw_b, opt_T, mx);
-    in (var, ctxt |> declare_var (x, opt_T, mx) |> #2) end);
+      val (_, ctxt') = ctxt |> declare_var (x, opt_T, mx);
+    in ((b, opt_T, mx), ctxt') end);
 
 in
 
@@ -1199,14 +1194,6 @@
 
 (** print context information **)
 
-val debug = Unsynchronized.ref false;
-
-val verbose = Unsynchronized.ref false;
-fun verb f x = if ! verbose then f (x ()) else [];
-
-fun setmp_verbose_CRITICAL f x = setmp_CRITICAL verbose true f x;
-
-
 (* local syntax *)
 
 val print_syntax = Syntax.print_syntax o syn_of;
@@ -1224,7 +1211,7 @@
           else cons (c, Logic.mk_equals (Const (c, T), t));
     val abbrevs = Name_Space.extern_table (space, Symtab.make (Symtab.fold add_abbr consts []));
   in
-    if null abbrevs andalso not (! verbose) then []
+    if null abbrevs then []
     else [Pretty.big_list "abbreviations:" (map (pretty_term_abbrev ctxt o #2) abbrevs)]
   end;
 
@@ -1238,7 +1225,7 @@
     val binds = Variable.binds_of ctxt;
     fun prt_bind (xi, (T, t)) = pretty_term_abbrev ctxt (Logic.mk_equals (Var (xi, T), t));
   in
-    if Vartab.is_empty binds andalso not (! verbose) then []
+    if Vartab.is_empty binds then []
     else [Pretty.big_list "term bindings:" (map prt_bind (Vartab.dest binds))]
   end;
 
@@ -1252,10 +1239,10 @@
     val local_facts = facts_of ctxt;
     val props = Facts.props local_facts;
     val facts =
-      (if null props then [] else [("unnamed", props)]) @
+      (if null props then [] else [("<unnamed>", props)]) @
       Facts.dest_static [] local_facts;
   in
-    if null facts andalso not (! verbose) then []
+    if null facts then []
     else [Pretty.big_list "facts:" (map #1 (sort_wrt (#1 o #2) (map (`(pretty_fact ctxt)) facts)))]
   end;
 
@@ -1278,8 +1265,9 @@
       ((if a = "" then [] else [Pretty.str (a ^ ":")]) @ map (Pretty.quote o prt_term) ts));
 
     fun prt_sect _ _ _ [] = []
-      | prt_sect s sep prt xs = [Pretty.block (Pretty.breaks (Pretty.str s ::
-            flat (Library.separate sep (map (Library.single o prt) xs))))];
+      | prt_sect s sep prt xs =
+          [Pretty.block (Pretty.breaks (Pretty.str s ::
+            flat (separate sep (map (single o prt) xs))))];
   in
     Pretty.block (Pretty.fbreaks
       (Pretty.str (name ^ ":") ::
@@ -1300,7 +1288,7 @@
           cons (name, (fixes, case_result c ctxt));
     val cases = fold add_case (cases_of ctxt) [];
   in
-    if null cases andalso not (! verbose) then []
+    if null cases then []
     else [Pretty.big_list "cases:" (map pretty_case cases)]
   end;
 
@@ -1311,6 +1299,8 @@
 
 (* core context *)
 
+val debug = Unsynchronized.ref false;
+val verbose = Unsynchronized.ref false;
 val prems_limit = Unsynchronized.ref ~1;
 
 fun pretty_ctxt ctxt =
@@ -1321,7 +1311,8 @@
 
       (*structures*)
       val structs = Local_Syntax.structs_of (syntax_of ctxt);
-      val prt_structs = if null structs then []
+      val prt_structs =
+        if null structs then []
         else [Pretty.block (Pretty.str "structures:" :: Pretty.brk 1 ::
           Pretty.commas (map Pretty.str structs))];
 
@@ -1332,7 +1323,8 @@
       val fixes =
         rev (filter_out ((can Name.dest_internal orf member (op =) structs) o #1)
           (Variable.fixes_of ctxt));
-      val prt_fixes = if null fixes then []
+      val prt_fixes =
+        if null fixes then []
         else [Pretty.block (Pretty.str "fixed variables:" :: Pretty.brk 1 ::
           Pretty.commas (map prt_fix fixes))];
 
@@ -1340,7 +1332,8 @@
       val prems = Assumption.all_prems_of ctxt;
       val len = length prems;
       val suppressed = len - ! prems_limit;
-      val prt_prems = if null prems then []
+      val prt_prems =
+        if null prems then []
         else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @
           map (Display.pretty_thm ctxt) (drop suppressed prems))];
     in prt_structs @ prt_fixes @ prt_prems end;
@@ -1350,6 +1343,9 @@
 
 fun pretty_context ctxt =
   let
+    val is_verbose = ! verbose;
+    fun verb f x = if is_verbose then f (x ()) else [];
+
     val prt_term = Syntax.pretty_term ctxt;
     val prt_typ = Syntax.pretty_typ ctxt;
     val prt_sort = Syntax.pretty_sort ctxt;
--- a/src/Pure/ML/ml_antiquote.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/ML/ml_antiquote.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -123,9 +123,11 @@
 val _ = inline "const"
   (Args.context -- Scan.lift Args.name_source -- Scan.optional
       (Scan.lift (Args.$$$ "(") |-- OuterParse.enum1' "," Args.typ --| Scan.lift (Args.$$$ ")")) []
-    >> (fn ((ctxt, c), Ts) =>
-      let val (c, _) = Term.dest_Const (ProofContext.read_const_proper ctxt c)
-      in ML_Syntax.atomic (ML_Syntax.print_term (ProofContext.mk_const ctxt (c, Ts))) end));
+    >> (fn ((ctxt, raw_c), Ts) =>
+      let
+        val Const (c, _) = ProofContext.read_const_proper ctxt raw_c;
+        val const = Const (c, Consts.instance (ProofContext.consts_of ctxt) (c, Ts));
+      in ML_Syntax.atomic (ML_Syntax.print_term const) end));
 
 val _ = inline "syntax_const"
   (Args.context -- Scan.lift Args.name >> (fn (ctxt, c) =>
--- a/src/Pure/Syntax/mixfix.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Syntax/mixfix.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -10,12 +10,9 @@
     NoSyn |
     Mixfix of string * int list * int |
     Delimfix of string |
-    InfixName of string * int |
-    InfixlName of string * int |
-    InfixrName of string * int |
-    Infix of int |    (*obsolete*)
-    Infixl of int |   (*obsolete*)
-    Infixr of int |   (*obsolete*)
+    Infix of string * int |
+    Infixl of string * int |
+    Infixr of string * int |
     Binder of string * int * int |
     Structure
   val binder_name: string -> string
@@ -27,9 +24,6 @@
   val literal: string -> mixfix
   val no_syn: 'a * 'b -> 'a * 'b * mixfix
   val pretty_mixfix: mixfix -> Pretty.T
-  val type_name: mixfix -> string -> string
-  val const_name: mixfix -> string -> string
-  val const_mixfix: string -> mixfix -> string * mixfix
   val mixfix_args: mixfix -> int
   val mixfixT: mixfix -> typ
 end;
@@ -51,12 +45,9 @@
   NoSyn |
   Mixfix of string * int list * int |
   Delimfix of string |
-  InfixName of string * int |
-  InfixlName of string * int |
-  InfixrName of string * int |
-  Infix of int |      (*obsolete*)
-  Infixl of int |     (*obsolete*)
-  Infixr of int |     (*obsolete*)
+  Infix of string * int |
+  Infixl of string * int |
+  Infixr of string * int |
   Binder of string * int * int |
   Structure;
 
@@ -81,12 +72,9 @@
   | pretty_mixfix (Mixfix (s, ps, p)) =
       parens (Pretty.breaks [quoted s, brackets (Pretty.commas (map int ps)), int p])
   | pretty_mixfix (Delimfix s) = parens [quoted s]
-  | pretty_mixfix (InfixName (s, p)) = parens (Pretty.breaks [keyword "infix", quoted s, int p])
-  | pretty_mixfix (InfixlName (s, p)) = parens (Pretty.breaks [keyword "infixl", quoted s, int p])
-  | pretty_mixfix (InfixrName (s, p)) = parens (Pretty.breaks [keyword "infixl", quoted s, int p])
-  | pretty_mixfix (Infix p) = parens (Pretty.breaks [keyword "infix", int p])
-  | pretty_mixfix (Infixl p) = parens (Pretty.breaks [keyword "infixl", int p])
-  | pretty_mixfix (Infixr p) = parens (Pretty.breaks [keyword "infixr", int p])
+  | pretty_mixfix (Infix (s, p)) = parens (Pretty.breaks [keyword "infix", quoted s, int p])
+  | pretty_mixfix (Infixl (s, p)) = parens (Pretty.breaks [keyword "infixl", quoted s, int p])
+  | pretty_mixfix (Infixr (s, p)) = parens (Pretty.breaks [keyword "infixl", quoted s, int p])
   | pretty_mixfix (Binder (s, p1, p2)) =
       parens (Pretty.breaks [keyword "binder", quoted s, brackets [int p1], int p2])
   | pretty_mixfix Structure = parens [keyword "structure"];
@@ -96,47 +84,12 @@
 
 (* syntax specifications *)
 
-fun strip ("'" :: c :: cs) = c :: strip cs
-  | strip ["'"] = []
-  | strip (c :: cs) = c :: strip cs
-  | strip [] = [];
-
-val strip_esc = implode o strip o Symbol.explode;
-
-fun deprecated c = (legacy_feature ("Unnamed infix operator " ^ quote c); c);
-
-fun type_name (InfixName _) = I
-  | type_name (InfixlName _) = I
-  | type_name (InfixrName _) = I
-  | type_name (Infix _) = deprecated o strip_esc
-  | type_name (Infixl _) = deprecated o strip_esc
-  | type_name (Infixr _) = deprecated o strip_esc
-  | type_name _ = I;
-
-fun const_name (InfixName _) = I
-  | const_name (InfixlName _) = I
-  | const_name (InfixrName _) = I
-  | const_name (Infix _) = prefix "op " o deprecated o strip_esc
-  | const_name (Infixl _) = prefix "op " o deprecated o strip_esc
-  | const_name (Infixr _) = prefix "op " o deprecated o strip_esc
-  | const_name _ = I;
-
-fun fix_mixfix c (Infix p) = InfixName (c, p)
-  | fix_mixfix c (Infixl p) = InfixlName (c, p)
-  | fix_mixfix c (Infixr p) = InfixrName (c, p)
-  | fix_mixfix _ mx = mx;
-
-fun const_mixfix c mx = (const_name mx c, fix_mixfix c mx);
-
 fun mixfix_args NoSyn = 0
   | mixfix_args (Mixfix (sy, _, _)) = SynExt.mfix_args sy
   | mixfix_args (Delimfix sy) = SynExt.mfix_args sy
-  | mixfix_args (InfixName (sy, _)) = 2 + SynExt.mfix_args sy
-  | mixfix_args (InfixlName (sy, _)) = 2 + SynExt.mfix_args sy
-  | mixfix_args (InfixrName (sy, _)) = 2 + SynExt.mfix_args sy
-  | mixfix_args (Infix _) = 2
-  | mixfix_args (Infixl _) = 2
-  | mixfix_args (Infixr _) = 2
+  | mixfix_args (Infix (sy, _)) = 2 + SynExt.mfix_args sy
+  | mixfix_args (Infixl (sy, _)) = 2 + SynExt.mfix_args sy
+  | mixfix_args (Infixr (sy, _)) = 2 + SynExt.mfix_args sy
   | mixfix_args (Binder _) = 1
   | mixfix_args Structure = 0;
 
@@ -148,27 +101,19 @@
 
 fun syn_ext_types type_decls =
   let
-    fun name_of (t, _, mx) = type_name mx t;
-
     fun mk_infix sy t p1 p2 p3 =
       SynExt.Mfix ("(_ " ^ sy ^ "/ _)",
         [SynExt.typeT, SynExt.typeT] ---> SynExt.typeT, t, [p1, p2], p3);
 
-    fun mfix_of decl =
-      let val t = name_of decl in
-        (case decl of
-          (_, _, NoSyn) => NONE
-        | (_, 2, InfixName (sy, p)) => SOME (mk_infix sy t (p + 1) (p + 1) p)
-        | (_, 2, InfixlName (sy, p)) => SOME (mk_infix sy t p (p + 1) p)
-        | (_, 2, InfixrName (sy, p)) => SOME (mk_infix sy t (p + 1) p p)
-        | (sy, 2, Infix p) => SOME (mk_infix sy t (p + 1) (p + 1) p)
-        | (sy, 2, Infixl p) => SOME (mk_infix sy t p (p + 1) p)
-        | (sy, 2, Infixr p) => SOME (mk_infix sy t (p + 1) p p)
-        | _ => error ("Bad mixfix declaration for type: " ^ quote t))
-      end;
+    fun mfix_of (_, _, NoSyn) = NONE
+      | mfix_of (t, 2, Infix (sy, p)) = SOME (mk_infix sy t (p + 1) (p + 1) p)
+      | mfix_of (t, 2, Infixl (sy, p)) = SOME (mk_infix sy t p (p + 1) p)
+      | mfix_of (t, 2, Infixr (sy, p)) = SOME (mk_infix sy t (p + 1) p p)
+      | mfix_of (t, _, _) =
+          error ("Bad mixfix declaration for type: " ^ quote t);  (* FIXME printable!? *)
 
     val mfix = map_filter mfix_of type_decls;
-    val xconsts = map name_of type_decls;
+    val xconsts = map #1 type_decls;
   in SynExt.syn_ext mfix xconsts ([], [], [], []) [] ([], []) end;
 
 
@@ -179,8 +124,6 @@
 
 fun syn_ext_consts is_logtype const_decls =
   let
-    fun name_of (c, _, mx) = const_name mx c;
-
     fun mk_infix sy ty c p1 p2 p3 =
       [SynExt.Mfix ("op " ^ sy, ty, c, [], SynExt.max_pri),
        SynExt.Mfix ("(_ " ^ sy ^ "/ _)", ty, c, [p1, p2], p3)];
@@ -189,33 +132,27 @@
           [Type ("idts", []), ty2] ---> ty3
       | binder_typ c _ = error ("Bad type of binder: " ^ quote c);
 
-    fun mfix_of decl =
-      let val c = name_of decl in
-        (case decl of
-          (_, _, NoSyn) => []
-        | (_, ty, Mixfix (sy, ps, p)) => [SynExt.Mfix (sy, ty, c, ps, p)]
-        | (_, ty, Delimfix sy) => [SynExt.Mfix (sy, ty, c, [], SynExt.max_pri)]
-        | (_, ty, InfixName (sy, p)) => mk_infix sy ty c (p + 1) (p + 1) p
-        | (_, ty, InfixlName (sy, p)) => mk_infix sy ty c p (p + 1) p
-        | (_, ty, InfixrName (sy, p)) => mk_infix sy ty c (p + 1) p p
-        | (sy, ty, Infix p) => mk_infix sy ty c (p + 1) (p + 1) p
-        | (sy, ty, Infixl p) => mk_infix sy ty c p (p + 1) p
-        | (sy, ty, Infixr p) => mk_infix sy ty c (p + 1) p p
-        | (_, ty, Binder (sy, p, q)) =>
-            [SynExt.Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, (binder_name c), [0, p], q)]
-        | _ => error ("Bad mixfix declaration for const: " ^ quote c))
-    end;
+    fun mfix_of (_, _, NoSyn) = []
+      | mfix_of (c, ty, Mixfix (sy, ps, p)) = [SynExt.Mfix (sy, ty, c, ps, p)]
+      | mfix_of (c, ty, Delimfix sy) = [SynExt.Mfix (sy, ty, c, [], SynExt.max_pri)]
+      | mfix_of (c, ty, Infix (sy, p)) = mk_infix sy ty c (p + 1) (p + 1) p
+      | mfix_of (c, ty, Infixl (sy, p)) = mk_infix sy ty c p (p + 1) p
+      | mfix_of (c, ty, Infixr (sy, p)) = mk_infix sy ty c (p + 1) p p
+      | mfix_of (c, ty, Binder (sy, p, q)) =
+          [SynExt.Mfix ("(3" ^ sy ^ "_./ _)", binder_typ c ty, (binder_name c), [0, p], q)]
+      | mfix_of (c, _, _) = error ("Bad mixfix declaration for const: " ^ quote c);
 
     fun binder (c, _, Binder _) = SOME (binder_name c, c)
       | binder _ = NONE;
 
     val mfix = maps mfix_of const_decls;
-    val xconsts = map name_of const_decls;
+    val xconsts = map #1 const_decls;
     val binders = map_filter binder const_decls;
-    val binder_trs = binders |> map (SynExt.stamp_trfun binder_stamp o
-        apsnd K o SynTrans.mk_binder_tr);
-    val binder_trs' = binders |> map (SynExt.stamp_trfun binder_stamp o
-        apsnd (K o SynTrans.non_typed_tr') o SynTrans.mk_binder_tr' o swap);
+    val binder_trs = binders
+      |> map (SynExt.stamp_trfun binder_stamp o apsnd K o SynTrans.mk_binder_tr);
+    val binder_trs' = binders
+      |> map (SynExt.stamp_trfun binder_stamp o
+          apsnd (K o SynTrans.non_typed_tr') o SynTrans.mk_binder_tr' o swap);
   in
     SynExt.syn_ext' true is_logtype
       mfix xconsts ([], binder_trs, binder_trs', []) [] ([], [])
--- a/src/Pure/Syntax/syn_trans.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Syntax/syn_trans.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -19,6 +19,7 @@
   val antiquote_tr': string -> term -> term
   val quote_tr': string -> term -> term
   val quote_antiquote_tr': string -> string -> string -> string * (term list -> term)
+  val update_name_tr': term -> term
   val mark_bound: string -> term
   val mark_boundT: string * typ -> term
   val bound_vars: (string * typ) list -> term -> term
@@ -187,6 +188,15 @@
   in (quoteN, tr) end;
 
 
+(* corresponding updates *)
+
+fun update_name_tr (Free (x, T) :: ts) = list_comb (Free (suffix "_update" x, T), ts)
+  | update_name_tr (Const (x, T) :: ts) = list_comb (Const (suffix "_update" x, T), ts)
+  | update_name_tr (((c as Const ("_constrain", _)) $ t $ ty) :: ts) =
+      list_comb (c $ update_name_tr [t] $ (Lexicon.const "fun" $ ty $ Lexicon.const "dummy"), ts)
+  | update_name_tr ts = raise TERM ("update_name_tr", ts);
+
+
 (* indexed syntax *)
 
 fun struct_ast_tr (*"_struct"*) [Ast.Appl [Ast.Constant "_index", ast]] = ast
@@ -444,6 +454,19 @@
   in (name, tr') end;
 
 
+(* corresponding updates *)
+
+fun upd_tr' (x_upd, T) =
+  (case try (unsuffix "_update") x_upd of
+    SOME x => (x, if T = dummyT then T else Term.domain_type T)
+  | NONE => raise Match);
+
+fun update_name_tr' (Free x) = Free (upd_tr' x)
+  | update_name_tr' ((c as Const ("_free", _)) $ Free x) = c $ Free (upd_tr' x)
+  | update_name_tr' (Const x) = Const (upd_tr' x)
+  | update_name_tr' _ = raise Match;
+
+
 (* indexed syntax *)
 
 fun index_ast_tr' (*"_index"*) [Ast.Appl [Ast.Constant "_struct", ast]] = ast
@@ -468,17 +491,31 @@
 (** Pure translations **)
 
 val pure_trfuns =
- ([("_constify", constify_ast_tr), ("_appl", appl_ast_tr), ("_applC", applC_ast_tr),
-   ("_lambda", lambda_ast_tr), ("_idtyp", idtyp_ast_tr), ("_idtypdummy", idtypdummy_ast_tr),
-   ("_bigimpl", bigimpl_ast_tr), ("_indexdefault", indexdefault_ast_tr),
-   ("_indexnum", indexnum_ast_tr), ("_indexvar", indexvar_ast_tr), ("_struct", struct_ast_tr)],
-  [("_abs", abs_tr), ("_aprop", aprop_tr), ("_ofclass", ofclass_tr),
-   ("_sort_constraint", sort_constraint_tr), ("_TYPE", type_tr),
-   ("_DDDOT", dddot_tr), ("_index", index_tr)],
-  ([]: (string * (term list -> term)) list),
-  [("_abs", abs_ast_tr'), ("_idts", idtyp_ast_tr' "_idts"),
-   ("_pttrns", idtyp_ast_tr' "_pttrns"), ("==>", impl_ast_tr'),
-   ("_index", index_ast_tr')]);
+  ([("_constify", constify_ast_tr),
+    ("_appl", appl_ast_tr),
+    ("_applC", applC_ast_tr),
+    ("_lambda", lambda_ast_tr),
+    ("_idtyp", idtyp_ast_tr),
+    ("_idtypdummy", idtypdummy_ast_tr),
+    ("_bigimpl", bigimpl_ast_tr),
+    ("_indexdefault", indexdefault_ast_tr),
+    ("_indexnum", indexnum_ast_tr),
+    ("_indexvar", indexvar_ast_tr),
+    ("_struct", struct_ast_tr)],
+   [("_abs", abs_tr),
+    ("_aprop", aprop_tr),
+    ("_ofclass", ofclass_tr),
+    ("_sort_constraint", sort_constraint_tr),
+    ("_TYPE", type_tr),
+    ("_DDDOT", dddot_tr),
+    ("_update_name", update_name_tr),
+    ("_index", index_tr)],
+   [],
+   [("_abs", abs_ast_tr'),
+    ("_idts", idtyp_ast_tr' "_idts"),
+    ("_pttrns", idtyp_ast_tr' "_pttrns"),
+    ("==>", impl_ast_tr'),
+    ("_index", index_ast_tr')]);
 
 val pure_trfunsT =
   [("_type_prop", type_prop_tr'), ("TYPE", type_tr'), ("_type_constraint_", type_constraint_tr')];
--- a/src/Pure/Syntax/syntax.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/Syntax/syntax.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -450,9 +450,9 @@
 fun guess_infix (Syntax ({gram, ...}, _)) c =
   (case Parser.guess_infix_lr gram c of
     SOME (s, l, r, j) => SOME
-     (if l then Mixfix.InfixlName (s, j)
-      else if r then Mixfix.InfixrName (s, j)
-      else Mixfix.InfixName (s, j))
+     (if l then Mixfix.Infixl (s, j)
+      else if r then Mixfix.Infixr (s, j)
+      else Mixfix.Infix (s, j))
   | NONE => NONE);
 
 
@@ -914,8 +914,8 @@
 
 end;
 
-structure BasicSyntax: BASIC_SYNTAX = Syntax;
-open BasicSyntax;
+structure Basic_Syntax: BASIC_SYNTAX = Syntax;
+open Basic_Syntax;
 
 forget_structure "Ast";
 forget_structure "SynExt";
--- a/src/Pure/pure_thy.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/pure_thy.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -309,11 +309,12 @@
     ("_indexdefault", typ "index",                     Delimfix ""),
     ("_indexvar",   typ "index",                       Delimfix "'\\<index>"),
     ("_struct",     typ "index => logic",              Mixfix ("\\<struct>_", [1000], 1000)),
+    ("_update_name", typ "idt",                        NoSyn),
     ("==>",         typ "prop => prop => prop",        Delimfix "op ==>"),
     (Term.dummy_patternN, typ "aprop",                 Delimfix "'_"),
     ("_sort_constraint", typ "type => prop",           Delimfix "(1SORT'_CONSTRAINT/(1'(_')))"),
     ("Pure.term",   typ "logic => prop",               Delimfix "TERM _"),
-    ("Pure.conjunction", typ "prop => prop => prop",   InfixrName ("&&&", 2))]
+    ("Pure.conjunction", typ "prop => prop => prop",   Infixr ("&&&", 2))]
   #> Sign.add_syntax_i applC_syntax
   #> Sign.add_modesyntax_i (Symbol.xsymbolsN, true)
    [("fun",      typ "type => type => type",   Mixfix ("(_/ \\<Rightarrow> _)", [1, 0], 0)),
@@ -325,9 +326,9 @@
     ("_idtypdummy", typ "type => idt",         Mixfix ("'_()\\<Colon>_", [], 0)),
     ("_type_constraint_", typ "'a",            NoSyn),
     ("_lambda",  typ "pttrns => 'a => logic",  Mixfix ("(3\\<lambda>_./ _)", [0, 3], 3)),
-    ("==",       typ "'a => 'a => prop",       InfixrName ("\\<equiv>", 2)),
+    ("==",       typ "'a => 'a => prop",       Infixr ("\\<equiv>", 2)),
     ("all_binder", typ "idts => prop => prop", Mixfix ("(3\\<And>_./ _)", [0, 0], 0)),
-    ("==>",      typ "prop => prop => prop",   InfixrName ("\\<Longrightarrow>", 1)),
+    ("==>",      typ "prop => prop => prop",   Infixr ("\\<Longrightarrow>", 1)),
     ("_DDDOT",   typ "aprop",                  Delimfix "\\<dots>"),
     ("_bigimpl", typ "asms => prop => prop",   Mixfix ("((1\\<lbrakk>_\\<rbrakk>)/ \\<Longrightarrow> _)", [0, 1], 1)),
     ("_DDDOT",   typ "logic",                  Delimfix "\\<dots>")]
@@ -336,7 +337,7 @@
   #> Sign.add_modesyntax_i ("HTML", false)
    [("_lambda", typ "pttrns => 'a => logic", Mixfix ("(3\\<lambda>_./ _)", [0, 3], 3))]
   #> Sign.add_consts_i
-   [(Binding.name "==", typ "'a => 'a => prop", InfixrName ("==", 2)),
+   [(Binding.name "==", typ "'a => 'a => prop", Infixr ("==", 2)),
     (Binding.name "==>", typ "prop => prop => prop", Mixfix ("(_/ ==> _)", [2, 1], 1)),
     (Binding.name "all", typ "('a => prop) => prop", Binder ("!!", 0, 0)),
     (Binding.name "prop", typ "prop => prop", NoSyn),
--- a/src/Pure/sign.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/Pure/sign.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -434,7 +434,7 @@
 fun add_types types thy = thy |> map_sign (fn (naming, syn, tsig, consts) =>
   let
     val syn' = Syntax.update_type_gram (map (fn (a, n, mx) => (Name.of_binding a, n, mx)) types) syn;
-    val decls = map (fn (a, n, mx) => (Binding.map_name (Syntax.type_name mx) a, n)) types;
+    val decls = map (fn (a, n, _) => (a, n)) types;
     val tsig' = fold (Type.add_type naming) decls tsig;
   in (naming, syn', tsig', consts) end);
 
@@ -450,12 +450,11 @@
 
 (* add type abbreviations *)
 
-fun gen_add_tyabbr parse_typ (a, vs, rhs, mx) thy =
+fun gen_add_tyabbr parse_typ (b, vs, rhs, mx) thy =
   thy |> map_sign (fn (naming, syn, tsig, consts) =>
     let
       val ctxt = ProofContext.init thy;
-      val syn' = Syntax.update_type_gram [(Name.of_binding a, length vs, mx)] syn;
-      val b = Binding.map_name (Syntax.type_name mx) a;
+      val syn' = Syntax.update_type_gram [(Name.of_binding b, length vs, mx)] syn;
       val abbr = (b, vs, certify_typ_mode Type.mode_syntax thy (parse_typ ctxt rhs))
         handle ERROR msg => cat_error msg ("in type abbreviation " ^ quote (Binding.str_of b));
       val tsig' = Type.add_abbrev naming abbr tsig;
@@ -471,8 +470,7 @@
   let
     val ctxt = ProofContext.init thy;
     fun prep (c, T, mx) = (c, certify_typ_mode Type.mode_syntax thy (parse_typ ctxt T), mx)
-      handle ERROR msg =>
-        cat_error msg ("in syntax declaration " ^ quote (Syntax.const_name mx c));
+      handle ERROR msg => cat_error msg ("in syntax declaration " ^ quote c);
   in thy |> map_syn (change_gram (is_logtype thy) mode (map prep args)) end;
 
 fun gen_add_syntax x = gen_syntax Syntax.update_const_gram x;
@@ -500,10 +498,8 @@
   let
     val ctxt = ProofContext.init thy;
     val prepT = Type.no_tvars o Term.no_dummyT o certify_typ thy o parse_typ ctxt;
-    fun prep (raw_b, raw_T, raw_mx) =
+    fun prep (b, raw_T, mx) =
       let
-        val (mx_name, mx) = Syntax.const_mixfix (Name.of_binding raw_b) raw_mx;
-        val b = Binding.map_name (K mx_name) raw_b;
         val c = full_name thy b;
         val c_syn = if authentic then Syntax.constN ^ c else Name.of_binding b;
         val T = (prepT raw_T handle TYPE (msg, _, _) => error msg) handle ERROR msg =>
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Cache_IO/cache_io.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -0,0 +1,121 @@
+(*  Title:      Tools/Cache_IO/cache_io.ML
+    Author:     Sascha Boehme, TU Muenchen
+
+Cache for output of external processes.
+*)
+
+signature CACHE_IO =
+sig
+  val with_tmp_file: string -> (Path.T -> 'a) -> 'a
+  val run: (Path.T -> string) -> Path.T -> string list
+  val run': (Path.T -> Path.T -> string) -> Path.T -> string list * string list
+
+  type cache
+  val make: Path.T -> cache
+  val cache_path_of: cache -> Path.T
+  val cached: cache -> (Path.T -> string) -> Path.T -> string list
+  val cached': cache -> (Path.T -> Path.T -> string) -> Path.T ->
+    string list * string list
+end
+
+structure Cache_IO : CACHE_IO =
+struct
+
+fun with_tmp_file name f =
+  let
+    val path = File.tmp_path (Path.explode (name ^ serial_string ()))
+    val x = Exn.capture f path
+    val _ = try File.rm path
+  in Exn.release x end
+
+fun run' make_cmd in_path =
+  with_tmp_file "cache-io-" (fn out_path =>
+    let
+      val (out2, _) = bash_output (make_cmd in_path out_path)
+      val out1 = the_default [] (try (rev o File.fold_lines cons out_path) [])
+    in (out1, split_lines out2) end)
+
+fun run make_cmd = snd o run' (fn in_path => fn _ => make_cmd in_path)
+
+
+
+abstype cache = Cache of {
+  path: Path.T,
+  table: (int * (int * int * int) Symtab.table) Synchronized.var }
+with
+
+
+fun cache_path_of (Cache {path, ...}) = path
+
+
+fun load cache_path =
+  let
+    fun err () = error ("Cache IO: corrupted cache file: " ^
+      File.shell_path cache_path)
+
+    fun int_of_string s =
+      (case read_int (explode s) of
+        (i, []) => i
+      | _ => err ())    
+
+    fun split line =
+      (case space_explode " " line of
+        [key, len1, len2] => (key, int_of_string len1, int_of_string len2)
+      | _ => err ())
+
+    fun parse line ((i, l), tab) =
+      if i = l
+      then
+        let val (key, l1, l2) = split line
+        in ((i+1, l+l1+l2+1), Symtab.update (key, (i+1, l1, l2)) tab) end
+      else ((i+1, l), tab)
+  in apfst fst (File.fold_lines parse cache_path ((1, 1), Symtab.empty)) end 
+
+fun make path =
+  let val table = if File.exists path then load path else (1, Symtab.empty)
+  in Cache {path=path, table=Synchronized.var (Path.implode path) table} end
+
+
+fun get_hash_key path =
+  let
+    val arg = File.shell_path path
+    val (out, res) = bash_output (getenv "COMPUTE_HASH_KEY" ^ " " ^ arg)
+  in
+    if res = 0 then hd (split_lines out)
+    else error ("Cache IO: failed to generate hash key for file " ^ arg)
+  end
+
+fun load_cached_result cache_path (p, len1, len2) =
+  let
+    fun load line (i, xsp) =
+      if i < p then (i+1, xsp)
+      else if i < p + len1 then (i+1, apfst (cons line) xsp)
+      else if i < p + len2 then (i+1, apsnd (cons line) xsp)
+      else (i, xsp)
+  in pairself rev (snd (File.fold_lines load cache_path (1, ([], [])))) end
+
+fun cached' (Cache {path=cache_path, table}) make_cmd in_path =
+  let val key = get_hash_key in_path
+  in
+    (case Symtab.lookup (snd (Synchronized.value table)) key of
+      SOME pos => load_cached_result cache_path pos
+    | NONE =>
+        let
+          val res as (out, err) = run' make_cmd in_path
+          val (l1, l2) = pairself length res
+          val header = key ^ " " ^ string_of_int l1 ^ " " ^ string_of_int l2
+          val lines = map (suffix "\n") (header :: out @ err)
+
+          val _ = Synchronized.change table (fn (p, tab) =>
+            if Symtab.defined tab key then (p, tab)
+            else
+              let val _ = File.append_list cache_path lines
+              in (p+l1+l2+1, Symtab.update (key, (p+1, l1, l2)) tab) end)
+        in res end)
+  end
+
+fun cached cache make_cmd =
+  snd o cached' cache (fn in_path => fn _ => make_cmd in_path)
+
+end
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Cache_IO/etc/settings	Wed Feb 17 09:51:46 2010 +0100
@@ -0,0 +1,4 @@
+ISABELLE_CACHE_IO="$COMPONENT"
+
+COMPUTE_HASH_KEY="$ISABELLE_CACHE_IO/lib/scripts/compute_hash_key"
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/Tools/Cache_IO/lib/scripts/compute_hash_key	Wed Feb 17 09:51:46 2010 +0100
@@ -0,0 +1,24 @@
+#!/usr/bin/env perl
+#
+# Author: Sascha Boehme, TU Muenchen
+#
+# Compute MD5 hash key.
+
+use strict;
+use warnings;
+use Digest::MD5;
+
+
+# argument
+
+my $file = $ARGV[0];
+
+
+# compute MD5 hash key
+
+my $md5 = Digest::MD5->new;
+open FILE, "<$file" or die "ERROR: Failed to open '$file' ($!)";
+$md5->addfile(*FILE);
+close FILE;
+print $md5->b64digest . "\n";
+
--- a/src/ZF/ind_syntax.ML	Wed Feb 17 09:48:53 2010 +0100
+++ b/src/ZF/ind_syntax.ML	Wed Feb 17 09:51:46 2010 +0100
@@ -65,15 +65,14 @@
   | dest_mem _ = error "Constructor specifications must have the form x:A";
 
 (*read a constructor specification*)
-fun read_construct ctxt (id, sprems, syn) =
+fun read_construct ctxt (id: string, sprems, syn: mixfix) =
     let val prems = map (Syntax.parse_term ctxt #> TypeInfer.constrain FOLogic.oT) sprems
           |> Syntax.check_terms ctxt
         val args = map (#1 o dest_mem) prems
         val T = (map (#2 o dest_Free) args) ---> iT
                 handle TERM _ => error
                     "Bad variable in constructor specification"
-        val name = Syntax.const_name syn id
-    in ((id,T,syn), name, args, prems) end;
+    in ((id,T,syn), id, args, prems) end;
 
 val read_constructs = map o map o read_construct;