merged
authorwenzelm
Mon, 06 Dec 2010 16:18:56 +0100
changeset 41019 b63cb15e96aa
parent 41018 83f241623b86 (diff)
parent 40967 5eb59b62e7de (current diff)
child 41020 f1e9db633212
merged
--- a/etc/components	Mon Dec 06 14:45:29 2010 +0100
+++ b/etc/components	Mon Dec 06 16:18:56 2010 +0100
@@ -17,3 +17,4 @@
 src/HOL/Library/Sum_Of_Squares
 src/HOL/Tools/SMT
 src/HOL/Tools/Predicate_Compile
+src/HOL/Mutabelle
--- a/etc/isar-keywords.el	Mon Dec 06 14:45:29 2010 +0100
+++ b/etc/isar-keywords.el	Mon Dec 06 16:18:56 2010 +0100
@@ -248,7 +248,7 @@
     "txt"
     "txt_raw"
     "typ"
-    "type_mapper"
+    "type_lifting"
     "type_notation"
     "typed_print_translation"
     "typedecl"
@@ -551,7 +551,7 @@
     "sublocale"
     "termination"
     "theorem"
-    "type_mapper"
+    "type_lifting"
     "typedef"))
 
 (defconst isar-keywords-qed
--- a/src/HOL/Datatype.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Datatype.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -13,6 +13,12 @@
   ("Tools/Datatype/datatype_realizer.ML")
 begin
 
+subsection {* Prelude: lifting over function space *}
+
+type_lifting map_fun
+  by (simp_all add: fun_eq_iff)
+
+
 subsection {* The datatype universe *}
 
 typedef (Node)
--- a/src/HOL/Fun.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Fun.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -7,6 +7,7 @@
 
 theory Fun
 imports Complete_Lattice
+uses ("Tools/type_lifting.ML")
 begin
 
 text{*As a simplification rule, it replaces all function equalities by
@@ -126,9 +127,6 @@
   "map_fun f g h x = g (h (f x))"
   by (simp add: map_fun_def)
 
-type_mapper map_fun
-  by (simp_all add: fun_eq_iff)
-
 
 subsection {* Injectivity and Bijectivity *}
 
@@ -774,7 +772,9 @@
   thus False by best
 qed
 
-subsection {* Proof tool setup *} 
+subsection {* Setup *} 
+
+subsubsection {* Proof tools *}
 
 text {* simplifies terms of the form
   f(...,x:=y,...,x:=z,...) to f(...,x:=z,...) *}
@@ -809,7 +809,7 @@
 *}
 
 
-subsection {* Code generator setup *}
+subsubsection {* Code generator *}
 
 types_code
   "fun"  ("(_ ->/ _)")
@@ -840,4 +840,9 @@
 code_const "id"
   (Haskell "id")
 
+
+subsubsection {* Functorial structure of types *}
+
+use "Tools/type_lifting.ML"
+
 end
--- a/src/HOL/HOL.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/HOL.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -32,7 +32,6 @@
   "Tools/async_manager.ML"
   "Tools/try.ML"
   ("Tools/cnf_funcs.ML")
-  ("Tools/type_mapper.ML")
   "~~/src/Tools/subtyping.ML"
 begin
 
@@ -714,7 +713,6 @@
   and [Pure.elim?] = iffD1 iffD2 impE
 
 use "Tools/hologic.ML"
-use "Tools/type_mapper.ML"
 
 
 subsubsection {* Atomizing meta-level connectives *}
--- a/src/HOL/IsaMakefile	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/IsaMakefile	Mon Dec 06 16:18:56 2010 +0100
@@ -147,7 +147,6 @@
   $(SRC)/Tools/solve_direct.ML \
   $(SRC)/Tools/value.ML \
   HOL.thy \
-  Tools/type_mapper.ML \
   Tools/hologic.ML \
   Tools/recfun_codegen.ML \
   Tools/simpdata.ML
@@ -243,6 +242,7 @@
   Tools/split_rule.ML \
   Tools/try.ML \
   Tools/typedef.ML \
+  Tools/type_lifting.ML \
   Transitive_Closure.thy \
   Typedef.thy \
   Wellfounded.thy
--- a/src/HOL/Library/Cset.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Library/Cset.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -188,7 +188,7 @@
   "map f (Cset.set xs) = Cset.set (remdups (List.map f xs))"
   by (simp add: set_def)
 
-type_mapper map
+type_lifting map
   by (simp_all add: image_image)
 
 definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a Cset.set \<Rightarrow> 'a Cset.set" where
--- a/src/HOL/Library/Dlist.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Library/Dlist.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -175,7 +175,7 @@
 
 section {* Functorial structure *}
 
-type_mapper map
+type_lifting map
   by (auto intro: dlist_eqI simp add: remdups_map_remdups comp_def)
 
 
--- a/src/HOL/Library/Mapping.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Library/Mapping.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -50,7 +50,7 @@
   "lookup (map f g m) = Option.map g \<circ> lookup m \<circ> f"
   by (simp add: map_def)
 
-type_mapper map
+type_lifting map
   by (auto intro!: mapping_eqI ext simp add: Option.map.compositionality Option.map.identity)
 
 
--- a/src/HOL/Library/Multiset.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Library/Multiset.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -1631,7 +1631,7 @@
   @{term "{#x+x|x:#M. x<c#}"}.
 *}
 
-type_mapper image_mset proof -
+type_lifting image_mset proof -
   fix f g A show "image_mset f (image_mset g A) = image_mset (\<lambda>x. f (g x)) A"
     by (induct A) simp_all
 next
--- a/src/HOL/List.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/List.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -881,7 +881,7 @@
   "length xs = length ys \<Longrightarrow> map snd (zip xs ys) = ys"
 by (induct rule:list_induct2, simp_all)
 
-type_mapper map
+type_lifting map
   by simp_all
 
 
--- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -369,7 +369,7 @@
     val problem =
       {state = st', goal = goal, subgoal = i,
        subgoal_count = Sledgehammer_Util.subgoal_count st,
-       facts = facts |> map Sledgehammer.Untranslated, only = true}
+       facts = facts |> map Sledgehammer.Untranslated}
     val time_limit =
       (case hard_timeout of
         NONE => I
--- a/src/HOL/Mirabelle/Tools/sledgehammer_tactic.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Mirabelle/Tools/sledgehammer_tactic.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -43,8 +43,7 @@
                         (prop_of goal))
     val problem =
       {state = Proof.init ctxt, goal = goal, subgoal = i,
-       facts = map Sledgehammer.Untranslated facts,
-       only = true, subgoal_count = n}
+       subgoal_count = n, facts = map Sledgehammer.Untranslated facts}
   in
     (case prover params (K "") problem of
       {outcome = NONE, used_facts, ...} => used_facts |> map fst |> SOME
--- a/src/HOL/Mirabelle/lib/Tools/mirabelle	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Mirabelle/lib/Tools/mirabelle	Mon Dec 06 16:18:56 2010 +0100
@@ -22,7 +22,7 @@
   echo "Usage: isabelle $PRG [OPTIONS] ACTIONS FILES"
   echo
   echo "  Options are:"
-  echo "    -L LOGIC     parent logic to use (default $ISABELLE_LOGIC)"
+  echo "    -L LOGIC     parent logic to use (default $MIRABELLE_LOGIC)"
   echo "    -T THEORY    parent theory to use (default $MIRABELLE_THEORY)"
   echo "    -O DIR       output directory for test data (default $out)"
   echo "    -t TIMEOUT   timeout for each action in seconds (default $timeout)"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mutabelle/etc/settings	Mon Dec 06 16:18:56 2010 +0100
@@ -0,0 +1,7 @@
+MUTABELLE_HOME="$COMPONENT"
+
+MUTABELLE_LOGIC=HOL
+MUTABELLE_IMPORT_THEORY=Complex_Main
+MUTABELLE_OUTPUT_PATH=/tmp/mutabelle
+
+ISABELLE_TOOLS="$ISABELLE_TOOLS:$COMPONENT/lib/Tools"
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Mutabelle/lib/Tools/mutabelle	Mon Dec 06 16:18:56 2010 +0100
@@ -0,0 +1,107 @@
+#!/usr/bin/env bash
+#
+# Author: Lukas Bulwahn
+#
+# DESCRIPTION: mutant-testing tool for counterexample generators and automated proof tools
+
+
+PRG="$(basename "$0")"
+
+function usage() {
+  echo
+  echo "Usage: isabelle $PRG [OPTIONS] THEORY"
+  echo
+  echo "  Options are:"
+  echo "    -L LOGIC     parent logic to use (default $MUTABELLE_LOGIC)"
+  echo "    -T THEORY    parent theory to use (default $MUTABELLE_IMPORT_THEORY)"
+  echo "    -O DIR       output directory for test data (default $MUTABELLE_OUTPUT_PATH)"
+  echo
+  echo "  THEORY is the name of the theory of which all theorems should be mutated and tested"
+  exit 1
+}
+
+
+## process command line
+
+# options
+
+while getopts "L:T:O:t:q?" OPT
+do
+  case "$OPT" in
+    L)
+      MIRABELLE_LOGIC="$OPTARG"
+      ;;
+    T)
+      MIRABELLE_IMPORT_THEORY="$OPTARG"
+      ;;
+    O)
+      MIRABELLE_OUTPUT_PATH="$OPTARG"
+      ;;
+    \?)
+      usage
+      ;;
+  esac
+done
+
+shift $(($OPTIND - 1))
+
+[ "$#" -ne 1 ] && usage
+
+MUTABELLE_TEST_THEORY="$1"
+
+## main
+
+echo "Starting Mutabelle..."
+
+# setup
+
+mkdir -p "$MIRABELLE_OUTPUT_PATH"
+
+echo "theory Mutabelle_Test
+imports $MUTABELLE_IMPORT_THEORY
+uses     
+  \"$MUTABELLE_HOME/mutabelle.ML\"
+  \"$MUTABELLE_HOME/mutabelle_extra.ML\"
+begin
+
+ML {*
+val mtds = [
+  MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.tester \"random\") \"random\",
+  MutabelleExtra.quickcheck_mtd (Config.put Quickcheck.tester \"exhaustive\") \"exhaustive\"
+]
+*}
+
+ML {*
+fun mutation_testing_of thy =
+  (MutabelleExtra.random_seed := 1.0;
+  MutabelleExtra.thms_of false thy
+  |> (fn thms => MutabelleExtra.mutate_theorems_and_write_report
+         @{theory} mtds thms (\"$MUTABELLE_OUTPUT_PATH/log\")))
+*}
+
+ML {*
+  mutation_testing_of @{theory $MUTABELLE_TEST_THEORY}
+*}
+
+end" > $MUTABELLE_OUTPUT_PATH/Mutabelle_Test.thy
+
+# execution
+
+$ISABELLE_PROCESS -e 'use_thy "$MUTABELLE_OUTPUT_PATH/Mutabelle_Test"' -q $MUTABELLE_LOGIC  > /dev/null 2>&1
+
+# make statistics
+
+GenuineCex_random=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_random: GenuineCex" | wc -l)
+NoCex_random=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_random: NoCex" | wc -l)
+Timeout_random=$(cat /$MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_random: Timeout" | wc -l)
+Errors_random=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_random: Error" | wc -l)
+
+GenuineCex_exh=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_exhaustive: GenuineCex" | wc -l)
+NoCex_exh=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_exhaustive: NoCex" | wc -l)
+Timeout_exh=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_exhaustive: Timeout" | wc -l)
+Errors_exh=$(cat $MUTABELLE_OUTPUT_PATH/log | grep "quickcheck_exhaustive: Error" | wc -l)
+
+
+echo "random      :" C: $GenuineCex_random N: $NoCex_random T: $Timeout_random E: $Errors_random
+echo "exhaustive  :" C: $GenuineCex_exh N: $NoCex_exh T: $Timeout_exh E: $Errors_exh
+
--- a/src/HOL/Mutabelle/mutabelle_extra.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Mutabelle/mutabelle_extra.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -25,6 +25,9 @@
 val solve_direct_mtd : mtd
 val try_mtd : mtd
 (*
+val sledgehammer_mtd : mtd
+*)
+(*
 val refute_mtd : mtd
 val nitpick_mtd : mtd
 *)
@@ -143,15 +146,26 @@
 
 fun invoke_try thy t =
   let
-    val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global thy) 
+    val state = Proof.theorem NONE (K I) (map (single o rpair []) [t]) (ProofContext.init_global thy)
   in
-    case Try.invoke_try NONE state of
+    case Try.invoke_try (SOME (seconds 5.0)) state of
       true => (Solved, ([], NONE))
     | false => (Unsolved, ([], NONE))
   end
 
 val try_mtd = ("try", invoke_try)
 
+(** sledgehammer **)
+(*
+fun invoke_sledgehammer thy t =
+  if can (Goal.prove_global thy (Term.add_free_names t [])  [] t)
+      (fn {context, ...} => Sledgehammer_Tactics.sledgehammer_with_metis_tac context 1) then
+    (Solved, ([], NONE))
+  else
+    (Unsolved, ([], NONE))
+
+val sledgehammer_mtd = ("sledgehammer", invoke_sledgehammer)
+*)
 (*
 fun invoke_refute thy t =
   let
@@ -287,7 +301,14 @@
  @{const_name "Pure.term"},
  @{const_name "top_class.top"},
  (*@{const_name "HOL.equal"},*)
- @{const_name "Quotient.Quot_True"}
+ @{const_name "Quotient.Quot_True"},
+ @{const_name "equal_fun_inst.equal_fun"},
+ @{const_name "equal_bool_inst.equal_bool"},
+ @{const_name "ord_fun_inst.less_eq_fun"},
+ @{const_name "ord_fun_inst.less_fun"},
+ @{const_name Metis.fequal},
+ @{const_name Meson.skolem},
+ @{const_name transfer_morphism}
  (*@{const_name "==>"}, @{const_name "=="}*)]
 
 val forbidden_mutant_consts =
@@ -426,11 +447,12 @@
       else
         mutants
     val mutants = mutants
-          |> take_random max_mutants
           |> map Mutabelle.freeze |> map freezeT
 (*          |> filter (not o is_forbidden_mutant) *)
           |> List.mapPartial (try (Sign.cert_term thy))
-          |> List.filter (is_some o try (cterm_of thy))
+          |> List.filter (is_some o try (Thm.cterm_of thy))
+          |> List.filter (is_some o try (Syntax.check_term (ProofContext.init_global thy)))
+          |> take_random max_mutants
     val _ = map (fn t => Output.urgent_message ("MUTANT: " ^ Syntax.string_of_term_global thy t)) mutants
   in
     create_entry thy thm exec mutants mtds
--- a/src/HOL/Nitpick_Examples/Mono_Nits.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Nitpick_Examples/Mono_Nits.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -8,93 +8,208 @@
 header {* Examples Featuring Nitpick's Monotonicity Check *}
 
 theory Mono_Nits
-imports Main
+imports Main (* "~/afp/thys/AVL-Trees/AVL2" "~/afp/thys/Huffman/Huffman" *)
 begin
 
 ML {*
-exception FAIL
+open Nitpick_Util
+open Nitpick_HOL
+open Nitpick_Preproc
 
+exception BUG
+
+val thy = @{theory}
+val ctxt = @{context}
+val stds = [(NONE, true)]
 val subst = []
-val defs = Nitpick_HOL.all_axioms_of @{context} subst |> #1
-val def_table = Nitpick_HOL.const_def_table @{context} subst defs
-val hol_ctxt : Nitpick_HOL.hol_context =
-  {thy = @{theory}, ctxt = @{context}, max_bisim_depth = ~1, boxes = [],
-   stds = [(NONE, true)], wfs = [], user_axioms = NONE, debug = false,
-   whacks = [], binary_ints = SOME false, destroy_constrs = false,
-   specialize = false, star_linear_preds = false, tac_timeout = NONE,
-   evals = [], case_names = [], def_table = def_table,
-   nondef_table = Symtab.empty, user_nondefs = [],
-   simp_table = Unsynchronized.ref Symtab.empty, psimp_table = Symtab.empty,
-   choice_spec_table = Symtab.empty, intro_table = Symtab.empty,
-   ground_thm_table = Inttab.empty, ersatz_table = [],
-   skolems = Unsynchronized.ref [], special_funs = Unsynchronized.ref [],
-   unrolled_preds = Unsynchronized.ref [], wf_cache = Unsynchronized.ref [],
-   constr_cache = Unsynchronized.ref []}
-(* term -> bool *)
+val case_names = case_const_names ctxt stds
+val (defs, built_in_nondefs, user_nondefs) = all_axioms_of ctxt subst
+val def_table = const_def_table ctxt subst defs
+val nondef_table = const_nondef_table (built_in_nondefs @ user_nondefs)
+val simp_table = Unsynchronized.ref (const_simp_table ctxt subst)
+val psimp_table = const_psimp_table ctxt subst
+val choice_spec_table = const_choice_spec_table ctxt subst
+val user_nondefs =
+  user_nondefs |> filter_out (is_choice_spec_axiom thy choice_spec_table)
+val intro_table = inductive_intro_table ctxt subst def_table
+val ground_thm_table = ground_theorem_table thy
+val ersatz_table = ersatz_table ctxt
+val hol_ctxt as {thy, ...} =
+  {thy = thy, ctxt = ctxt, max_bisim_depth = ~1, boxes = [],
+   stds = stds, wfs = [], user_axioms = NONE, debug = false,
+   whacks = [], binary_ints = SOME false, destroy_constrs = true,
+   specialize = false, star_linear_preds = false,
+   tac_timeout = NONE, evals = [], case_names = case_names,
+   def_table = def_table, nondef_table = nondef_table,
+   user_nondefs = user_nondefs, simp_table = simp_table,
+   psimp_table = psimp_table, choice_spec_table = choice_spec_table,
+   intro_table = intro_table, ground_thm_table = ground_thm_table,
+   ersatz_table = ersatz_table, skolems = Unsynchronized.ref [],
+   special_funs = Unsynchronized.ref [], unrolled_preds = Unsynchronized.ref [],
+   wf_cache = Unsynchronized.ref [], constr_cache = Unsynchronized.ref []}
+val binarize = false
+
 fun is_mono t =
-  Nitpick_Mono.formulas_monotonic hol_ctxt false @{typ 'a} ([t], [])
+  Nitpick_Mono.formulas_monotonic hol_ctxt binarize @{typ 'a} ([t], [])
+
 fun is_const t =
   let val T = fastype_of t in
     is_mono (Logic.mk_implies (Logic.mk_equals (Free ("dummyP", T), t),
                                @{const False}))
   end
-fun mono t = is_mono t orelse raise FAIL
-fun nonmono t = not (is_mono t) orelse raise FAIL
-fun const t = is_const t orelse raise FAIL
-fun nonconst t = not (is_const t) orelse raise FAIL
+
+fun mono t = is_mono t orelse raise BUG
+fun nonmono t = not (is_mono t) orelse raise BUG
+fun const t = is_const t orelse raise BUG
+fun nonconst t = not (is_const t) orelse raise BUG
 *}
 
-ML {* const @{term "A::('a\<Rightarrow>'b)"} *}
-ML {* const @{term "(A::'a set) = A"} *}
-ML {* const @{term "(A::'a set set) = A"} *}
-ML {* const @{term "(\<lambda>x::'a set. x a)"} *}
-ML {* const @{term "{{a::'a}} = C"} *}
-ML {* const @{term "{f::'a\<Rightarrow>nat} = {g::'a\<Rightarrow>nat}"} *}
-ML {* const @{term "A \<union> (B::'a set)"} *}
-ML {* const @{term "P (a::'a)"} *}
-ML {* const @{term "\<lambda>a::'a. b (c (d::'a)) (e::'a) (f::'a)"} *}
-ML {* const @{term "\<forall>A::'a set. A a"} *}
-ML {* const @{term "\<forall>A::'a set. P A"} *}
+ML {* Nitpick_Mono.trace := false *}
+
+ML {* const @{term "A\<Colon>('a\<Rightarrow>'b)"} *}
+(*
+ML {* const @{term "(A\<Colon>'a set) = A"} *}
+ML {* const @{term "(A\<Colon>'a set set) = A"} *}
+ML {* const @{term "(\<lambda>x\<Colon>'a set. x a)"} *}
+ML {* const @{term "{{a\<Colon>'a}} = C"} *}
+ML {* const @{term "{f\<Colon>'a\<Rightarrow>nat} = {g\<Colon>'a\<Rightarrow>nat}"} *}
+ML {* const @{term "A \<union> (B\<Colon>'a set)"} *}
+ML {* const @{term "\<lambda>A B x\<Colon>'a. A x \<or> B x"} *}
+ML {* const @{term "P (a\<Colon>'a)"} *}
+ML {* const @{term "\<lambda>a\<Colon>'a. b (c (d\<Colon>'a)) (e\<Colon>'a) (f\<Colon>'a)"} *}
+ML {* const @{term "\<forall>A\<Colon>'a set. A a"} *}
+ML {* const @{term "\<forall>A\<Colon>'a set. P A"} *}
 ML {* const @{term "P \<or> Q"} *}
-ML {* const @{term "A \<union> B = (C::'a set)"} *}
-ML {* const @{term "(if P then (A::'a set) else B) = C"} *}
-ML {* const @{term "let A = (C::'a set) in A \<union> B"} *}
-ML {* const @{term "THE x::'b. P x"} *}
-ML {* const @{term "(\<lambda>x::'a. False)"} *}
-ML {* const @{term "(\<lambda>x::'a. True)"} *}
-ML {* const @{term "Let (a::'a) A"} *}
-ML {* const @{term "A (a::'a)"} *}
-ML {* const @{term "insert (a::'a) A = B"} *}
-ML {* const @{term "- (A::'a set)"} *}
-ML {* const @{term "finite (A::'a set)"} *}
-ML {* const @{term "\<not> finite (A::'a set)"} *}
-ML {* const @{term "finite (A::'a set set)"} *}
-ML {* const @{term "\<lambda>a::'a. A a \<and> \<not> B a"} *}
-ML {* const @{term "A < (B::'a set)"} *}
-ML {* const @{term "A \<le> (B::'a set)"} *}
-ML {* const @{term "[a::'a]"} *}
-ML {* const @{term "[a::'a set]"} *}
-ML {* const @{term "[A \<union> (B::'a set)]"} *}
-ML {* const @{term "[A \<union> (B::'a set)] = [C]"} *}
-ML {* const @{term "{(\<lambda>x::'a. x = a)} = C"} *}
+ML {* const @{term "A \<union> B = (C\<Colon>'a set)"} *}
+ML {* const @{term "(\<lambda>A B x\<Colon>'a. A x \<or> B x) A B = C"} *}
+ML {* const @{term "(if P then (A\<Colon>'a set) else B) = C"} *}
+ML {* const @{term "let A = (C\<Colon>'a set) in A \<union> B"} *}
+ML {* const @{term "THE x\<Colon>'b. P x"} *}
+ML {* const @{term "(\<lambda>x\<Colon>'a. False)"} *}
+ML {* const @{term "(\<lambda>x\<Colon>'a. True)"} *}
+ML {* const @{term "(\<lambda>x\<Colon>'a. False) = (\<lambda>x\<Colon>'a. False)"} *}
+ML {* const @{term "(\<lambda>x\<Colon>'a. True) = (\<lambda>x\<Colon>'a. True)"} *}
+ML {* const @{term "Let (a\<Colon>'a) A"} *}
+ML {* const @{term "A (a\<Colon>'a)"} *}
+ML {* const @{term "insert (a\<Colon>'a) A = B"} *}
+ML {* const @{term "- (A\<Colon>'a set)"} *}
+ML {* const @{term "finite (A\<Colon>'a set)"} *}
+ML {* const @{term "\<not> finite (A\<Colon>'a set)"} *}
+ML {* const @{term "finite (A\<Colon>'a set set)"} *}
+ML {* const @{term "\<lambda>a\<Colon>'a. A a \<and> \<not> B a"} *}
+ML {* const @{term "A < (B\<Colon>'a set)"} *}
+ML {* const @{term "A \<le> (B\<Colon>'a set)"} *}
+ML {* const @{term "[a\<Colon>'a]"} *}
+ML {* const @{term "[a\<Colon>'a set]"} *}
+ML {* const @{term "[A \<union> (B\<Colon>'a set)]"} *}
+ML {* const @{term "[A \<union> (B\<Colon>'a set)] = [C]"} *}
+ML {* const @{term "{(\<lambda>x\<Colon>'a. x = a)} = C"} *}
+ML {* const @{term "(\<lambda>a\<Colon>'a. \<not> A a) = B"} *}
+ML {* const @{prop "\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> \<not> f a"} *}
+ML {* const @{term "\<lambda>A B x\<Colon>'a. A x \<and> B x \<and> A = B"} *}
+ML {* const @{term "p = (\<lambda>(x\<Colon>'a) (y\<Colon>'a). P x \<or> \<not> Q y)"} *}
+ML {* const @{term "p = (\<lambda>(x\<Colon>'a) (y\<Colon>'a). p x y \<Colon> bool)"} *}
+ML {* const @{term "p = (\<lambda>A B x. A x \<and> \<not> B x) (\<lambda>x. True) (\<lambda>y. x \<noteq> y)"} *}
+ML {* const @{term "p = (\<lambda>y. x \<noteq> y)"} *}
+ML {* const @{term "(\<lambda>x. (p\<Colon>'a\<Rightarrow>bool\<Rightarrow>bool) x False)"} *}
+ML {* const @{term "(\<lambda>x y. (p\<Colon>'a\<Rightarrow>'a\<Rightarrow>bool\<Rightarrow>bool) x y False)"} *}
+ML {* const @{term "f = (\<lambda>x\<Colon>'a. P x \<longrightarrow> Q x)"} *}
+ML {* const @{term "\<forall>a\<Colon>'a. P a"} *}
+
+ML {* nonconst @{term "\<forall>P (a\<Colon>'a). P a"} *}
+ML {* nonconst @{term "THE x\<Colon>'a. P x"} *}
+ML {* nonconst @{term "SOME x\<Colon>'a. P x"} *}
+ML {* nonconst @{term "(\<lambda>A B x\<Colon>'a. A x \<or> B x) = myunion"} *}
+ML {* nonconst @{term "(\<lambda>x\<Colon>'a. False) = (\<lambda>x\<Colon>'a. True)"} *}
+ML {* nonconst @{prop "\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h"} *}
+
+ML {* mono @{prop "Q (\<forall>x\<Colon>'a set. P x)"} *}
+ML {* mono @{prop "P (a\<Colon>'a)"} *}
+ML {* mono @{prop "{a} = {b\<Colon>'a}"} *}
+ML {* mono @{prop "P (a\<Colon>'a) \<and> P \<union> P = P"} *}
+ML {* mono @{prop "\<forall>F\<Colon>'a set set. P"} *}
+ML {* mono @{prop "\<not> (\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h)"} *}
+ML {* mono @{prop "\<not> Q (\<forall>x\<Colon>'a set. P x)"} *}
+ML {* mono @{prop "\<not> (\<forall>x\<Colon>'a. P x)"} *}
+ML {* mono @{prop "myall P = (P = (\<lambda>x\<Colon>'a. True))"} *}
+ML {* mono @{prop "myall P = (P = (\<lambda>x\<Colon>'a. False))"} *}
+ML {* mono @{prop "\<forall>x\<Colon>'a. P x"} *}
+ML {* mono @{term "(\<lambda>A B x\<Colon>'a. A x \<or> B x) \<noteq> myunion"} *}
+
+ML {* nonmono @{prop "A = (\<lambda>x::'a. True) \<and> A = (\<lambda>x. False)"} *}
+ML {* nonmono @{prop "\<forall>F f g (h\<Colon>'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h"} *}
+*)
+
+ML {*
+val preproc_timeout = SOME (seconds 5.0)
+val mono_timeout = SOME (seconds 1.0)
 
-ML {* nonconst @{term "\<forall>P (a::'a). P a"} *}
-ML {* nonconst @{term "\<forall>a::'a. P a"} *}
-ML {* nonconst @{term "(\<lambda>a::'a. \<not> A a) = B"} *}
-ML {* nonconst @{term "THE x::'a. P x"} *}
-ML {* nonconst @{term "SOME x::'a. P x"} *}
+fun all_unconcealed_theorems_of thy =
+  let val facts = Global_Theory.facts_of thy in
+    Facts.fold_static
+        (fn (name, ths) =>
+            if Facts.is_concealed facts name then I
+            else append (map (`(Thm.get_name_hint)) ths))
+        facts []
+  end
+
+fun is_forbidden_theorem name =
+  length (space_explode "." name) <> 2 orelse
+  String.isPrefix "type_definition" (List.last (space_explode "." name)) orelse
+  String.isSuffix "_def" name orelse
+  String.isSuffix "_raw" name
+
+fun theorems_of thy =
+  filter (fn (name, th) =>
+             not (is_forbidden_theorem name) andalso
+             (theory_of_thm th, thy) |> pairself Context.theory_name |> op =)
+         (all_unconcealed_theorems_of thy)
 
-ML {* mono @{prop "Q (\<forall>x::'a set. P x)"} *}
-ML {* mono @{prop "P (a::'a)"} *}
-ML {* mono @{prop "{a} = {b::'a}"} *}
-ML {* mono @{prop "P (a::'a) \<and> P \<union> P = P"} *}
-ML {* mono @{prop "\<forall>F::'a set set. P"} *}
-ML {* mono @{prop "\<not> (\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h)"} *}
-ML {* mono @{prop "\<not> Q (\<forall>x::'a set. P x)"} *}
-ML {* mono @{prop "\<not> (\<forall>x::'a. P x)"} *}
+fun check_formulas tsp =
+  let
+    fun is_type_actually_monotonic T =
+      Nitpick_Mono.formulas_monotonic hol_ctxt binarize T tsp
+    val free_Ts = fold Term.add_tfrees (op @ tsp) [] |> map TFree
+    val (mono_free_Ts, nonmono_free_Ts) =
+      time_limit mono_timeout (List.partition is_type_actually_monotonic)
+                 free_Ts
+  in
+    if not (null mono_free_Ts) then "MONO"
+    else if not (null nonmono_free_Ts) then "NONMONO"
+    else "NIX"
+  end
+  handle TimeLimit.TimeOut => "TIMEOUT"
+       | NOT_SUPPORTED _ => "UNSUP"
+       | _ => "UNKNOWN"
 
-ML {* nonmono @{prop "\<forall>x::'a. P x"} *}
-ML {* nonmono @{prop "myall P = (P = (\<lambda>x::'a. True))"} *}
-ML {* nonmono @{prop "\<forall>F f g (h::'a set). F f \<and> F g \<and> \<not> f a \<and> g a \<longrightarrow> F h"} *}
+fun check_theory thy =
+  let
+    val finitizes = [(NONE, SOME false)]
+    val monos = [(NONE, SOME false)]
+    val path = File.tmp_path (Context.theory_name thy ^ ".out" |> Path.explode)
+    val _ = File.write path ""
+    fun check_theorem (name, th) =
+      let
+        val t = th |> prop_of |> Type.legacy_freeze |> close_form
+        val neg_t = Logic.mk_implies (t, @{prop False})
+        val (nondef_ts, def_ts, _, _, _) =
+          time_limit preproc_timeout
+              (preprocess_formulas hol_ctxt finitizes monos []) neg_t
+        val res = name ^ ": " ^ check_formulas (nondef_ts, def_ts)
+      in File.append path (res ^ "\n"); writeln res end
+      handle TimeLimit.TimeOut => ()
+  in thy |> theorems_of |> List.app check_theorem end
+*}
+
+(*
+ML {* check_theory @{theory AVL2} *}
+ML {* check_theory @{theory Fun} *}
+ML {* check_theory @{theory Huffman} *}
+ML {* check_theory @{theory List} *}
+ML {* check_theory @{theory Map} *}
+ML {* check_theory @{theory Relation} *}
+*)
+
+ML {* getenv "ISABELLE_TMP" *}
 
 end
--- a/src/HOL/Option.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Option.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -81,7 +81,7 @@
     "map f o sum_case g h = sum_case (map f o g) (map f o h)"
   by (rule ext) (simp split: sum.split)
 
-type_mapper Option.map proof -
+type_lifting Option.map proof -
   fix f g x
   show "Option.map f (Option.map g x) = Option.map (\<lambda>x. f (g x)) x"
     by (cases x) simp_all
--- a/src/HOL/Product_Type.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Product_Type.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -835,7 +835,7 @@
   "map_pair f g (a, b) = (f a, g b)"
   by (simp add: map_pair_def)
 
-type_mapper map_pair
+type_lifting map_pair
   by (simp_all add: split_paired_all)
 
 lemma fst_map_pair [simp]:
--- a/src/HOL/Quickcheck.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Quickcheck.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -129,7 +129,6 @@
 use "Tools/quickcheck_generators.ML"
 setup Quickcheck_Generators.setup
 
-declare [[quickcheck_tester = random]]
 
 subsection {* Code setup *}
 
--- a/src/HOL/Sum_Type.thy	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Sum_Type.thy	Mon Dec 06 16:18:56 2010 +0100
@@ -95,7 +95,7 @@
   "sum_map f1 f2 (Inl a) = Inl (f1 a)"
 | "sum_map f1 f2 (Inr a) = Inr (f2 a)"
 
-type_mapper sum_map proof -
+type_lifting sum_map proof -
   fix f g h i s
   show "sum_map f g (sum_map h i s) = sum_map (\<lambda>x. f (h x)) (\<lambda>x. g (i x)) s"
     by (cases s) simp_all
--- a/src/HOL/Tools/Nitpick/nitpick.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -270,7 +270,7 @@
     val intro_table = inductive_intro_table ctxt subst def_table
     val ground_thm_table = ground_theorem_table thy
     val ersatz_table = ersatz_table ctxt
-    val (hol_ctxt as {wf_cache, ...}) =
+    val hol_ctxt as {wf_cache, ...} =
       {thy = thy, ctxt = ctxt, max_bisim_depth = max_bisim_depth, boxes = boxes,
        stds = stds, wfs = wfs, user_axioms = user_axioms, debug = debug,
        whacks = whacks, binary_ints = binary_ints,
--- a/src/HOL/Tools/Nitpick/nitpick_mono.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/Nitpick/nitpick_mono.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -23,23 +23,27 @@
 open Nitpick_Util
 open Nitpick_HOL
 
-type var = int
+structure PL = PropLogic
 
 datatype sign = Plus | Minus
-datatype sign_atom = S of sign | V of var
+
+type var = int
 
-type literal = var * sign
+datatype annotation = Gen | New | Fls | Tru
+datatype annotation_atom = A of annotation | V of var
+
+type assign_literal = var * (sign * annotation)
 
 datatype mtyp =
   MAlpha |
-  MFun of mtyp * sign_atom * mtyp |
+  MFun of mtyp * annotation_atom * mtyp |
   MPair of mtyp * mtyp |
   MType of string * mtyp list |
   MRec of string * typ list
 
 datatype mterm =
   MRaw of term * mtyp |
-  MAbs of string * typ * mtyp * sign_atom * mterm |
+  MAbs of string * typ * mtyp * annotation_atom * mterm |
   MApp of mterm * mterm
 
 type mdata =
@@ -58,22 +62,29 @@
 val trace = Unsynchronized.ref false
 fun trace_msg msg = if !trace then tracing (msg ()) else ()
 
+fun string_for_sign Plus = "+"
+  | string_for_sign Minus = "-"
+
+fun negate_sign Plus = Minus
+  | negate_sign Minus = Plus
+
 val string_for_var = signed_string_of_int
 fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>"
   | string_for_vars sep xs = space_implode sep (map string_for_var xs)
 fun subscript_string_for_vars sep xs =
   if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>"
 
-fun string_for_sign Plus = "+"
-  | string_for_sign Minus = "-"
+fun string_for_annotation Gen = "G"
+  | string_for_annotation New = "N"
+  | string_for_annotation Fls = "F"
+  | string_for_annotation Tru = "T"
 
-fun xor sn1 sn2 = if sn1 = sn2 then Plus else Minus
-val negate = xor Minus
+fun string_for_annotation_atom (A a) = string_for_annotation a
+  | string_for_annotation_atom (V x) = string_for_var x
 
-fun string_for_sign_atom (S sn) = string_for_sign sn
-  | string_for_sign_atom (V x) = string_for_var x
-
-fun string_for_literal (x, sn) = string_for_var x ^ " = " ^ string_for_sign sn
+fun string_for_assign_literal (x, (sn, a)) =
+  string_for_var x ^ (case sn of Plus => " = " | Minus => " \<noteq> ") ^
+  string_for_annotation a
 
 val bool_M = MType (@{type_name bool}, [])
 val dummy_M = MType (nitpick_prefix ^ "dummy", [])
@@ -101,9 +112,9 @@
            "_"
          else case M of
              MAlpha => "\<alpha>"
-           | MFun (M1, a, M2) =>
+           | MFun (M1, aa, M2) =>
              aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^
-             string_for_sign_atom a ^ "\<^esup> " ^ aux prec M2
+             string_for_annotation_atom aa ^ "\<^esup> " ^ aux prec M2
            | MPair (M1, M2) => aux (prec + 1) M1 ^ " \<times> " ^ aux prec M2
            | MType (s, []) =>
              if s = @{type_name prop} orelse s = @{type_name bool} then "o"
@@ -133,16 +144,16 @@
         (if need_parens then "(" else "") ^
         (case m of
            MRaw (t, M) => Syntax.string_of_term ctxt t ^ mtype_annotation M
-         | MAbs (s, _, M, a, m) =>
+         | MAbs (s, _, M, aa, m) =>
            "\<lambda>" ^ s ^ mtype_annotation M ^ ".\<^bsup>" ^
-           string_for_sign_atom a ^ "\<^esup> " ^ aux prec m
+           string_for_annotation_atom aa ^ "\<^esup> " ^ aux prec m
          | MApp (m1, m2) => aux prec m1 ^ " " ^ aux (prec + 1) m2) ^
         (if need_parens then ")" else "")
       end
   in aux 0 end
 
 fun mtype_of_mterm (MRaw (_, M)) = M
-  | mtype_of_mterm (MAbs (_, _, M, a, m)) = MFun (M, a, mtype_of_mterm m)
+  | mtype_of_mterm (MAbs (_, _, M, aa, m)) = MFun (M, aa, mtype_of_mterm m)
   | mtype_of_mterm (MApp (m1, _)) =
     case mtype_of_mterm m1 of
       MFun (_, _, M12) => M12
@@ -185,11 +196,11 @@
   | exists_alpha_sub_mtype_fresh (MRec _) = true
 
 fun constr_mtype_for_binders z Ms =
-  fold_rev (fn M => curry3 MFun M (S Minus)) Ms (MRec z)
+  fold_rev (fn M => curry3 MFun M (A Gen)) Ms (MRec z)
 
 fun repair_mtype _ _ MAlpha = MAlpha
-  | repair_mtype cache seen (MFun (M1, a, M2)) =
-    MFun (repair_mtype cache seen M1, a, repair_mtype cache seen M2)
+  | repair_mtype cache seen (MFun (M1, aa, M2)) =
+    MFun (repair_mtype cache seen M1, aa, repair_mtype cache seen M2)
   | repair_mtype cache seen (MPair Mp) =
     MPair (pairself (repair_mtype cache seen) Mp)
   | repair_mtype cache seen (MType (s, Ms)) =
@@ -218,6 +229,8 @@
   | is_fin_fun_supported_type @{typ bool} = true
   | is_fin_fun_supported_type (Type (@{type_name option}, _)) = true
   | is_fin_fun_supported_type _ = false
+
+(* TODO: clean this up *)
 fun fin_fun_body _ _ (t as @{term False}) = SOME t
   | fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t
   | fin_fun_body dom_T ran_T
@@ -233,17 +246,19 @@
                 $ (Const (@{const_name unknown}, ran_T)) $ (t0 $ t1 $ t2 $ t3)))
   | fin_fun_body _ _ _ = NONE
 
+(* ### FIXME: make sure well-annotated! *)
+
 fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus
                             T1 T2 =
   let
     val M1 = fresh_mtype_for_type mdata all_minus T1
     val M2 = fresh_mtype_for_type mdata all_minus T2
-    val a = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
-               is_fin_fun_supported_type (body_type T2) then
-              V (Unsynchronized.inc max_fresh)
-            else
-              S Minus
-  in (M1, a, M2) end
+    val aa = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso
+                is_fin_fun_supported_type (body_type T2) then
+               V (Unsynchronized.inc max_fresh)
+             else
+               A Gen
+  in (M1, aa, M2) end
 and fresh_mtype_for_type (mdata as {hol_ctxt as {ctxt, ...}, binarize, alpha_T,
                                     datatype_mcache, constr_mcache, ...})
                          all_minus =
@@ -294,14 +309,23 @@
       | _ => MType (simple_string_of_typ T, [])
   in do_type end
 
+val ground_and_sole_base_constrs = [] (* FIXME: [@{const_name Nil}, @{const_name None}], cf. lists_empty *)
+
 fun prodM_factors (MPair (M1, M2)) = maps prodM_factors [M1, M2]
   | prodM_factors M = [M]
 fun curried_strip_mtype (MFun (M1, _, M2)) =
     curried_strip_mtype M2 |>> append (prodM_factors M1)
   | curried_strip_mtype M = ([], M)
 fun sel_mtype_from_constr_mtype s M =
-  let val (arg_Ms, dataM) = curried_strip_mtype M in
-    MFun (dataM, S Minus,
+  let
+    val (arg_Ms, dataM) = curried_strip_mtype M
+    val a = if member (op =) ground_and_sole_base_constrs
+                             (constr_name_for_sel_like s) then
+              Fls
+            else
+              Gen
+  in
+    MFun (dataM, A a,
           case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n)
   end
 
@@ -323,230 +347,437 @@
   x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize
     |> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s
 
-fun resolve_sign_atom lits (V x) =
-    x |> AList.lookup (op =) lits |> Option.map S |> the_default (V x)
-  | resolve_sign_atom _ a = a
-fun resolve_mtype lits =
+fun resolve_annotation_atom asgs (V x) =
+    x |> AList.lookup (op =) asgs |> Option.map A |> the_default (V x)
+  | resolve_annotation_atom _ aa = aa
+fun resolve_mtype asgs =
   let
     fun aux MAlpha = MAlpha
-      | aux (MFun (M1, a, M2)) = MFun (aux M1, resolve_sign_atom lits a, aux M2)
+      | aux (MFun (M1, aa, M2)) =
+        MFun (aux M1, resolve_annotation_atom asgs aa, aux M2)
       | aux (MPair Mp) = MPair (pairself aux Mp)
       | aux (MType (s, Ms)) = MType (s, map aux Ms)
       | aux (MRec z) = MRec z
   in aux end
 
-datatype comp_op = Eq | Leq
+datatype comp_op = Eq | Neq | Leq
 
-type comp = sign_atom * sign_atom * comp_op * var list
-type sign_expr = literal list
+type comp = annotation_atom * annotation_atom * comp_op * var list
+type assign_clause = assign_literal list
 
-type constraint_set = literal list * comp list * sign_expr list
+type constraint_set = comp list * assign_clause list
 
 fun string_for_comp_op Eq = "="
+  | string_for_comp_op Neq = "\<noteq>"
   | string_for_comp_op Leq = "\<le>"
 
-fun string_for_sign_expr [] = "\<bot>"
-  | string_for_sign_expr lits =
-    space_implode " \<or> " (map string_for_literal lits)
+fun string_for_comp (aa1, aa2, cmp, xs) =
+  string_for_annotation_atom aa1 ^ " " ^ string_for_comp_op cmp ^
+  subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_annotation_atom aa2
+
+fun string_for_assign_clause NONE = "\<top>"
+  | string_for_assign_clause (SOME []) = "\<bot>"
+  | string_for_assign_clause (SOME asgs) =
+    space_implode " \<or> " (map string_for_assign_literal asgs)
 
-fun do_literal _ NONE = NONE
-  | do_literal (x, sn) (SOME lits) =
-    case AList.lookup (op =) lits x of
-      SOME sn' => if sn = sn' then SOME lits else NONE
-    | NONE => SOME ((x, sn) :: lits)
+fun add_assign_literal (x, (sn, a)) clauses =
+  if exists (fn [(x', (sn', a'))] =>
+                x = x' andalso ((sn = sn' andalso a <> a') orelse
+                                (sn <> sn' andalso a = a'))
+              | _ => false) clauses then
+    NONE
+  else
+    SOME ([(x, (sn, a))] :: clauses)
+
+fun add_assign_disjunct _ NONE = NONE
+  | add_assign_disjunct asg (SOME asgs) = SOME (insert (op =) asg asgs)
+
+fun add_assign_clause NONE = I
+  | add_assign_clause (SOME clause) = insert (op =) clause
 
-fun do_sign_atom_comp Eq [] a1 a2 (accum as (lits, comps)) =
-    (case (a1, a2) of
-       (S sn1, S sn2) => if sn1 = sn2 then SOME accum else NONE
-     | (V x1, S sn2) =>
-       Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits))
-     | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps)
-     | _ => do_sign_atom_comp Eq [] a2 a1 accum)
-  | do_sign_atom_comp Leq [] a1 a2 (accum as (lits, comps)) =
-    (case (a1, a2) of
-       (_, S Minus) => SOME accum
-     | (S Plus, _) => SOME accum
-     | (S Minus, S Plus) => NONE
-     | (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps)
-     | _ => do_sign_atom_comp Eq [] a1 a2 accum)
-  | do_sign_atom_comp cmp xs a1 a2 (lits, comps) =
-    SOME (lits, insert (op =) (a1, a2, cmp, xs) comps)
+fun annotation_comp Eq a1 a2 = (a1 = a2)
+  | annotation_comp Neq a1 a2 = (a1 <> a2)
+  | annotation_comp Leq a1 a2 = (a1 = a2 orelse a2 = Gen)
+
+fun sign_for_comp_op Eq = Plus
+  | sign_for_comp_op Neq = Minus
+  | sign_for_comp_op Leq = raise BAD ("sign_for_comp_op", "unexpected \"Leq\"")
+
+fun do_annotation_atom_comp Leq [] aa1 aa2 (cset as (comps, clauses)) =
+    (case (aa1, aa2) of
+       (A a1, A a2) => if annotation_comp Leq a1 a2 then SOME cset else NONE
+     | _ => SOME (insert (op =) (aa1, aa2, Leq, []) comps, clauses))
+  | do_annotation_atom_comp cmp [] aa1 aa2 (cset as (comps, clauses)) =
+    (case (aa1, aa2) of
+       (A a1, A a2) => if annotation_comp cmp a1 a2 then SOME cset else NONE
+     | (V x1, A a2) =>
+       clauses |> add_assign_literal (x1, (sign_for_comp_op cmp, a2))
+               |> Option.map (pair comps)
+     | (A _, V _) => do_annotation_atom_comp cmp [] aa2 aa1 cset
+     | (V _, V _) => SOME (insert (op =) (aa1, aa2, cmp, []) comps, clauses))
+  | do_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) =
+    SOME (insert (op =) (aa1, aa2, cmp, xs) comps, clauses)
+
+fun add_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) =
+  (trace_msg (fn () => "*** Add " ^ string_for_comp (aa1, aa2, cmp, xs));
+   case do_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) of
+     NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
+   | SOME cset => cset)
 
 fun do_mtype_comp _ _ _ _ NONE = NONE
-  | do_mtype_comp _ _ MAlpha MAlpha accum = accum
-  | do_mtype_comp Eq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
-                  (SOME accum) =
-     accum |> do_sign_atom_comp Eq xs a1 a2 |> do_mtype_comp Eq xs M11 M21
-           |> do_mtype_comp Eq xs M12 M22
-  | do_mtype_comp Leq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22))
-                  (SOME accum) =
+  | do_mtype_comp _ _ MAlpha MAlpha cset = cset
+  | do_mtype_comp Eq xs (MFun (M11, aa1, M12)) (MFun (M21, aa2, M22))
+                  (SOME cset) =
+    cset |> do_annotation_atom_comp Eq xs aa1 aa2
+         |> do_mtype_comp Eq xs M11 M21 |> do_mtype_comp Eq xs M12 M22
+  | do_mtype_comp Leq xs (MFun (M11, aa1, M12)) (MFun (M21, aa2, M22))
+                  (SOME cset) =
     (if exists_alpha_sub_mtype M11 then
-       accum |> do_sign_atom_comp Leq xs a1 a2
-             |> do_mtype_comp Leq xs M21 M11
-             |> (case a2 of
-                   S Minus => I
-                 | S Plus => do_mtype_comp Leq xs M11 M21
-                 | V x => do_mtype_comp Leq (x :: xs) M11 M21)
+       cset |> do_annotation_atom_comp Leq xs aa1 aa2
+            |> do_mtype_comp Leq xs M21 M11
+            |> (case aa2 of
+                  A Gen => I
+                | A _ => do_mtype_comp Leq xs M11 M21
+                | V x => do_mtype_comp Leq (x :: xs) M11 M21)
      else
-       SOME accum)
+       SOME cset)
     |> do_mtype_comp Leq xs M12 M22
   | do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22))
-                  accum =
-    (accum |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
+                  cset =
+    (cset |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)]
      handle ListPair.UnequalLengths =>
             raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2], []))
-  | do_mtype_comp _ _ (MType _) (MType _) accum =
-    accum (* no need to compare them thanks to the cache *)
+  | do_mtype_comp _ _ (MType _) (MType _) cset =
+    cset (* no need to compare them thanks to the cache *)
   | do_mtype_comp cmp _ M1 M2 _ =
     raise MTYPE ("Nitpick_Mono.do_mtype_comp (" ^ string_for_comp_op cmp ^ ")",
                  [M1, M2], [])
 
-fun add_mtype_comp cmp M1 M2 ((lits, comps, sexps) : constraint_set) =
-    (trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
-                         string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
-     case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of
-       NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
-     | SOME (lits, comps) => (lits, comps, sexps))
+fun add_mtype_comp cmp M1 M2 cset =
+  (trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^
+                       string_for_comp_op cmp ^ " " ^ string_for_mtype M2);
+   case SOME cset |> do_mtype_comp cmp [] M1 M2 of
+     NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
+   | SOME cset => cset)
 
 val add_mtypes_equal = add_mtype_comp Eq
 val add_is_sub_mtype = add_mtype_comp Leq
 
 fun do_notin_mtype_fv _ _ _ NONE = NONE
-  | do_notin_mtype_fv Minus _ MAlpha accum = accum
+  | do_notin_mtype_fv Minus _ MAlpha cset = cset
   | do_notin_mtype_fv Plus [] MAlpha _ = NONE
-  | do_notin_mtype_fv Plus [(x, sn)] MAlpha (SOME (lits, sexps)) =
-    SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps)
-  | do_notin_mtype_fv Plus sexp MAlpha (SOME (lits, sexps)) =
-    SOME (lits, insert (op =) sexp sexps)
-  | do_notin_mtype_fv sn sexp (MFun (M1, S sn', M2)) accum =
-    accum |> (if sn' = Plus andalso sn = Plus then
-                do_notin_mtype_fv Plus sexp M1
-              else
-                I)
-          |> (if sn' = Minus orelse sn = Plus then
-                do_notin_mtype_fv Minus sexp M1
-              else
-                I)
-          |> do_notin_mtype_fv sn sexp M2
-  | do_notin_mtype_fv Plus sexp (MFun (M1, V x, M2)) accum =
-    accum |> (case do_literal (x, Minus) (SOME sexp) of
-                NONE => I
-              | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
-          |> do_notin_mtype_fv Minus sexp M1
-          |> do_notin_mtype_fv Plus sexp M2
-  | do_notin_mtype_fv Minus sexp (MFun (M1, V x, M2)) accum =
-    accum |> (case do_literal (x, Plus) (SOME sexp) of
-                NONE => I
-              | SOME sexp' => do_notin_mtype_fv Plus sexp' M1)
-          |> do_notin_mtype_fv Minus sexp M2
-  | do_notin_mtype_fv sn sexp (MPair (M1, M2)) accum =
-    accum |> fold (do_notin_mtype_fv sn sexp) [M1, M2]
-  | do_notin_mtype_fv sn sexp (MType (_, Ms)) accum =
-    accum |> fold (do_notin_mtype_fv sn sexp) Ms
-  | do_notin_mtype_fv _ _ M _ =
-    raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
+  | do_notin_mtype_fv Plus [asg] MAlpha (SOME clauses) =
+    clauses |> add_assign_literal asg
+  | do_notin_mtype_fv Plus unless MAlpha (SOME clauses) =
+    SOME (insert (op =) unless clauses)
+  | do_notin_mtype_fv sn unless (MFun (M1, A a, M2)) cset =
+    cset |> (if a <> Gen andalso sn = Plus then do_notin_mtype_fv Plus unless M1
+             else I)
+         |> (if a = Gen orelse sn = Plus then do_notin_mtype_fv Minus unless M1
+             else I)
+         |> do_notin_mtype_fv sn unless M2
+  | do_notin_mtype_fv Plus unless (MFun (M1, V x, M2)) cset =
+    cset |> (case add_assign_disjunct (x, (Plus, Gen)) (SOME unless) of
+               NONE => I
+             | SOME unless' => do_notin_mtype_fv Plus unless' M1)
+         |> do_notin_mtype_fv Minus unless M1
+         |> do_notin_mtype_fv Plus unless M2
+  | do_notin_mtype_fv Minus unless (MFun (M1, V x, M2)) cset =
+    cset |> (case fold (fn a => add_assign_disjunct (x, (Plus, a))) [Fls, Tru]
+                       (SOME unless) of
+               NONE => I
+             | SOME unless' => do_notin_mtype_fv Plus unless' M1)
+         |> do_notin_mtype_fv Minus unless M2
+  | do_notin_mtype_fv sn unless (MPair (M1, M2)) cset =
+    cset |> fold (do_notin_mtype_fv sn unless) [M1, M2]
+  | do_notin_mtype_fv sn unless (MType (_, Ms)) cset =
+    cset |> fold (do_notin_mtype_fv sn unless) Ms
+ | do_notin_mtype_fv _ _ M _ =
+   raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], [])
 
-fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) =
-    (trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
-                         (case sn of Minus => "concrete" | Plus => "complete"));
-     case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of
-       NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
-     | SOME (lits, sexps) => (lits, comps, sexps))
+fun add_notin_mtype_fv sn unless M (comps, clauses) =
+  (trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^
+                       (case sn of Minus => "concrete" | Plus => "complete"));
+   case SOME clauses |> do_notin_mtype_fv sn unless M of
+     NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ())
+   | SOME clauses => (comps, clauses))
 
 val add_mtype_is_concrete = add_notin_mtype_fv Minus
 val add_mtype_is_complete = add_notin_mtype_fv Plus
 
-val bool_from_minus = true
-
-fun bool_from_sign Plus = not bool_from_minus
-  | bool_from_sign Minus = bool_from_minus
-fun sign_from_bool b = if b = bool_from_minus then Minus else Plus
+val bool_table =
+  [(Gen, (false, false)),
+   (New, (false, true)),
+   (Fls, (true, false)),
+   (Tru, (true, true))]
 
-fun prop_for_literal (x, sn) =
-  (not (bool_from_sign sn) ? PropLogic.Not) (PropLogic.BoolVar x)
-fun prop_for_sign_atom_eq (S sn', sn) =
-    if sn = sn' then PropLogic.True else PropLogic.False
-  | prop_for_sign_atom_eq (V x, sn) = prop_for_literal (x, sn)
-fun prop_for_sign_expr xs = PropLogic.exists (map prop_for_literal xs)
-fun prop_for_exists_eq xs sn =
-  PropLogic.exists (map (fn x => prop_for_literal (x, sn)) xs)
-fun prop_for_comp (a1, a2, Eq, []) =
-    PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []),
-                    prop_for_comp (a2, a1, Leq, []))
-  | prop_for_comp (a1, a2, Leq, []) =
-    PropLogic.SOr (prop_for_sign_atom_eq (a1, Plus),
-                   prop_for_sign_atom_eq (a2, Minus))
-  | prop_for_comp (a1, a2, cmp, xs) =
-    PropLogic.SOr (prop_for_exists_eq xs Minus, prop_for_comp (a1, a2, cmp, []))
+fun fst_var n = 2 * n
+fun snd_var n = 2 * n + 1
 
-fun literals_from_assignments max_var assigns lits =
-  fold (fn x => fn accum =>
-           if AList.defined (op =) lits x then
-             accum
-           else case assigns x of
-             SOME b => (x, sign_from_bool b) :: accum
-           | NONE => accum) (max_var downto 1) lits
+val bools_from_annotation = AList.lookup (op =) bool_table #> the
+val annotation_from_bools = AList.find (op =) bool_table #> the_single
 
-fun string_for_comp (a1, a2, cmp, xs) =
-  string_for_sign_atom a1 ^ " " ^ string_for_comp_op cmp ^
-  subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_sign_atom a2
-
-fun print_problem lits comps sexps =
-  trace_msg (fn () => "*** Problem:\n" ^
-                      cat_lines (map string_for_literal lits @
-                                 map string_for_comp comps @
-                                 map string_for_sign_expr sexps))
+fun prop_for_bool b = if b then PL.True else PL.False
+fun prop_for_bool_var_equality (v1, v2) =
+  PL.SAnd (PL.SOr (PL.BoolVar v1, PL.SNot (PL.BoolVar v2)),
+           PL.SOr (PL.SNot (PL.BoolVar v1), PL.BoolVar v2))
+fun prop_for_assign (x, a) =
+  let val (b1, b2) = bools_from_annotation a in
+    PL.SAnd (PL.BoolVar (fst_var x) |> not b1 ? PL.SNot,
+             PL.BoolVar (snd_var x) |> not b2 ? PL.SNot)
+  end
+fun prop_for_assign_literal (x, (Plus, a)) = prop_for_assign (x, a)
+  | prop_for_assign_literal (x, (Minus, a)) = PL.SNot (prop_for_assign (x, a))
+fun prop_for_atom_assign (A a', a) = prop_for_bool (a = a')
+  | prop_for_atom_assign (V x, a) = prop_for_assign_literal (x, (Plus, a))
+fun prop_for_atom_equality (aa1, A a2) = prop_for_atom_assign (aa1, a2)
+  | prop_for_atom_equality (A a1, aa2) = prop_for_atom_assign (aa2, a1)
+  | prop_for_atom_equality (V x1, V x2) =
+    PL.SAnd (prop_for_bool_var_equality (pairself fst_var (x1, x2)),
+             prop_for_bool_var_equality (pairself snd_var (x1, x2)))
+val prop_for_assign_clause = PL.exists o map prop_for_assign_literal
+fun prop_for_exists_var_assign_literal xs a =
+  PL.exists (map (fn x => prop_for_assign_literal (x, (Plus, a))) xs)
+fun prop_for_comp (aa1, aa2, Eq, []) =
+    PL.SAnd (prop_for_comp (aa1, aa2, Leq, []),
+             prop_for_comp (aa2, aa1, Leq, []))
+  | prop_for_comp (aa1, aa2, Neq, []) =
+    PL.SNot (prop_for_comp (aa1, aa2, Eq, []))
+  | prop_for_comp (aa1, aa2, Leq, []) =
+    PL.SOr (prop_for_atom_equality (aa1, aa2), prop_for_atom_assign (aa2, Gen))
+  | prop_for_comp (aa1, aa2, cmp, xs) =
+    PL.SOr (prop_for_exists_var_assign_literal xs Gen,
+            prop_for_comp (aa1, aa2, cmp, []))
 
-fun print_solution lits =
-  let val (pos, neg) = List.partition (curry (op =) Plus o snd) lits in
-    trace_msg (fn () => "*** Solution:\n" ^
-                        "+: " ^ commas (map (string_for_var o fst) pos) ^ "\n" ^
-                        "-: " ^ commas (map (string_for_var o fst) neg))
-  end
+fun extract_assigns max_var assigns asgs =
+  fold (fn x => fn accum =>
+           if AList.defined (op =) asgs x then
+             accum
+           else case (fst_var x, snd_var x) |> pairself assigns of
+             (NONE, NONE) => accum
+           | bp => (x, annotation_from_bools (pairself (the_default false) bp))
+                   :: accum)
+       (max_var downto 1) asgs
+
+fun print_problem comps clauses =
+  trace_msg (fn () => "*** Problem:\n" ^
+                      cat_lines (map string_for_comp comps @
+                                 map (string_for_assign_clause o SOME) clauses))
 
-fun solve max_var (lits, comps, sexps) =
+fun print_solution asgs =
+  trace_msg (fn () => "*** Solution:\n" ^
+      (asgs
+       |> map swap
+       |> AList.group (op =)
+       |> map (fn (a, xs) => string_for_annotation a ^ ": " ^
+                             string_for_vars ", " (sort int_ord xs))
+       |> space_implode "\n"))
+
+(* The ML solver timeout should correspond more or less to the overhead of
+   invoking an external prover. *)
+val ml_solver_timeout = SOME (seconds 0.02)
+
+fun solve tac_timeout max_var (comps, clauses) =
   let
+    val asgs =
+      map_filter (fn [(x, (Plus, a))] => SOME (x, a) | _ => NONE) clauses
     fun do_assigns assigns =
-      SOME (literals_from_assignments max_var assigns lits
-            |> tap print_solution)
-    val _ = print_problem lits comps sexps
-    val prop = PropLogic.all (map prop_for_literal lits @
-                              map prop_for_comp comps @
-                              map prop_for_sign_expr sexps)
-    val default_val = bool_from_sign Minus
+      SOME (extract_assigns max_var assigns asgs |> tap print_solution)
+    val _ = print_problem comps clauses
+    val prop =
+      PL.all (map prop_for_comp comps @ map prop_for_assign_clause clauses)
   in
-    if PropLogic.eval (K default_val) prop then
-      do_assigns (K (SOME default_val))
+    if PL.eval (K false) prop then
+      do_assigns (K (SOME false))
+    else if PL.eval (K true) prop then
+      do_assigns (K (SOME true))
     else
       let
         (* use the first ML solver (to avoid startup overhead) *)
-        val solvers = !SatSolver.solvers
-                      |> filter (member (op =) ["dptsat", "dpll"] o fst)
+        val (ml_solvers, nonml_solvers) =
+          !SatSolver.solvers
+          |> List.partition (member (op =) ["dptsat", "dpll"] o fst)
+        val res =
+          if null nonml_solvers then
+            time_limit tac_timeout (snd (hd ml_solvers)) prop
+          else
+            time_limit ml_solver_timeout (snd (hd ml_solvers)) prop
+            handle TimeLimit.TimeOut =>
+                   time_limit tac_timeout (SatSolver.invoke_solver "auto") prop
       in
-        case snd (hd solvers) prop of
+        case res of
           SatSolver.SATISFIABLE assigns => do_assigns assigns
-        | _ => NONE
+        | _ => (trace_msg (K "*** Unsolvable"); NONE)
       end
+      handle TimeLimit.TimeOut => (trace_msg (K "*** Timed out"); NONE)
   end
 
-type mtype_schema = mtyp * constraint_set
-type mtype_context =
+type mcontext =
   {bound_Ts: typ list,
    bound_Ms: mtyp list,
+   frame: (int * annotation_atom) list,
    frees: (styp * mtyp) list,
    consts: (styp * mtyp) list}
 
-type accumulator = mtype_context * constraint_set
+fun string_for_bound ctxt Ms (j, aa) =
+  Syntax.string_of_term ctxt (Bound (length Ms - j - 1)) ^ " :\<^bsup>" ^
+  string_for_annotation_atom aa ^ "\<^esup> " ^
+  string_for_mtype (nth Ms (length Ms - j - 1))
+fun string_for_free relevant_frees ((s, _), M) =
+  if member (op =) relevant_frees s then SOME (s ^ " : " ^ string_for_mtype M)
+  else NONE
+fun string_for_mcontext ctxt t {bound_Ms, frame, frees, ...} =
+  (map_filter (string_for_free (Term.add_free_names t [])) frees @
+   map (string_for_bound ctxt bound_Ms) frame)
+  |> commas |> enclose "[" "]"
+
+val initial_gamma =
+  {bound_Ts = [], bound_Ms = [], frame = [], frees = [], consts = []}
+
+fun push_bound aa T M {bound_Ts, bound_Ms, frame, frees, consts} =
+  {bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms,
+   frame = frame @ [(length bound_Ts, aa)], frees = frees, consts = consts}
+fun pop_bound {bound_Ts, bound_Ms, frame, frees, consts} =
+  {bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms,
+   frame = frame |> filter_out (fn (j, _) => j = length bound_Ts - 1),
+   frees = frees, consts = consts}
+  handle List.Empty => initial_gamma (* FIXME: needed? *)
+
+fun set_frame frame ({bound_Ts, bound_Ms, frees, consts, ...} : mcontext) =
+  {bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame, frees = frees,
+   consts = consts}
+
+(* FIXME: make sure tracing messages are complete *)
+
+fun add_comp_frame aa cmp = fold (add_annotation_atom_comp cmp [] aa o snd)
 
-val initial_gamma = {bound_Ts = [], bound_Ms = [], frees = [], consts = []}
+fun add_bound_frame j frame =
+  let
+    val (new_frame, gen_frame) = List.partition (curry (op =) j o fst) frame
+  in
+    add_comp_frame (A New) Leq new_frame
+    #> add_comp_frame (A Gen) Eq gen_frame
+  end
+
+fun fresh_frame ({max_fresh, ...} : mdata) fls tru =
+  map (apsnd (fn aa =>
+                 case (aa, fls, tru) of
+                   (A Fls, SOME a, _) => A a
+                 | (A Tru, _, SOME a) => A a
+                 | (A Gen, _, _) => A Gen
+                 | _ => V (Unsynchronized.inc max_fresh)))
+
+fun conj_clauses res_aa aa1 aa2 =
+  [[(aa1, (Neq, Tru)), (aa2, (Neq, Tru)), (res_aa, (Eq, Tru))],
+   [(aa1, (Neq, Fls)), (res_aa, (Eq, Fls))],
+   [(aa2, (Neq, Fls)), (res_aa, (Eq, Fls))],
+   [(aa1, (Neq, Gen)), (aa2, (Eq, Fls)), (res_aa, (Eq, Gen))],
+   [(aa1, (Neq, New)), (aa2, (Eq, Fls)), (res_aa, (Eq, Gen))],
+   [(aa1, (Eq, Fls)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))],
+   [(aa1, (Eq, Fls)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]]
+
+fun disj_clauses res_aa aa1 aa2 =
+  [[(aa1, (Neq, Tru)), (res_aa, (Eq, Tru))],
+   [(aa2, (Neq, Tru)), (res_aa, (Eq, Tru))],
+   [(aa1, (Neq, Fls)), (aa2, (Neq, Fls)), (res_aa, (Eq, Fls))],
+   [(aa1, (Neq, Gen)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+   [(aa1, (Neq, New)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+   [(aa1, (Eq, Tru)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))],
+   [(aa1, (Eq, Tru)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]]
+
+fun imp_clauses res_aa aa1 aa2 =
+  [[(aa1, (Neq, Fls)), (res_aa, (Eq, Tru))],
+   [(aa2, (Neq, Tru)), (res_aa, (Eq, Tru))],
+   [(aa1, (Neq, Tru)), (aa2, (Neq, Fls)), (res_aa, (Eq, Fls))],
+   [(aa1, (Neq, Gen)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+   [(aa1, (Neq, New)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))],
+   [(aa1, (Eq, Fls)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))],
+   [(aa1, (Eq, Fls)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]]
 
-fun push_bound T M {bound_Ts, bound_Ms, frees, consts} =
-  {bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms, frees = frees,
-   consts = consts}
-fun pop_bound {bound_Ts, bound_Ms, frees, consts} =
-  {bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms, frees = frees,
-   consts = consts}
-  handle List.Empty => initial_gamma (* FIXME: needed? *)
+val meta_conj_triple = ("\<and>", conj_clauses, @{const Pure.conjunction})
+val meta_imp_triple = ("\<implies>", imp_clauses, @{const "==>"})
+val conj_triple = ("\<and>", conj_clauses, @{const conj})
+val disj_triple = ("\<or>", disj_clauses, @{const disj})
+val imp_triple = ("\<implies>", imp_clauses, @{const implies})
+
+fun add_annotation_clause_from_quasi_clause _ NONE = NONE
+  | add_annotation_clause_from_quasi_clause [] accum = accum
+  | add_annotation_clause_from_quasi_clause ((aa, (cmp, a)) :: rest) accum =
+    case aa of
+      A a' => if annotation_comp cmp a' a then NONE
+              else add_annotation_clause_from_quasi_clause rest accum
+    | V x => add_annotation_clause_from_quasi_clause rest accum
+             |> Option.map (cons (x, (sign_for_comp_op cmp, a)))
+
+fun assign_clause_from_quasi_clause unless =
+  add_annotation_clause_from_quasi_clause unless (SOME [])
+
+fun add_connective_var conn mk_quasi_clauses res_aa aa1 aa2 =
+  (trace_msg (fn () => "*** Add " ^ string_for_annotation_atom res_aa ^ " = " ^
+                       string_for_annotation_atom aa1 ^ " " ^ conn ^ " " ^
+                       string_for_annotation_atom aa2);
+   fold (add_assign_clause o assign_clause_from_quasi_clause)
+        (mk_quasi_clauses res_aa aa1 aa2))
+fun add_connective_frames conn mk_quasi_clauses res_frame frame1 frame2 =
+  fold I (map3 (fn (_, res_aa) => fn (_, aa1) => fn (_, aa2) =>
+                   add_connective_var conn mk_quasi_clauses res_aa aa1 aa2)
+               res_frame frame1 frame2)
+
+fun kill_unused_in_frame is_in (accum as ({frame, ...}, _)) =
+  let val (used_frame, unused_frame) = List.partition is_in frame in
+    accum |>> set_frame used_frame
+          ||> add_comp_frame (A Gen) Eq unused_frame
+  end
+
+fun split_frame is_in_fun (gamma as {frame, ...}, cset) =
+  let
+    fun bubble fun_frame arg_frame [] cset =
+        ((rev fun_frame, rev arg_frame), cset)
+      | bubble fun_frame arg_frame ((bound as (_, aa)) :: rest) cset =
+        if is_in_fun bound then
+          bubble (bound :: fun_frame) arg_frame rest
+                 (cset |> add_comp_frame aa Leq arg_frame)
+        else
+          bubble fun_frame (bound :: arg_frame) rest cset
+  in cset |> bubble [] [] frame ||> pair gamma end
+
+fun add_annotation_atom_comp_alt _ (A Gen) _ _ = I
+  | add_annotation_atom_comp_alt _ (A _) _ _ =
+    (trace_msg (K "*** Expected G"); raise UNSOLVABLE ())
+  | add_annotation_atom_comp_alt cmp (V x) aa1 aa2 =
+    add_annotation_atom_comp cmp [x] aa1 aa2
+
+fun add_arg_order1 ((_, aa), (_, prev_aa)) = I
+  add_annotation_atom_comp_alt Neq prev_aa (A Gen) aa
+fun add_app1 fun_aa ((_, res_aa), (_, arg_aa)) = I
+  let
+    val clause = [(arg_aa, (Eq, New)), (res_aa, (Eq, Gen))]
+                 |> assign_clause_from_quasi_clause
+  in
+    trace_msg (fn () => "*** Add " ^ string_for_assign_clause clause);
+    apsnd (add_assign_clause clause)
+    #> add_annotation_atom_comp_alt Leq arg_aa fun_aa res_aa
+  end
+fun add_app _ [] [] = I
+  | add_app fun_aa res_frame arg_frame =
+    add_comp_frame (A New) Leq arg_frame
+    #> fold add_arg_order1 (tl arg_frame ~~ (fst (split_last arg_frame)))
+    #> fold (add_app1 fun_aa) (res_frame ~~ arg_frame)
+
+fun consider_connective mdata (conn, mk_quasi_clauses, t0) do_t1 do_t2
+                        (accum as ({frame, ...}, _)) =
+  let
+    val mtype_for = fresh_mtype_for_type mdata false
+    val frame1 = fresh_frame mdata (SOME Tru) NONE frame
+    val frame2 = fresh_frame mdata (SOME Fls) NONE frame
+    val (m1, accum) = accum |>> set_frame frame1 |> do_t1
+    val (m2, accum) = accum |>> set_frame frame2 |> do_t2
+  in
+    (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
+     accum |>> set_frame frame
+           ||> apsnd (add_connective_frames conn mk_quasi_clauses frame frame1
+                                            frame2))
+  end
 
 fun consider_term (mdata as {hol_ctxt = {thy, ctxt, stds, ...}, alpha_T,
                              max_fresh, ...}) =
@@ -557,21 +788,23 @@
         <= length ts
       | _ => true
     val mtype_for = fresh_mtype_for_type mdata false
-    fun plus_set_mtype_for_dom M =
-      MFun (M, S (if exists_alpha_sub_mtype M then Plus else Minus), bool_M)
     fun do_all T (gamma, cset) =
       let
         val abs_M = mtype_for (domain_type (domain_type T))
+        val x = Unsynchronized.inc max_fresh
         val body_M = mtype_for (body_type T)
       in
-        (MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M),
-         (gamma, cset |> add_mtype_is_complete abs_M))
+        (MFun (MFun (abs_M, V x, body_M), A Gen, body_M),
+         (gamma, cset |> add_mtype_is_complete [(x, (Plus, Tru))] abs_M))
       end
     fun do_equals T (gamma, cset) =
-      let val M = mtype_for (domain_type T) in
-        (MFun (M, S Minus, MFun (M, V (Unsynchronized.inc max_fresh),
-                                 mtype_for (nth_range_type 2 T))),
-         (gamma, cset |> add_mtype_is_concrete M))
+      let
+        val M = mtype_for (domain_type T)
+        val x = Unsynchronized.inc max_fresh
+      in
+        (MFun (M, A Gen, MFun (M, V x, mtype_for (nth_range_type 2 T))),
+         (gamma, cset |> add_mtype_is_concrete [] M
+                      |> add_annotation_atom_comp Leq [] (A Fls) (V x)))
       end
     fun do_robust_set_operation T (gamma, cset) =
       let
@@ -580,7 +813,7 @@
         val M2 = mtype_for set_T
         val M3 = mtype_for set_T
       in
-        (MFun (M1, S Minus, MFun (M2, S Minus, M3)),
+        (MFun (M1, A Gen, MFun (M2, A Gen, M3)),
          (gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3))
       end
     fun do_fragile_set_operation T (gamma, cset) =
@@ -589,164 +822,175 @@
         val set_M = mtype_for set_T
         fun custom_mtype_for (T as Type (@{type_name fun}, [T1, T2])) =
             if T = set_T then set_M
-            else MFun (custom_mtype_for T1, S Minus, custom_mtype_for T2)
+            else MFun (custom_mtype_for T1, A Gen, custom_mtype_for T2)
           | custom_mtype_for T = mtype_for T
       in
-        (custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete set_M))
+        (custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete [] set_M))
       end
     fun do_pair_constr T accum =
       case mtype_for (nth_range_type 2 T) of
         M as MPair (a_M, b_M) =>
-        (MFun (a_M, S Minus, MFun (b_M, S Minus, M)), accum)
+        (MFun (a_M, A Gen, MFun (b_M, A Gen, M)), accum)
       | M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M], [])
     fun do_nth_pair_sel n T =
       case mtype_for (domain_type T) of
         M as MPair (a_M, b_M) =>
-        pair (MFun (M, S Minus, if n = 0 then a_M else b_M))
+        pair (MFun (M, A Gen, if n = 0 then a_M else b_M))
       | M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M], [])
-    fun do_bounded_quantifier t0 abs_s abs_T connective_t bound_t body_t accum =
-      let
-        val abs_M = mtype_for abs_T
-        val (bound_m, accum) =
-          accum |>> push_bound abs_T abs_M |> do_term bound_t
-        val expected_bound_M = plus_set_mtype_for_dom abs_M
-        val (body_m, accum) =
-          accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m)
-                |> do_term body_t ||> apfst pop_bound
-        val bound_M = mtype_of_mterm bound_m
-        val (M1, a, _) = dest_MFun bound_M
-      in
-        (MApp (MRaw (t0, MFun (bound_M, S Minus, bool_M)),
-               MAbs (abs_s, abs_T, M1, a,
-                     MApp (MApp (MRaw (connective_t,
-                                       mtype_for (fastype_of connective_t)),
-                                 MApp (bound_m, MRaw (Bound 0, M1))),
-                           body_m))), accum)
-      end
-    and do_term t (accum as ({bound_Ts, bound_Ms, frees, consts}, cset)) =
-        (trace_msg (fn () => "  \<Gamma> \<turnstile> " ^
-                             Syntax.string_of_term ctxt t ^ " : _?");
-         case t of
-           Const (x as (s, T)) =>
-           (case AList.lookup (op =) consts x of
-              SOME M => (M, accum)
-            | NONE =>
-              if not (could_exist_alpha_subtype alpha_T T) then
-                (mtype_for T, accum)
-              else case s of
-                @{const_name all} => do_all T accum
-              | @{const_name "=="} => do_equals T accum
-              | @{const_name All} => do_all T accum
-              | @{const_name Ex} =>
-                let val set_T = domain_type T in
-                  do_term (Abs (Name.uu, set_T,
-                                @{const Not} $ (HOLogic.mk_eq
-                                    (Abs (Name.uu, domain_type set_T,
-                                          @{const False}),
-                                     Bound 0)))) accum
-                  |>> mtype_of_mterm
-                end
-              | @{const_name HOL.eq} => do_equals T accum
-              | @{const_name The} =>
-                (trace_msg (K "*** The"); raise UNSOLVABLE ())
-              | @{const_name Eps} =>
-                (trace_msg (K "*** Eps"); raise UNSOLVABLE ())
-              | @{const_name If} =>
-                do_robust_set_operation (range_type T) accum
-                |>> curry3 MFun bool_M (S Minus)
-              | @{const_name Pair} => do_pair_constr T accum
-              | @{const_name fst} => do_nth_pair_sel 0 T accum
-              | @{const_name snd} => do_nth_pair_sel 1 T accum
-              | @{const_name Id} =>
-                (MFun (mtype_for (domain_type T), S Minus, bool_M), accum)
-              | @{const_name converse} =>
+    and do_connect triple t1 t2 =
+      consider_connective mdata triple (do_term t1) (do_term t2)
+    and do_term t
+            (accum as (gamma as {bound_Ts, bound_Ms, frame, frees, consts},
+                       cset)) =
+      (trace_msg (fn () => "  " ^ string_for_mcontext ctxt t gamma ^
+                           " \<turnstile> " ^ Syntax.string_of_term ctxt t ^
+                           " : _?");
+       case t of
+         @{const False} =>
+         (MRaw (t, bool_M), accum ||> add_comp_frame (A Fls) Leq frame)
+       | Const (@{const_name None}, T) =>
+         (MRaw (t, mtype_for T), accum ||> add_comp_frame (A Fls) Leq frame)
+       | @{const True} =>
+         (MRaw (t, bool_M), accum ||> add_comp_frame (A Tru) Leq frame)
+       | (t0 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t2 =>
+         (* hack to exploit symmetry of equality when typing "insert" *)
+         (if t2 = Bound 0 then do_term @{term True}
+          else do_term (t0 $ t2 $ Bound 0)) accum
+       | Const (x as (s, T)) =>
+         (case AList.lookup (op =) consts x of
+            SOME M => (M, accum)
+          | NONE =>
+            if not (could_exist_alpha_subtype alpha_T T) then
+              (mtype_for T, accum)
+            else case s of
+              @{const_name all} => do_all T accum
+            | @{const_name "=="} => do_equals T accum
+            | @{const_name All} => do_all T accum
+            | @{const_name Ex} =>
+              let val set_T = domain_type T in
+                do_term (Abs (Name.uu, set_T,
+                              @{const Not} $ (HOLogic.mk_eq
+                                  (Abs (Name.uu, domain_type set_T,
+                                        @{const False}),
+                                   Bound 0)))) accum
+                |>> mtype_of_mterm
+              end
+            | @{const_name HOL.eq} => do_equals T accum
+            | @{const_name The} =>
+              (trace_msg (K "*** The"); raise UNSOLVABLE ())
+            | @{const_name Eps} =>
+              (trace_msg (K "*** Eps"); raise UNSOLVABLE ())
+            | @{const_name If} =>
+              do_robust_set_operation (range_type T) accum
+              |>> curry3 MFun bool_M (A Gen)
+            | @{const_name Pair} => do_pair_constr T accum
+            | @{const_name fst} => do_nth_pair_sel 0 T accum
+            | @{const_name snd} => do_nth_pair_sel 1 T accum
+            | @{const_name Id} =>
+              (MFun (mtype_for (domain_type T), A Gen, bool_M), accum)
+            | @{const_name converse} =>
+              let
+                val x = Unsynchronized.inc max_fresh
+                fun mtype_for_set T =
+                  MFun (mtype_for (domain_type T), V x, bool_M)
+                val ab_set_M = domain_type T |> mtype_for_set
+                val ba_set_M = range_type T |> mtype_for_set
+              in
+                (MFun (ab_set_M, A Gen, ba_set_M),
+                 accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+              end
+            | @{const_name trancl} => do_fragile_set_operation T accum
+            | @{const_name rel_comp} =>
+              let
+                val x = Unsynchronized.inc max_fresh
+                fun mtype_for_set T =
+                  MFun (mtype_for (domain_type T), V x, bool_M)
+                val bc_set_M = domain_type T |> mtype_for_set
+                val ab_set_M = domain_type (range_type T) |> mtype_for_set
+                val ac_set_M = nth_range_type 2 T |> mtype_for_set
+              in
+                (MFun (bc_set_M, A Gen, MFun (ab_set_M, A Gen, ac_set_M)),
+                 accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+              end
+            | @{const_name image} =>
+              let
+                val x = Unsynchronized.inc max_fresh
+                val a_M = mtype_for (domain_type (domain_type T))
+                val b_M = mtype_for (range_type (domain_type T))
+              in
+                (MFun (MFun (a_M, A Gen, b_M), A Gen,
+                       MFun (MFun (a_M, V x, bool_M), A Gen,
+                             MFun (b_M, V x, bool_M))),
+                 accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+              end
+            | @{const_name finite} =>
+              let
+                val M1 = mtype_for (domain_type (domain_type T))
+                val a = if exists_alpha_sub_mtype M1 then Fls else Gen
+              in (MFun (MFun (M1, A a, bool_M), A Gen, bool_M), accum) end
+            | @{const_name Sigma} =>
+              let
+                val x = Unsynchronized.inc max_fresh
+                fun mtype_for_set T =
+                  MFun (mtype_for (domain_type T), V x, bool_M)
+                val a_set_T = domain_type T
+                val a_M = mtype_for (domain_type a_set_T)
+                val b_set_M =
+                  mtype_for_set (range_type (domain_type (range_type T)))
+                val a_set_M = mtype_for_set a_set_T
+                val a_to_b_set_M = MFun (a_M, A Gen, b_set_M)
+                val ab_set_M = mtype_for_set (nth_range_type 2 T)
+              in
+                (MFun (a_set_M, A Gen, MFun (a_to_b_set_M, A Gen, ab_set_M)),
+                 accum ||> add_annotation_atom_comp Neq [] (V x) (A New))
+              end
+            | _ =>
+              if s = @{const_name safe_The} then
                 let
-                  val x = Unsynchronized.inc max_fresh
-                  fun mtype_for_set T =
-                    MFun (mtype_for (domain_type T), V x, bool_M)
-                  val ab_set_M = domain_type T |> mtype_for_set
-                  val ba_set_M = range_type T |> mtype_for_set
-                in (MFun (ab_set_M, S Minus, ba_set_M), accum) end
-              | @{const_name trancl} => do_fragile_set_operation T accum
-              | @{const_name rel_comp} =>
-                let
-                  val x = Unsynchronized.inc max_fresh
-                  fun mtype_for_set T =
-                    MFun (mtype_for (domain_type T), V x, bool_M)
-                  val bc_set_M = domain_type T |> mtype_for_set
-                  val ab_set_M = domain_type (range_type T) |> mtype_for_set
-                  val ac_set_M = nth_range_type 2 T |> mtype_for_set
-                in
-                  (MFun (bc_set_M, S Minus, MFun (ab_set_M, S Minus, ac_set_M)),
-                   accum)
-                end
-              | @{const_name image} =>
-                let
-                  val a_M = mtype_for (domain_type (domain_type T))
-                  val b_M = mtype_for (range_type (domain_type T))
-                in
-                  (MFun (MFun (a_M, S Minus, b_M), S Minus,
-                         MFun (plus_set_mtype_for_dom a_M, S Minus,
-                               plus_set_mtype_for_dom b_M)), accum)
-                end
-              | @{const_name finite} =>
-                let val M1 = mtype_for (domain_type (domain_type T)) in
-                  (MFun (plus_set_mtype_for_dom M1, S Minus, bool_M), accum)
-                end
-              | @{const_name Sigma} =>
-                let
-                  val x = Unsynchronized.inc max_fresh
-                  fun mtype_for_set T =
-                    MFun (mtype_for (domain_type T), V x, bool_M)
-                  val a_set_T = domain_type T
-                  val a_M = mtype_for (domain_type a_set_T)
-                  val b_set_M = mtype_for_set (range_type (domain_type
-                                                               (range_type T)))
-                  val a_set_M = mtype_for_set a_set_T
-                  val a_to_b_set_M = MFun (a_M, S Minus, b_set_M)
-                  val ab_set_M = mtype_for_set (nth_range_type 2 T)
-                in
-                  (MFun (a_set_M, S Minus,
-                         MFun (a_to_b_set_M, S Minus, ab_set_M)), accum)
-                end
-              | _ =>
-                if s = @{const_name safe_The} then
-                  let
-                    val a_set_M = mtype_for (domain_type T)
-                    val a_M = dest_MFun a_set_M |> #1
-                  in (MFun (a_set_M, S Minus, a_M), accum) end
-                else if s = @{const_name ord_class.less_eq} andalso
-                        is_set_type (domain_type T) then
-                  do_fragile_set_operation T accum
-                else if is_sel s then
-                  (mtype_for_sel mdata x, accum)
-                else if is_constr ctxt stds x then
-                  (mtype_for_constr mdata x, accum)
-                else if is_built_in_const thy stds x then
-                  (fresh_mtype_for_type mdata true T, accum)
-                else
-                  let val M = mtype_for T in
-                    (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
-                          frees = frees, consts = (x, M) :: consts}, cset))
-                  end) |>> curry MRaw t
+                  val a_set_M = mtype_for (domain_type T)
+                  val a_M = dest_MFun a_set_M |> #1
+                in (MFun (a_set_M, A Gen, a_M), accum) end
+              else if s = @{const_name ord_class.less_eq} andalso
+                      is_set_type (domain_type T) then
+                do_fragile_set_operation T accum
+              else if is_sel s then
+                (mtype_for_sel mdata x, accum)
+              else if is_constr ctxt stds x then
+                (mtype_for_constr mdata x, accum)
+              else if is_built_in_const thy stds x then
+                (fresh_mtype_for_type mdata true T, accum)
+              else
+                let val M = mtype_for T in
+                  (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame,
+                        frees = frees, consts = (x, M) :: consts}, cset))
+                end)
+           |>> curry MRaw t
+           ||> apsnd (add_comp_frame (A Gen) Eq frame)
          | Free (x as (_, T)) =>
            (case AList.lookup (op =) frees x of
               SOME M => (M, accum)
             | NONE =>
               let val M = mtype_for T in
-                (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms,
+                (M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame,
                       frees = (x, M) :: frees, consts = consts}, cset))
-              end) |>> curry MRaw t
+              end)
+           |>> curry MRaw t ||> apsnd (add_comp_frame (A Gen) Eq frame)
          | Var _ => (trace_msg (K "*** Var"); raise UNSOLVABLE ())
-         | Bound j => (MRaw (t, nth bound_Ms j), accum)
+         | Bound j =>
+           (MRaw (t, nth bound_Ms j),
+            accum ||> add_bound_frame (length bound_Ts - j - 1) frame)
          | Abs (s, T, t') =>
            (case fin_fun_body T (fastype_of1 (T :: bound_Ts, t')) t' of
               SOME t' =>
               let
                 val M = mtype_for T
-                val a = V (Unsynchronized.inc max_fresh)
-                val (m', accum) = do_term t' (accum |>> push_bound T M)
-              in (MAbs (s, T, M, a, m'), accum |>> pop_bound) end
+                val x = Unsynchronized.inc max_fresh
+                val (m', accum) = do_term t' (accum |>> push_bound (V x) T M)
+              in
+                (MAbs (s, T, M, V x, m'),
+                 accum |>> pop_bound
+                       ||> add_annotation_atom_comp Leq [] (A Fls) (V x))
+              end
             | NONE =>
               ((case t' of
                   t1' $ Bound 0 =>
@@ -764,38 +1008,49 @@
                handle SAME () =>
                       let
                         val M = mtype_for T
-                        val (m', accum) = do_term t' (accum |>> push_bound T M)
-                      in
-                        (MAbs (s, T, M, S Minus, m'), accum |>> pop_bound)
-                      end))
-         | (t0 as Const (@{const_name All}, _))
-           $ Abs (s', T', (t10 as @{const HOL.implies}) $ (t11 $ Bound 0) $ t12) =>
-           do_bounded_quantifier t0 s' T' t10 t11 t12 accum
-         | (t0 as Const (@{const_name Ex}, _))
-           $ Abs (s', T', (t10 as @{const HOL.conj}) $ (t11 $ Bound 0) $ t12) =>
-           do_bounded_quantifier t0 s' T' t10 t11 t12 accum
+                        val x = Unsynchronized.inc max_fresh
+                        val (m', accum) =
+                          do_term t' (accum |>> push_bound (V x) T M)
+                      in (MAbs (s, T, M, V x, m'), accum |>> pop_bound) end))
+         | @{const Not} $ t1 => do_connect imp_triple t1 @{const False} accum
+         | @{const conj} $ t1 $ t2 => do_connect conj_triple t1 t2 accum
+         | @{const disj} $ t1 $ t2 => do_connect disj_triple t1 t2 accum
+         | @{const implies} $ t1 $ t2 => do_connect imp_triple t1 t2 accum
          | Const (@{const_name Let}, _) $ t1 $ t2 =>
            do_term (betapply (t2, t1)) accum
          | t1 $ t2 =>
            let
-             val (m1, accum) = do_term t1 accum
-             val (m2, accum) = do_term t2 accum
+             fun is_in t (j, _) = loose_bvar1 (t, length bound_Ts - j - 1)
+             val accum as ({frame, ...}, _) =
+               accum |> kill_unused_in_frame (is_in t)
+             val ((frame1a, frame1b), accum) = accum |> split_frame (is_in t1)
+             val frame2a = frame1a |> map (apsnd (K (A Gen)))
+             val frame2b =
+               frame1b |> map (apsnd (fn _ => V (Unsynchronized.inc max_fresh)))
+             val frame2 = frame2a @ frame2b
+             val (m1, accum) = accum |>> set_frame frame1a |> do_term t1
+             val (m2, accum) = accum |>> set_frame frame2 |> do_term t2
            in
              let
-               val M11 = mtype_of_mterm m1 |> dest_MFun |> #1
+               val (M11, aa, _) = mtype_of_mterm m1 |> dest_MFun
                val M2 = mtype_of_mterm m2
-             in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end
+             in
+               (MApp (m1, m2),
+                accum |>> set_frame frame
+                      ||> add_is_sub_mtype M2 M11
+                      ||> add_app aa frame1b frame2b)
+             end
            end)
-        |> tap (fn (m, _) => trace_msg (fn () => "  \<Gamma> \<turnstile> " ^
-                                                 string_for_mterm ctxt m))
+        |> tap (fn (m, (gamma, _)) =>
+                   trace_msg (fn () => "  " ^ string_for_mcontext ctxt t gamma ^
+                                       " \<turnstile> " ^
+                                       string_for_mterm ctxt m))
   in do_term end
 
-fun force_minus_funs 0 _ = I
-  | force_minus_funs n (M as MFun (M1, _, M2)) =
-    add_mtypes_equal M (MFun (M1, S Minus, M2))
-    #> force_minus_funs (n - 1) M2
-  | force_minus_funs _ M =
-    raise MTYPE ("Nitpick_Mono.force_minus_funs", [M], [])
+fun force_gen_funs 0 _ = I
+  | force_gen_funs n (M as MFun (M1, _, M2)) =
+    add_mtypes_equal M (MFun (M1, A Gen, M2)) #> force_gen_funs (n - 1) M2
+  | force_gen_funs _ M = raise MTYPE ("Nitpick_Mono.force_gen_funs", [M], [])
 fun consider_general_equals mdata def (x as (_, T)) t1 t2 accum =
   let
     val (m1, accum) = consider_term mdata t1 accum
@@ -805,97 +1060,89 @@
     val accum = accum ||> add_mtypes_equal M1 M2
     val body_M = fresh_mtype_for_type mdata false (nth_range_type 2 T)
     val m = MApp (MApp (MRaw (Const x,
-                MFun (M1, S Minus, MFun (M2, S Minus, body_M))), m1), m2)
+                           MFun (M1, A Gen, MFun (M2, A Gen, body_M))), m1), m2)
   in
-    (m, if def then
-          let val (head_m, arg_ms) = strip_mcomb m1 in
-            accum ||> force_minus_funs (length arg_ms) (mtype_of_mterm head_m)
-          end
-        else
-          accum)
+    (m, (if def then
+           let val (head_m, arg_ms) = strip_mcomb m1 in
+             accum ||> force_gen_funs (length arg_ms) (mtype_of_mterm head_m)
+           end
+         else
+           accum))
   end
 
-fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) =
+fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, max_fresh,
+                                        ...}) =
   let
     val mtype_for = fresh_mtype_for_type mdata false
     val do_term = consider_term mdata
-    fun do_formula sn t accum =
-        let
-          fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
-            let
-              val abs_M = mtype_for abs_T
-              val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
-              val (body_m, accum) =
-                accum ||> side_cond ? add_mtype_is_complete abs_M
-                      |>> push_bound abs_T abs_M |> do_formula sn body_t
-              val body_M = mtype_of_mterm body_m
-            in
-              (MApp (MRaw (Const quant_x,
-                           MFun (MFun (abs_M, S Minus, body_M), S Minus,
-                                 body_M)),
-                     MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
-               accum |>> pop_bound)
-            end
-          fun do_equals x t1 t2 =
-            case sn of
-              Plus => do_term t accum
-            | Minus => consider_general_equals mdata false x t1 t2 accum
-        in
-          (trace_msg (fn () => "  \<Gamma> \<turnstile> " ^
-                               Syntax.string_of_term ctxt t ^ " : o\<^sup>" ^
-                               string_for_sign sn ^ "?");
-           case t of
-             Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
-             do_quantifier x s1 T1 t1
-           | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
-           | @{const Trueprop} $ t1 =>
-             let val (m1, accum) = do_formula sn t1 accum in
-               (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
-                      m1), accum)
-             end
-           | @{const Not} $ t1 =>
-             let val (m1, accum) = do_formula (negate sn) t1 accum in
-               (MApp (MRaw (@{const Not}, mtype_for (bool_T --> bool_T)), m1),
-                accum)
-             end
-           | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
-             do_quantifier x s1 T1 t1
-           | Const (x0 as (@{const_name Ex}, T0))
-             $ (t1 as Abs (s1, T1, t1')) =>
-             (case sn of
-                Plus => do_quantifier x0 s1 T1 t1'
-              | Minus =>
-                (* FIXME: Move elsewhere *)
-                do_term (@{const Not}
-                         $ (HOLogic.eq_const (domain_type T0) $ t1
-                            $ Abs (Name.uu, T1, @{const False}))) accum)
-           | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
-             do_equals x t1 t2
-           | Const (@{const_name Let}, _) $ t1 $ t2 =>
-             do_formula sn (betapply (t2, t1)) accum
-           | (t0 as Const (s0, _)) $ t1 $ t2 =>
-             if s0 = @{const_name "==>"} orelse
-                s0 = @{const_name Pure.conjunction} orelse
-                s0 = @{const_name HOL.conj} orelse
-                s0 = @{const_name HOL.disj} orelse
-                s0 = @{const_name HOL.implies} then
-               let
-                 val impl = (s0 = @{const_name "==>"} orelse
-                             s0 = @{const_name HOL.implies})
-                 val (m1, accum) = do_formula (sn |> impl ? negate) t1 accum
-                 val (m2, accum) = do_formula sn t2 accum
-               in
-                 (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2),
-                  accum)
-               end
-             else
-               do_term t accum
-           | _ => do_term t accum)
-        end
-        |> tap (fn (m, _) =>
-                   trace_msg (fn () => "\<Gamma> \<turnstile> " ^
-                                       string_for_mterm ctxt m ^ " : o\<^sup>" ^
-                                       string_for_sign sn))
+    fun do_formula sn t (accum as (gamma, _)) =
+      let
+        fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t =
+          let
+            val abs_M = mtype_for abs_T
+            val x = Unsynchronized.inc max_fresh
+            val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex}))
+            fun ann () = if quant_s = @{const_name Ex} then Fls else Tru
+            val (body_m, accum) =
+              accum ||> side_cond
+                        ? add_mtype_is_complete [(x, (Plus, ann ()))] abs_M
+                    |>> push_bound (V x) abs_T abs_M |> do_formula sn body_t
+            val body_M = mtype_of_mterm body_m
+          in
+            (MApp (MRaw (Const quant_x,
+                         MFun (MFun (abs_M, A Gen, body_M), A Gen, body_M)),
+                   MAbs (abs_s, abs_T, abs_M, A Gen, body_m)),
+             accum |>> pop_bound)
+          end
+        fun do_connect triple neg1 t1 t2 =
+          consider_connective mdata triple
+              (do_formula (sn |> neg1 ? negate_sign) t1) (do_formula sn t2)
+        fun do_equals x t1 t2 =
+          case sn of
+            Plus => do_term t accum
+          | Minus => consider_general_equals mdata false x t1 t2 accum
+      in
+        trace_msg (fn () => "  " ^ string_for_mcontext ctxt t gamma ^
+                            " \<turnstile> " ^ Syntax.string_of_term ctxt t ^
+                            " : o\<^sup>" ^ string_for_sign sn ^ "?");
+        case t of
+          Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
+          do_quantifier x s1 T1 t1
+        | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2
+        | @{const Trueprop} $ t1 =>
+          let val (m1, accum) = do_formula sn t1 accum in
+            (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)), m1),
+             accum)
+          end
+        | Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) =>
+          do_quantifier x s1 T1 t1
+        | Const (x0 as (@{const_name Ex}, T0)) $ (t1 as Abs (s1, T1, t1')) =>
+          (case sn of
+             Plus => do_quantifier x0 s1 T1 t1'
+           | Minus =>
+             (* FIXME: Move elsewhere *)
+             do_term (@{const Not}
+                      $ (HOLogic.eq_const (domain_type T0) $ t1
+                         $ Abs (Name.uu, T1, @{const False}))) accum)
+        | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 => do_equals x t1 t2
+        | Const (@{const_name Let}, _) $ t1 $ t2 =>
+          do_formula sn (betapply (t2, t1)) accum
+        | @{const Pure.conjunction} $ t1 $ t2 =>
+          do_connect meta_conj_triple false t1 t2 accum
+        | @{const "==>"} $ t1 $ t2 =>
+          do_connect meta_imp_triple true t1 t2 accum
+        | @{const Not} $ t1 =>
+          do_connect imp_triple true t1 @{const False} accum
+        | @{const conj} $ t1 $ t2 => do_connect conj_triple false t1 t2 accum
+        | @{const disj} $ t1 $ t2 => do_connect disj_triple false t1 t2 accum
+        | @{const implies} $ t1 $ t2 => do_connect imp_triple true t1 t2 accum
+        | _ => do_term t accum
+      end
+      |> tap (fn (m, (gamma, _)) =>
+                 trace_msg (fn () => string_for_mcontext ctxt t gamma ^
+                                     " \<turnstile> " ^
+                                     string_for_mterm ctxt m ^ " : o\<^sup>" ^
+                                     string_for_sign sn))
   in do_formula end
 
 (* The harmless axiom optimization below is somewhat too aggressive in the face
@@ -928,12 +1175,12 @@
         let
           val abs_M = mtype_for abs_T
           val (body_m, accum) =
-            accum |>> push_bound abs_T abs_M |> do_formula body_t
+            accum |>> push_bound (A Gen) abs_T abs_M |> do_formula body_t
           val body_M = mtype_of_mterm body_m
         in
-          (MApp (MRaw (quant_t,
-                       MFun (MFun (abs_M, S Minus, body_M), S Minus, body_M)),
-                 MAbs (abs_s, abs_T, abs_M, S Minus, body_m)),
+          (MApp (MRaw (quant_t, MFun (MFun (abs_M, A Gen, body_M), A Gen,
+                       body_M)),
+                 MAbs (abs_s, abs_T, abs_M, A Gen, body_m)),
            accum |>> pop_bound)
         end
       and do_conjunction t0 t1 t2 accum =
@@ -951,50 +1198,50 @@
           (MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum)
         end
       and do_formula t accum =
-          case t of
-            (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
-            do_all t0 s1 T1 t1 accum
-          | @{const Trueprop} $ t1 =>
-            let val (m1, accum) = do_formula t1 accum in
-              (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)),
-                     m1), accum)
-            end
-          | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
-            consider_general_equals mdata true x t1 t2 accum
-          | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
-          | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
-            do_conjunction t0 t1 t2 accum
-          | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
-            do_all t0 s0 T1 t1 accum
-          | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
-            consider_general_equals mdata true x t1 t2 accum
-          | (t0 as @{const HOL.conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
-          | (t0 as @{const HOL.implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
-          | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
-                             \do_formula", [t])
+        case t of
+          (t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) =>
+          do_all t0 s1 T1 t1 accum
+        | @{const Trueprop} $ t1 =>
+          let val (m1, accum) = do_formula t1 accum in
+            (MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)), m1),
+             accum)
+          end
+        | Const (x as (@{const_name "=="}, _)) $ t1 $ t2 =>
+          consider_general_equals mdata true x t1 t2 accum
+        | (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+        | (t0 as @{const Pure.conjunction}) $ t1 $ t2 =>
+          do_conjunction t0 t1 t2 accum
+        | (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) =>
+          do_all t0 s0 T1 t1 accum
+        | Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 =>
+          consider_general_equals mdata true x t1 t2 accum
+        | (t0 as @{const conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum
+        | (t0 as @{const implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum
+        | _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\
+                           \do_formula", [t])
     in do_formula t end
 
-fun string_for_mtype_of_term ctxt lits t M =
-  Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype lits M)
+fun string_for_mtype_of_term ctxt asgs t M =
+  Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype asgs M)
 
-fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) =
+fun print_mcontext ctxt asgs ({frees, consts, ...} : mcontext) =
   trace_msg (fn () =>
-      map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @
-      map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts
+      map (fn (x, M) => string_for_mtype_of_term ctxt asgs (Free x) M) frees @
+      map (fn (x, M) => string_for_mtype_of_term ctxt asgs (Const x) M) consts
       |> cat_lines)
 
 fun amass f t (ms, accum) =
   let val (m, accum) = f t accum in (m :: ms, accum) end
 
-fun infer which no_harmless (hol_ctxt as {ctxt, ...}) binarize alpha_T
-          (nondef_ts, def_ts) =
+fun infer which no_harmless (hol_ctxt as {ctxt, tac_timeout, ...}) binarize
+          alpha_T (nondef_ts, def_ts) =
   let
     val _ = trace_msg (fn () => "****** " ^ which ^ " analysis: " ^
                                 string_for_mtype MAlpha ^ " is " ^
                                 Syntax.string_of_typ ctxt alpha_T)
     val mdata as {max_fresh, constr_mcache, ...} =
       initial_mdata hol_ctxt binarize no_harmless alpha_T
-    val accum = (initial_gamma, ([], [], []))
+    val accum = (initial_gamma, ([], []))
     val (nondef_ms, accum) =
       ([], accum) |> amass (consider_general_formula mdata Plus) (hd nondef_ts)
                   |> fold (amass (consider_nondefinitional_axiom mdata))
@@ -1002,9 +1249,9 @@
     val (def_ms, (gamma, cset)) =
       ([], accum) |> fold (amass (consider_definitional_axiom mdata)) def_ts
   in
-    case solve (!max_fresh) cset of
-      SOME lits => (print_mtype_context ctxt lits gamma;
-                    SOME (lits, (nondef_ms, def_ms), !constr_mcache))
+    case solve tac_timeout (!max_fresh) cset of
+      SOME asgs => (print_mcontext ctxt asgs gamma;
+                    SOME (asgs, (nondef_ms, def_ms), !constr_mcache))
     | _ => NONE
   end
   handle UNSOLVABLE () => NONE
@@ -1019,23 +1266,23 @@
 fun fin_fun_constr T1 T2 =
   (@{const_name FinFun}, (T1 --> T2) --> Type (@{type_name fin_fun}, [T1, T2]))
 
-fun finitize_funs (hol_ctxt as {thy, ctxt, stds, constr_cache, ...})
-                  binarize finitizes alpha_T tsp =
+fun finitize_funs (hol_ctxt as {thy, ctxt, stds, constr_cache, ...}) binarize
+                  finitizes alpha_T tsp =
   case infer "Finiteness" true hol_ctxt binarize alpha_T tsp of
-    SOME (lits, msp, constr_mtypes) =>
-    if forall (curry (op =) Minus o snd) lits then
+    SOME (asgs, msp, constr_mtypes) =>
+    if forall (curry (op =) Gen o snd) asgs then
       tsp
     else
       let
-        fun should_finitize T a =
+        fun should_finitize T aa =
           case triple_lookup (type_match thy) finitizes T of
             SOME (SOME false) => false
-          | _ => resolve_sign_atom lits a = S Plus
+          | _ => resolve_annotation_atom asgs aa = A Fls
         fun type_from_mtype T M =
           case (M, T) of
             (MAlpha, _) => T
-          | (MFun (M1, a, M2), Type (@{type_name fun}, Ts)) =>
-            Type (if should_finitize T a then @{type_name fin_fun}
+          | (MFun (M1, aa, M2), Type (@{type_name fun}, Ts)) =>
+            Type (if should_finitize T aa then @{type_name fin_fun}
                   else @{type_name fun}, map2 type_from_mtype Ts [M1, M2])
           | (MPair (M1, M2), Type (@{type_name prod}, Ts)) =>
             Type (@{type_name prod}, map2 type_from_mtype Ts [M1, M2])
@@ -1106,14 +1353,14 @@
                 | _ => raise TYPE ("Nitpick_Mono.finitize_funs.term_from_mterm",
                                    [T1], [])
             in betapply (t1', coerce_term hol_ctxt new_Ts T2' T2 t2) end
-          | MAbs (s, old_T, M, a, m') =>
+          | MAbs (s, old_T, M, aa, m') =>
             let
               val new_T = type_from_mtype old_T M
               val t' = term_from_mterm (new_T :: new_Ts) (old_T :: old_Ts) m'
               val T' = fastype_of1 (new_T :: new_Ts, t')
             in
               Abs (s, new_T, t')
-              |> should_finitize (new_T --> T') a
+              |> should_finitize (new_T --> T') aa
                  ? construct_value ctxt stds (fin_fun_constr new_T T') o single
             end
       in
--- a/src/HOL/Tools/SMT/smt_setup_solvers.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_setup_solvers.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -19,7 +19,7 @@
   else raise SMT_Failure.SMT (SMT_Failure.Other_Failure ("Solver " ^
     quote solver_name ^ " failed. Enable SMT tracing by setting the " ^
     "configuration option " ^ quote SMT_Config.traceN ^ " and " ^
-    " see the trace for details."))
+    "see the trace for details."))
 
 fun on_first_line test_outcome solver_name lines =
   let
@@ -40,7 +40,8 @@
   outcome =
     on_first_line (outcome_of "Unsatisfiable." "Satisfiable." "Unknown."),
   cex_parser = NONE,
-  reconstruct = NONE }
+  reconstruct = NONE,
+  default_max_relevant = 250 }
 
 
 (* Yices *)
@@ -53,7 +54,8 @@
   interface = SMTLIB_Interface.interface,
   outcome = on_first_line (outcome_of "unsat" "sat" "unknown"),
   cex_parser = NONE,
-  reconstruct = NONE }
+  reconstruct = NONE,
+  default_max_relevant = 275 }
 
 
 (* Z3 *)
@@ -85,7 +87,8 @@
   interface = Z3_Interface.interface,
   outcome = z3_on_last_line,
   cex_parser = SOME Z3_Model.parse_counterex,
-  reconstruct = SOME Z3_Proof_Reconstruction.reconstruct }
+  reconstruct = SOME Z3_Proof_Reconstruction.reconstruct,
+  default_max_relevant = 225 }
 
 
 (* overall setup *)
--- a/src/HOL/Tools/SMT/smt_solver.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/SMT/smt_solver.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -21,7 +21,8 @@
     cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
       term list * term list) option,
     reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
-      (int list * thm) * Proof.context) option }
+      (int list * thm) * Proof.context) option,
+    default_max_relevant: int }
 
   (*registry*)
   type solver = bool option -> Proof.context -> (int * thm) list ->
@@ -32,6 +33,7 @@
   val available_solvers_of: Proof.context -> string list
   val is_locally_installed: Proof.context -> string -> bool
   val is_remotely_available: Proof.context -> string -> bool
+  val default_max_relevant: Proof.context -> string -> int
 
   (*filter*)
   val smt_filter: bool -> Time.time -> Proof.state -> ('a * thm) list -> int ->
@@ -70,7 +72,8 @@
   cex_parser: (Proof.context -> SMT_Translate.recon -> string list ->
     term list * term list) option,
   reconstruct: (Proof.context -> SMT_Translate.recon -> string list ->
-    (int list * thm) * Proof.context) option }
+    (int list * thm) * Proof.context) option,
+  default_max_relevant: int }
 
 
 (* interface to external solvers *)
@@ -215,7 +218,7 @@
 
 fun gen_solver name info rm ctxt irules =
   let
-    val {env_var, is_remote, options, interface, reconstruct} = info
+    val {env_var, is_remote, options, interface, reconstruct, ...} = info
     val {extra_norm, translate} = interface
     val (with_datatypes, translate') =
       set_has_datatypes (Config.get ctxt C.datatypes) translate
@@ -244,7 +247,8 @@
   options: Proof.context -> string list,
   interface: interface,
   reconstruct: string list * SMT_Translate.recon -> Proof.context ->
-    (int list * thm) * Proof.context }
+    (int list * thm) * Proof.context,
+  default_max_relevant: int }
 
 structure Solvers = Generic_Data
 (
@@ -279,13 +283,14 @@
 fun add_solver cfg =
   let
     val {name, env_var, is_remote, options, interface, outcome, cex_parser,
-      reconstruct} = cfg
+      reconstruct, default_max_relevant} = cfg
 
     fun core_oracle () = cfalse
 
     fun solver ocl = { env_var=env_var, is_remote=is_remote, options=options,
       interface=interface,
-      reconstruct=finish (outcome name) cex_parser reconstruct ocl }
+      reconstruct=finish (outcome name) cex_parser reconstruct ocl,
+      default_max_relevant=default_max_relevant }
   in
     Thm.add_oracle (Binding.name name, core_oracle) #-> (fn (_, ocl) =>
     Context.theory_map (Solvers.map (Symtab.update_new (name, solver ocl)))) #>
@@ -294,9 +299,12 @@
 
 end
 
+fun get_info ctxt name =
+  the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
+
 fun name_and_solver_of ctxt =
   let val name = C.solver_of ctxt
-  in (name, the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)) end
+  in (name, get_info ctxt name) end
 
 val solver_name_of = fst o name_and_solver_of
 fun solver_of ctxt =
@@ -306,11 +314,12 @@
 val available_solvers_of = Symtab.keys o Solvers.get o Context.Proof
 
 fun is_locally_installed ctxt name =
-  let val {env_var, ...} = the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name)
+  let val {env_var, ...} = get_info ctxt name
   in is_some (get_local_solver env_var) end
 
-fun is_remotely_available ctxt name =
-  #is_remote (the (Symtab.lookup (Solvers.get (Context.Proof ctxt)) name))
+val is_remotely_available = #is_remote oo get_info
+
+val default_max_relevant = #default_max_relevant oo get_info
 
 
 (* filter *)
@@ -347,7 +356,8 @@
     (xrules, map snd xrules)
     ||> distinct (op =) o fst o smt_solver rm ctxt' o append irs o map_index I
     |-> map_filter o try o nth
-    |> (fn xs => {outcome=NONE, used_facts=xs, run_time_in_msecs=NONE})
+    |> (fn xs => {outcome=NONE, used_facts=if solver_name_of ctxt = "z3" (* FIXME *) then xs
+      else xrules, run_time_in_msecs=NONE})
   end
   handle SMT_Failure.SMT fail => {outcome=SOME fail, used_facts=[],
     run_time_in_msecs=NONE}
--- a/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -39,8 +39,7 @@
      goal: thm,
      subgoal: int,
      subgoal_count: int,
-     facts: fact list,
-     only: bool}
+     facts: fact list}
 
   type prover_result =
     {outcome: failure option,
@@ -118,13 +117,15 @@
   end
 
 (* FUDGE *)
-val smt_default_max_relevant = 225
 val auto_max_relevant_divisor = 2
 
 fun default_max_relevant_for_prover ctxt name =
   let val thy = ProofContext.theory_of ctxt in
-    if is_smt_prover ctxt name then smt_default_max_relevant
-    else #default_max_relevant (get_atp thy name)
+    if is_smt_prover ctxt name then
+      SMT_Solver.default_max_relevant ctxt
+          (perhaps (try (unprefix remote_prefix)) name)
+    else
+      #default_max_relevant (get_atp thy name)
   end
 
 (* These are typically simplified away by "Meson.presimplify". Equality is
@@ -222,8 +223,7 @@
    goal: thm,
    subgoal: int,
    subgoal_count: int,
-   facts: fact list,
-   only: bool}
+   facts: fact list}
 
 type prover_result =
   {outcome: failure option,
@@ -285,18 +285,15 @@
 
 fun run_atp auto atp_name
         {exec, required_execs, arguments, has_incomplete_mode, proof_delims,
-         known_failures, default_max_relevant, explicit_forall,
-         use_conjecture_for_hypotheses}
-        ({debug, verbose, overlord, full_types, explicit_apply,
-          max_relevant, isar_proof, isar_shrink_factor, timeout, ...} : params)
-        minimize_command
-        ({state, goal, subgoal, facts, only, ...} : prover_problem) =
+         known_failures, explicit_forall, use_conjecture_for_hypotheses, ...}
+        ({debug, verbose, overlord, full_types, explicit_apply, isar_proof,
+          isar_shrink_factor, timeout, ...} : params)
+        minimize_command ({state, goal, subgoal, facts, ...} : prover_problem) =
   let
     val ctxt = Proof.context_of state
     val (_, hyp_ts, concl_t) = strip_subgoal goal subgoal
     val facts =
-      facts |> not only ? take (the_default default_max_relevant max_relevant)
-            |> map (translated_fact ctxt)
+      facts |> map (translated_fact ctxt)
     val dest_dir = if overlord then getenv "ISABELLE_HOME_USER"
                    else Config.get ctxt dest_dir
     val problem_prefix = Config.get ctxt problem_prefix
@@ -480,11 +477,7 @@
           else
             ()
         val {outcome, used_facts, run_time_in_msecs} =
-          TimeLimit.timeLimit iter_timeout
-              (SMT_Solver.smt_filter remote iter_timeout state facts) i
-          handle TimeLimit.TimeOut =>
-                 {outcome = SOME SMT_Failure.Time_Out, used_facts = [],
-                  run_time_in_msecs = NONE}
+          SMT_Solver.smt_filter remote iter_timeout state facts i
         val _ =
           if verbose andalso is_some outcome then
             "SMT outcome: " ^ SMT_Failure.string_of_failure ctxt (the outcome)
@@ -553,9 +546,10 @@
       #> Config.put SMT_Config.monomorph_limit smt_monomorph_limit
     val state = state |> Proof.map_context repair_context
     val thy = Proof.theory_of state
-    val smt_fact = Option.map (apsnd (Thm.transfer thy)) o try dest_Untranslated
+    val get_fact = Option.map (apsnd (Thm.transfer thy)) o try dest_Untranslated
+    val facts = facts |> map_filter get_fact
     val {outcome, used_facts, run_time_in_msecs} =
-      smt_filter_loop params remote state subgoal (map_filter smt_fact facts)
+      smt_filter_loop params remote state subgoal facts
     val (chained_lemmas, other_lemmas) = split_used_facts (map fst used_facts)
     val outcome = outcome |> Option.map failure_from_smt_failure
     val message =
@@ -591,19 +585,21 @@
   end
 
 fun run_prover (params as {blocking, debug, max_relevant, timeout, expect, ...})
-               auto minimize_command
-               (problem as {state, goal, subgoal, subgoal_count, facts, ...})
-               name =
+               auto minimize_command only
+               {state, goal, subgoal, subgoal_count, facts} name =
   let
     val ctxt = Proof.context_of state
     val birth_time = Time.now ()
     val death_time = Time.+ (birth_time, timeout)
     val max_relevant =
       the_default (default_max_relevant_for_prover ctxt name) max_relevant
-    val num_facts = Int.min (length facts, max_relevant)
+    val num_facts = length facts |> not only ? Integer.min max_relevant
     val desc =
       prover_description ctxt params name num_facts subgoal subgoal_count goal
     val prover = get_prover ctxt auto name
+    val problem =
+      {state = state, goal = goal, subgoal = subgoal,
+       subgoal_count = subgoal_count, facts = take num_facts facts}
     fun go () =
       let
         fun really_go () =
@@ -616,8 +612,12 @@
           else
             (really_go ()
              handle ERROR message => ("unknown", "Error: " ^ message ^ "\n")
-                  | exn => ("unknown", "Internal error:\n" ^
-                                       ML_Compiler.exn_message exn ^ "\n"))
+                  | exn =>
+                    if Exn.is_interrupt exn then
+                      reraise exn
+                    else
+                      ("unknown", "Internal error:\n" ^
+                                  ML_Compiler.exn_message exn ^ "\n"))
         val _ =
           if expect = "" orelse outcome_code = expect then
             ()
@@ -686,8 +686,8 @@
               |> map maybe_translate
             val problem =
               {state = state, goal = goal, subgoal = i, subgoal_count = n,
-               facts = facts, only = only}
-            val run_prover = run_prover params auto minimize_command
+               facts = facts}
+            val run_prover = run_prover params auto minimize_command only
           in
             if debug then
               Output.urgent_message (label ^ plural_s (length provers) ^ ": " ^
--- a/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Dec 06 14:45:29 2010 +0100
+++ b/src/HOL/Tools/Sledgehammer/sledgehammer_minimize.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -2,7 +2,7 @@
     Author:     Philipp Meyer, TU Muenchen
     Author:     Jasmin Blanchette, TU Muenchen
 
-Minimization of fact list for Metis using ATPs.
+Minimization of fact list for Metis using external provers.
 *)
 
 signature SLEDGEHAMMER_MINIMIZE =
@@ -59,19 +59,27 @@
     val {goal, ...} = Proof.goal state
     val problem =
       {state = state, goal = goal, subgoal = i, subgoal_count = n,
-       facts = facts, only = true}
+       facts = facts}
     val result as {outcome, used_facts, ...} = prover params (K "") problem
   in
-    Output.urgent_message (case outcome of
-                SOME failure => string_for_failure failure
-              | NONE =>
-                if length used_facts = length facts then "Found proof."
-                else "Found proof with " ^ n_facts used_facts ^ ".");
+    Output.urgent_message
+        (case outcome of
+           SOME failure => string_for_failure failure
+         | NONE =>
+           if length used_facts = length facts then "Found proof."
+           else "Found proof with " ^ n_facts used_facts ^ ".");
     result
   end
 
 (* minimalization of facts *)
 
+(* The sublinear algorithm works well in almost all situations, except when the
+   external prover cannot return the list of used facts and hence returns all
+   facts as used. In that case, the binary algorithm is much more
+   appropriate. We can usually detect that situation just by looking at the
+   number of used facts reported by the prover. *)
+val binary_threshold = 50
+
 fun filter_used_facts used = filter (member (op =) used o fst)
 
 fun sublinear_minimize _ [] p = p
@@ -82,9 +90,35 @@
                          (filter_used_facts used_facts seen, result)
     | _ => sublinear_minimize test xs (x :: seen, result)
 
-(* Give the ATP some slack. The ATP gets further slack because the Sledgehammer
-   preprocessing time is included in the estimate below but isn't part of the
-   timeout. *)
+fun binary_minimize test xs =
+  let
+    fun p xs = #outcome (test xs : prover_result) = NONE
+    fun split [] p = p
+      | split [h] (l, r) = (h :: l, r)
+      | split (h1 :: h2 :: t) (l, r) = split t (h1 :: l, h2 :: r)
+    fun min _ [] = raise Empty
+      | min _ [s0] = [s0]
+      | min sup xs =
+        let val (l0, r0) = split xs ([], []) in
+          if p (sup @ l0) then
+            min sup l0
+          else if p (sup @ r0) then
+            min sup r0
+          else
+            let
+              val l = min (sup @ r0) l0
+              val r = min (sup @ l) r0
+            in l @ r end
+        end
+    val xs =
+      case min [] xs of
+        [x] => if p [] then [] else [x]
+      | xs => xs
+  in (xs, test xs) end
+
+(* Give the external prover some slack. The ATP gets further slack because the
+   Sledgehammer preprocessing time is included in the estimate below but isn't
+   part of the timeout. *)
 val fudge_msecs = 1000
 
 fun minimize_facts {provers = [], ...} _ _ _ _ = error "No prover is set."
@@ -95,7 +129,8 @@
     val ctxt = Proof.context_of state
     val prover = get_prover ctxt false prover_name
     val msecs = Time.toMilliseconds timeout
-    val _ = Output.urgent_message ("Sledgehammer minimize: " ^ quote prover_name ^ ".")
+    val _ = Output.urgent_message ("Sledgehammer minimize: " ^
+                                   quote prover_name ^ ".")
     val {goal, ...} = Proof.goal state
     val (_, hyp_ts, concl_t) = strip_subgoal goal i
     val explicit_apply =
@@ -112,9 +147,12 @@
          val new_timeout =
            Int.min (msecs, Time.toMilliseconds time + fudge_msecs)
            |> Time.fromMilliseconds
+         val facts = filter_used_facts used_facts facts
          val (min_thms, {message, ...}) =
-           sublinear_minimize (do_test new_timeout)
-               (filter_used_facts used_facts facts) ([], result)
+           if length facts >= binary_threshold then
+             binary_minimize (do_test new_timeout) facts
+           else
+             sublinear_minimize (do_test new_timeout) facts ([], result)
          val n = length min_thms
          val _ = Output.urgent_message (cat_lines
            ["Minimized: " ^ string_of_int n ^ " fact" ^ plural_s n] ^
@@ -132,7 +170,7 @@
               \can increase the time limit using the \"timeout\" \
               \option (e.g., \"timeout = " ^
               string_of_int (10 + msecs div 1000) ^ "\").")
-     | {message, ...} => (NONE, "ATP error: " ^ message))
+     | {message, ...} => (NONE, "Prover error: " ^ message))
     handle ERROR msg => (NONE, "Error: " ^ msg)
   end
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Tools/type_lifting.ML	Mon Dec 06 16:18:56 2010 +0100
@@ -0,0 +1,217 @@
+(*  Title:      HOL/Tools/type_lifting.ML
+    Author:     Florian Haftmann, TU Muenchen
+
+Functorial structure of types.
+*)
+
+signature TYPE_LIFTING =
+sig
+  val find_atomic: theory -> typ -> (typ * (bool * bool)) list
+  val construct_mapper: theory -> (string * bool -> term)
+    -> bool -> typ -> typ -> term
+  val type_lifting: string option -> term -> theory -> Proof.state
+  type entry
+  val entries: theory -> entry Symtab.table
+end;
+
+structure Type_Lifting : TYPE_LIFTING =
+struct
+
+val compositionalityN = "compositionality";
+val identityN = "identity";
+
+(** functorial mappers and their properties **)
+
+(* bookkeeping *)
+
+type entry = { mapper: string, variances: (sort * (bool * bool)) list,
+  compositionality: thm, identity: thm };
+
+structure Data = Theory_Data(
+  type T = entry Symtab.table
+  val empty = Symtab.empty
+  fun merge (xy : T * T) = Symtab.merge (K true) xy
+  val extend = I
+);
+
+val entries = Data.get;
+
+
+(* type analysis *)
+
+fun find_atomic thy T =
+  let
+    val variances_of = Option.map #variances o Symtab.lookup (Data.get thy);
+    fun add_variance is_contra T =
+      AList.map_default (op =) (T, (false, false))
+        ((if is_contra then apsnd else apfst) (K true));
+    fun analyze' is_contra (_, (co, contra)) T =
+      (if co then analyze is_contra T else I)
+      #> (if contra then analyze (not is_contra) T else I)
+    and analyze is_contra (T as Type (tyco, Ts)) = (case variances_of tyco
+          of NONE => add_variance is_contra T
+           | SOME variances => fold2 (analyze' is_contra) variances Ts)
+      | analyze is_contra T = add_variance is_contra T;
+  in analyze false T [] end;
+
+fun construct_mapper thy atomic =
+  let
+    val lookup = the o Symtab.lookup (Data.get thy);
+    fun constructs is_contra (_, (co, contra)) T T' =
+      (if co then [construct is_contra T T'] else [])
+      @ (if contra then [construct (not is_contra) T T'] else [])
+    and construct is_contra (T as Type (tyco, Ts)) (T' as Type (_, Ts')) =
+          let
+            val { mapper, variances, ... } = lookup tyco;
+            val args = maps (fn (arg_pattern, (T, T')) =>
+              constructs is_contra arg_pattern T T')
+                (variances ~~ (Ts ~~ Ts'));
+            val (U, U') = if is_contra then (T', T) else (T, T');
+          in list_comb (Const (mapper, map fastype_of args ---> U --> U'), args) end
+      | construct is_contra (TFree (v, _)) (TFree _) = atomic (v, is_contra);
+  in construct end;
+
+
+(* mapper properties *)
+
+fun make_compositionality_prop variances (tyco, mapper) =
+  let
+    fun invents n k nctxt =
+      let
+        val names = Name.invents nctxt n k;
+      in (names, fold Name.declare names nctxt) end;
+    val (((vs1, vs2), vs3), _) = Name.context
+      |> invents Name.aT (length variances)
+      ||>> invents Name.aT (length variances)
+      ||>> invents Name.aT (length variances);
+    fun mk_Ts vs = map2 (fn v => fn (sort, _) => TFree (v, sort))
+      vs variances;
+    val (Ts1, Ts2, Ts3) = (mk_Ts vs1, mk_Ts vs2, mk_Ts vs3);
+    fun mk_argT ((T, T'), (_, (co, contra))) =
+      (if co then [(T --> T')] else [])
+      @ (if contra then [(T' --> T)] else []);
+    val contras = maps (fn (_, (co, contra)) =>
+      (if co then [false] else []) @ (if contra then [true] else [])) variances;
+    val Ts21 = maps mk_argT ((Ts2 ~~ Ts1) ~~ variances);
+    val Ts32 = maps mk_argT ((Ts3 ~~ Ts2) ~~ variances);
+    val ((names21, names32), nctxt) = Name.context
+      |> invents "f" (length Ts21)
+      ||>> invents "f" (length Ts32);
+    val T1 = Type (tyco, Ts1);
+    val T2 = Type (tyco, Ts2);
+    val T3 = Type (tyco, Ts3);
+    val x = Free (the_single (Name.invents nctxt (Long_Name.base_name tyco) 1), T3);
+    val (args21, args32) = (names21 ~~ Ts21, names32 ~~ Ts32);
+    val args31 = map2 (fn is_contra => fn ((f21, T21), (f32, T32)) =>
+      if not is_contra then
+        Abs ("x", domain_type T32, Free (f21, T21) $ (Free (f32, T32) $ Bound 0))
+      else
+        Abs ("x", domain_type T21, Free (f32, T32) $ (Free (f21, T21) $ Bound 0))
+      ) contras (args21 ~~ args32)
+    fun mk_mapper T T' args = list_comb (Const (mapper,
+      map fastype_of args ---> T --> T'), args);
+    val lhs = mk_mapper T2 T1 (map Free args21) $
+      (mk_mapper T3 T2 (map Free args32) $ x);
+    val rhs = mk_mapper T3 T1 args31 $ x;
+  in (map Free (args21 @ args32) @ [x], (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhs, rhs)) end;
+
+fun make_identity_prop variances (tyco, mapper) =
+  let
+    val vs = Name.invents Name.context Name.aT (length variances);
+    val Ts = map2 (fn v => fn (sort, _) => TFree (v, sort)) vs variances;
+    fun bool_num b = if b then 1 else 0;
+    fun mk_argT (T, (_, (co, contra))) =
+      replicate (bool_num co + bool_num contra) (T --> T)
+    val Ts' = maps mk_argT (Ts ~~ variances)
+    val T = Type (tyco, Ts);
+    val x = Free (Long_Name.base_name tyco, T);
+    val lhs = list_comb (Const (mapper, Ts' ---> T --> T),
+      map (fn T => Abs ("x", domain_type T, Bound 0)) Ts') $ x;
+  in (x, (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhs, x)) end;
+
+
+(* analyzing and registering mappers *)
+
+fun consume eq x [] = (false, [])
+  | consume eq x (ys as z :: zs) = if eq (x, z) then (true, zs) else (false, ys);
+
+fun split_mapper_typ "fun" T =
+      let
+        val (Ts', T') = strip_type T;
+        val (Ts'', T'') = split_last Ts';
+        val (Ts''', T''') = split_last Ts'';
+      in (Ts''', T''', T'' --> T') end
+  | split_mapper_typ tyco T =
+      let
+        val (Ts', T') = strip_type T;
+        val (Ts'', T'') = split_last Ts';
+      in (Ts'', T'', T') end;
+
+fun analyze_variances thy tyco T =
+  let
+    fun bad_typ () = error ("Bad mapper type: " ^ Syntax.string_of_typ_global thy T);
+    val (Ts, T1, T2) = split_mapper_typ tyco T
+      handle List.Empty => bad_typ ();
+    val _ = pairself
+      ((fn tyco' => if tyco' = tyco then () else bad_typ ()) o fst o dest_Type) (T1, T2)
+    val (vs1, vs2) = pairself (map dest_TFree o snd o dest_Type) (T1, T2)
+      handle TYPE _ => bad_typ ();
+    val _ = if has_duplicates (eq_fst (op =)) (vs1 @ vs2)
+      then bad_typ () else ();
+    fun check_variance_pair (var1 as (v1, sort1), var2 as (v2, sort2)) =
+      let
+        val coT = TFree var1 --> TFree var2;
+        val contraT = TFree var2 --> TFree var1;
+        val sort = Sign.inter_sort thy (sort1, sort2);
+      in
+        consume (op =) coT
+        ##>> consume (op =) contraT
+        #>> pair sort
+      end;
+    val (variances, left_variances) = fold_map check_variance_pair (vs1 ~~ vs2) Ts;
+    val _ = if null left_variances then () else bad_typ ();
+  in variances end;
+
+fun gen_type_lifting prep_term some_prfx raw_t thy =
+  let
+    val (mapper, T) = case prep_term thy raw_t
+     of Const cT => cT
+      | t => error ("No constant: " ^ Syntax.string_of_term_global thy t);
+    val prfx = the_default (Long_Name.base_name mapper) some_prfx;
+    val _ = Type.no_tvars T;
+    fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
+      | add_tycos _ = I;
+    val tycos = add_tycos T [];
+    val tyco = if tycos = ["fun"] then "fun"
+      else case remove (op =) "fun" tycos
+       of [tyco] => tyco
+        | _ => error ("Bad number of type constructors: " ^ Syntax.string_of_typ_global thy T);
+    val variances = analyze_variances thy tyco T;
+    val compositionality_prop = uncurry (fold_rev Logic.all)
+      (make_compositionality_prop variances (tyco, mapper));
+    val identity_prop = uncurry Logic.all
+      (make_identity_prop variances (tyco, mapper));
+    val qualify = Binding.qualify true prfx o Binding.name;
+    fun after_qed [single_compositionality, single_identity] lthy =
+      lthy
+      |> Local_Theory.note ((qualify compositionalityN, []), single_compositionality)
+      ||>> Local_Theory.note ((qualify identityN, []), single_identity)
+      |-> (fn ((_, [compositionality]), (_, [identity])) =>
+          (Local_Theory.background_theory o Data.map)
+            (Symtab.update (tyco, { mapper = mapper, variances = variances,
+              compositionality = compositionality, identity = identity })));
+  in
+    thy
+    |> Named_Target.theory_init
+    |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) [compositionality_prop, identity_prop])
+  end
+
+val type_lifting = gen_type_lifting Sign.cert_term;
+val type_lifting_cmd = gen_type_lifting Syntax.read_term_global;
+
+val _ =
+  Outer_Syntax.command "type_lifting" "register operations managing the functorial structure of a type" Keyword.thy_goal
+    (Scan.option (Parse.name --| Parse.$$$ ":") -- Parse.term
+      >> (fn (prfx, t) => Toplevel.print o (Toplevel.theory_to_proof (type_lifting_cmd prfx t))));
+
+end;
--- a/src/HOL/Tools/type_mapper.ML	Mon Dec 06 14:45:29 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,217 +0,0 @@
-(*  Title:      HOL/Tools/type_mapper.ML
-    Author:     Florian Haftmann, TU Muenchen
-
-Functorial mappers on types.
-*)
-
-signature TYPE_MAPPER =
-sig
-  val find_atomic: theory -> typ -> (typ * (bool * bool)) list
-  val construct_mapper: theory -> (string * bool -> term)
-    -> bool -> typ -> typ -> term
-  val type_mapper: string option -> term -> theory -> Proof.state
-  type entry
-  val entries: theory -> entry Symtab.table
-end;
-
-structure Type_Mapper : TYPE_MAPPER =
-struct
-
-val compositionalityN = "compositionality";
-val identityN = "identity";
-
-(** functorial mappers and their properties **)
-
-(* bookkeeping *)
-
-type entry = { mapper: string, variances: (sort * (bool * bool)) list,
-  compositionality: thm, identity: thm };
-
-structure Data = Theory_Data(
-  type T = entry Symtab.table
-  val empty = Symtab.empty
-  fun merge (xy : T * T) = Symtab.merge (K true) xy
-  val extend = I
-);
-
-val entries = Data.get;
-
-
-(* type analysis *)
-
-fun find_atomic thy T =
-  let
-    val variances_of = Option.map #variances o Symtab.lookup (Data.get thy);
-    fun add_variance is_contra T =
-      AList.map_default (op =) (T, (false, false))
-        ((if is_contra then apsnd else apfst) (K true));
-    fun analyze' is_contra (_, (co, contra)) T =
-      (if co then analyze is_contra T else I)
-      #> (if contra then analyze (not is_contra) T else I)
-    and analyze is_contra (T as Type (tyco, Ts)) = (case variances_of tyco
-          of NONE => add_variance is_contra T
-           | SOME variances => fold2 (analyze' is_contra) variances Ts)
-      | analyze is_contra T = add_variance is_contra T;
-  in analyze false T [] end;
-
-fun construct_mapper thy atomic =
-  let
-    val lookup = the o Symtab.lookup (Data.get thy);
-    fun constructs is_contra (_, (co, contra)) T T' =
-      (if co then [construct is_contra T T'] else [])
-      @ (if contra then [construct (not is_contra) T T'] else [])
-    and construct is_contra (T as Type (tyco, Ts)) (T' as Type (_, Ts')) =
-          let
-            val { mapper, variances, ... } = lookup tyco;
-            val args = maps (fn (arg_pattern, (T, T')) =>
-              constructs is_contra arg_pattern T T')
-                (variances ~~ (Ts ~~ Ts'));
-            val (U, U') = if is_contra then (T', T) else (T, T');
-          in list_comb (Const (mapper, map fastype_of args ---> U --> U'), args) end
-      | construct is_contra (TFree (v, _)) (TFree _) = atomic (v, is_contra);
-  in construct end;
-
-
-(* mapper properties *)
-
-fun make_compositionality_prop variances (tyco, mapper) =
-  let
-    fun invents n k nctxt =
-      let
-        val names = Name.invents nctxt n k;
-      in (names, fold Name.declare names nctxt) end;
-    val (((vs1, vs2), vs3), _) = Name.context
-      |> invents Name.aT (length variances)
-      ||>> invents Name.aT (length variances)
-      ||>> invents Name.aT (length variances);
-    fun mk_Ts vs = map2 (fn v => fn (sort, _) => TFree (v, sort))
-      vs variances;
-    val (Ts1, Ts2, Ts3) = (mk_Ts vs1, mk_Ts vs2, mk_Ts vs3);
-    fun mk_argT ((T, T'), (_, (co, contra))) =
-      (if co then [(T --> T')] else [])
-      @ (if contra then [(T' --> T)] else []);
-    val contras = maps (fn (_, (co, contra)) =>
-      (if co then [false] else []) @ (if contra then [true] else [])) variances;
-    val Ts21 = maps mk_argT ((Ts2 ~~ Ts1) ~~ variances);
-    val Ts32 = maps mk_argT ((Ts3 ~~ Ts2) ~~ variances);
-    val ((names21, names32), nctxt) = Name.context
-      |> invents "f" (length Ts21)
-      ||>> invents "f" (length Ts32);
-    val T1 = Type (tyco, Ts1);
-    val T2 = Type (tyco, Ts2);
-    val T3 = Type (tyco, Ts3);
-    val x = Free (the_single (Name.invents nctxt (Long_Name.base_name tyco) 1), T3);
-    val (args21, args32) = (names21 ~~ Ts21, names32 ~~ Ts32);
-    val args31 = map2 (fn is_contra => fn ((f21, T21), (f32, T32)) =>
-      if not is_contra then
-        Abs ("x", domain_type T32, Free (f21, T21) $ (Free (f32, T32) $ Bound 0))
-      else
-        Abs ("x", domain_type T21, Free (f32, T32) $ (Free (f21, T21) $ Bound 0))
-      ) contras (args21 ~~ args32)
-    fun mk_mapper T T' args = list_comb (Const (mapper,
-      map fastype_of args ---> T --> T'), args);
-    val lhs = mk_mapper T2 T1 (map Free args21) $
-      (mk_mapper T3 T2 (map Free args32) $ x);
-    val rhs = mk_mapper T3 T1 args31 $ x;
-  in (map Free (args21 @ args32) @ [x], (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhs, rhs)) end;
-
-fun make_identity_prop variances (tyco, mapper) =
-  let
-    val vs = Name.invents Name.context Name.aT (length variances);
-    val Ts = map2 (fn v => fn (sort, _) => TFree (v, sort)) vs variances;
-    fun bool_num b = if b then 1 else 0;
-    fun mk_argT (T, (_, (co, contra))) =
-      replicate (bool_num co + bool_num contra) (T --> T)
-    val Ts' = maps mk_argT (Ts ~~ variances)
-    val T = Type (tyco, Ts);
-    val x = Free (Long_Name.base_name tyco, T);
-    val lhs = list_comb (Const (mapper, Ts' ---> T --> T),
-      map (fn T => Abs ("x", domain_type T, Bound 0)) Ts') $ x;
-  in (x, (HOLogic.mk_Trueprop o HOLogic.mk_eq) (lhs, x)) end;
-
-
-(* analyzing and registering mappers *)
-
-fun consume eq x [] = (false, [])
-  | consume eq x (ys as z :: zs) = if eq (x, z) then (true, zs) else (false, ys);
-
-fun split_mapper_typ "fun" T =
-      let
-        val (Ts', T') = strip_type T;
-        val (Ts'', T'') = split_last Ts';
-        val (Ts''', T''') = split_last Ts'';
-      in (Ts''', T''', T'' --> T') end
-  | split_mapper_typ tyco T =
-      let
-        val (Ts', T') = strip_type T;
-        val (Ts'', T'') = split_last Ts';
-      in (Ts'', T'', T') end;
-
-fun analyze_variances thy tyco T =
-  let
-    fun bad_typ () = error ("Bad mapper type: " ^ Syntax.string_of_typ_global thy T);
-    val (Ts, T1, T2) = split_mapper_typ tyco T
-      handle List.Empty => bad_typ ();
-    val _ = pairself
-      ((fn tyco' => if tyco' = tyco then () else bad_typ ()) o fst o dest_Type) (T1, T2)
-    val (vs1, vs2) = pairself (map dest_TFree o snd o dest_Type) (T1, T2)
-      handle TYPE _ => bad_typ ();
-    val _ = if has_duplicates (eq_fst (op =)) (vs1 @ vs2)
-      then bad_typ () else ();
-    fun check_variance_pair (var1 as (v1, sort1), var2 as (v2, sort2)) =
-      let
-        val coT = TFree var1 --> TFree var2;
-        val contraT = TFree var2 --> TFree var1;
-        val sort = Sign.inter_sort thy (sort1, sort2);
-      in
-        consume (op =) coT
-        ##>> consume (op =) contraT
-        #>> pair sort
-      end;
-    val (variances, left_variances) = fold_map check_variance_pair (vs1 ~~ vs2) Ts;
-    val _ = if null left_variances then () else bad_typ ();
-  in variances end;
-
-fun gen_type_mapper prep_term some_prfx raw_t thy =
-  let
-    val (mapper, T) = case prep_term thy raw_t
-     of Const cT => cT
-      | t => error ("No constant: " ^ Syntax.string_of_term_global thy t);
-    val prfx = the_default (Long_Name.base_name mapper) some_prfx;
-    val _ = Type.no_tvars T;
-    fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts
-      | add_tycos _ = I;
-    val tycos = add_tycos T [];
-    val tyco = if tycos = ["fun"] then "fun"
-      else case remove (op =) "fun" tycos
-       of [tyco] => tyco
-        | _ => error ("Bad number of type constructors: " ^ Syntax.string_of_typ_global thy T);
-    val variances = analyze_variances thy tyco T;
-    val compositionality_prop = uncurry (fold_rev Logic.all)
-      (make_compositionality_prop variances (tyco, mapper));
-    val identity_prop = uncurry Logic.all
-      (make_identity_prop variances (tyco, mapper));
-    val qualify = Binding.qualify true prfx o Binding.name;
-    fun after_qed [single_compositionality, single_identity] lthy =
-      lthy
-      |> Local_Theory.note ((qualify compositionalityN, []), single_compositionality)
-      ||>> Local_Theory.note ((qualify identityN, []), single_identity)
-      |-> (fn ((_, [compositionality]), (_, [identity])) =>
-          (Local_Theory.background_theory o Data.map)
-            (Symtab.update (tyco, { mapper = mapper, variances = variances,
-              compositionality = compositionality, identity = identity })));
-  in
-    thy
-    |> Named_Target.theory_init
-    |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) [compositionality_prop, identity_prop])
-  end
-
-val type_mapper = gen_type_mapper Sign.cert_term;
-val type_mapper_cmd = gen_type_mapper Syntax.read_term_global;
-
-val _ =
-  Outer_Syntax.command "type_mapper" "register functorial mapper for type with its properties" Keyword.thy_goal
-    (Scan.option (Parse.name --| Parse.$$$ ":") -- Parse.term
-      >> (fn (prfx, t) => Toplevel.print o (Toplevel.theory_to_proof (type_mapper_cmd prfx t))));
-
-end;