merged
authorwenzelm
Thu, 29 Oct 2009 18:17:26 +0100
changeset 33333 78faaec3209f
parent 33332 9b5286c0fb14 (diff)
parent 33317 b4534348b8fd (current diff)
child 33334 cba65e4bf565
merged
src/HOL/IsaMakefile
src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML
--- a/src/HOL/Code_Numeral.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Code_Numeral.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -3,7 +3,7 @@
 header {* Type of target language numerals *}
 
 theory Code_Numeral
-imports Nat_Numeral Divides
+imports Nat_Numeral Nat_Transfer Divides
 begin
 
 text {*
--- a/src/HOL/Divides.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Divides.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -6,7 +6,7 @@
 header {* The division operators div and mod *}
 
 theory Divides
-imports Nat_Numeral
+imports Nat_Numeral Nat_Transfer
 uses
   "~~/src/Provers/Arith/assoc_fold.ML"
   "~~/src/Provers/Arith/cancel_numerals.ML"
--- a/src/HOL/Fact.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Fact.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -8,7 +8,7 @@
 header{*Factorial Function*}
 
 theory Fact
-imports Nat_Transfer
+imports Main
 begin
 
 class fact =
@@ -266,9 +266,6 @@
 lemma of_nat_fact_not_zero [simp]: "of_nat (fact n) \<noteq> (0::'a::semiring_char_0)"
 by auto
 
-class ordered_semiring_1 = ordered_semiring + semiring_1
-class ordered_semiring_1_strict = ordered_semiring_strict + semiring_1
-
 lemma of_nat_fact_gt_zero [simp]: "(0::'a::{ordered_semidom}) < of_nat(fact n)" by auto
 
 lemma of_nat_fact_ge_zero [simp]: "(0::'a::ordered_semidom) \<le> of_nat(fact n)"
--- a/src/HOL/Fun.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Fun.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -7,7 +7,6 @@
 
 theory Fun
 imports Complete_Lattice
-uses ("Tools/transfer.ML")
 begin
 
 text{*As a simplification rule, it replaces all function equalities by
@@ -604,16 +603,6 @@
 *}
 
 
-subsection {* Generic transfer procedure *}
-
-definition TransferMorphism:: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> bool"
-  where "TransferMorphism a B \<longleftrightarrow> True"
-
-use "Tools/transfer.ML"
-
-setup Transfer.setup
-
-
 subsection {* Code generator setup *}
 
 types_code
--- a/src/HOL/GCD.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/GCD.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -28,7 +28,7 @@
 header {* GCD *}
 
 theory GCD
-imports Fact
+imports Fact Parity
 begin
 
 declare One_nat_def [simp del]
--- a/src/HOL/Int.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Int.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -1984,6 +1984,135 @@
 lemmas half_gt_zero [simp] = half_gt_zero_iff [THEN iffD2, standard]
 
 
+subsection {* The divides relation *}
+
+lemma zdvd_anti_sym:
+    "0 < m ==> 0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)"
+  apply (simp add: dvd_def, auto)
+  apply (simp add: mult_assoc zero_less_mult_iff zmult_eq_1_iff)
+  done
+
+lemma zdvd_dvd_eq: assumes "a \<noteq> 0" and "(a::int) dvd b" and "b dvd a" 
+  shows "\<bar>a\<bar> = \<bar>b\<bar>"
+proof-
+  from `a dvd b` obtain k where k:"b = a*k" unfolding dvd_def by blast 
+  from `b dvd a` obtain k' where k':"a = b*k'" unfolding dvd_def by blast 
+  from k k' have "a = a*k*k'" by simp
+  with mult_cancel_left1[where c="a" and b="k*k'"]
+  have kk':"k*k' = 1" using `a\<noteq>0` by (simp add: mult_assoc)
+  hence "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1" by (simp add: zmult_eq_1_iff)
+  thus ?thesis using k k' by auto
+qed
+
+lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)"
+  apply (subgoal_tac "m = n + (m - n)")
+   apply (erule ssubst)
+   apply (blast intro: dvd_add, simp)
+  done
+
+lemma zdvd_reduce: "(k dvd n + k * m) = (k dvd (n::int))"
+apply (rule iffI)
+ apply (erule_tac [2] dvd_add)
+ apply (subgoal_tac "n = (n + k * m) - k * m")
+  apply (erule ssubst)
+  apply (erule dvd_diff)
+  apply(simp_all)
+done
+
+lemma dvd_imp_le_int:
+  fixes d i :: int
+  assumes "i \<noteq> 0" and "d dvd i"
+  shows "\<bar>d\<bar> \<le> \<bar>i\<bar>"
+proof -
+  from `d dvd i` obtain k where "i = d * k" ..
+  with `i \<noteq> 0` have "k \<noteq> 0" by auto
+  then have "1 \<le> \<bar>k\<bar>" and "0 \<le> \<bar>d\<bar>" by auto
+  then have "\<bar>d\<bar> * 1 \<le> \<bar>d\<bar> * \<bar>k\<bar>" by (rule mult_left_mono)
+  with `i = d * k` show ?thesis by (simp add: abs_mult)
+qed
+
+lemma zdvd_not_zless:
+  fixes m n :: int
+  assumes "0 < m" and "m < n"
+  shows "\<not> n dvd m"
+proof
+  from assms have "0 < n" by auto
+  assume "n dvd m" then obtain k where k: "m = n * k" ..
+  with `0 < m` have "0 < n * k" by auto
+  with `0 < n` have "0 < k" by (simp add: zero_less_mult_iff)
+  with k `0 < n` `m < n` have "n * k < n * 1" by simp
+  with `0 < n` `0 < k` show False unfolding mult_less_cancel_left by auto
+qed
+
+lemma zdvd_mult_cancel: assumes d:"k * m dvd k * n" and kz:"k \<noteq> (0::int)"
+  shows "m dvd n"
+proof-
+  from d obtain h where h: "k*n = k*m * h" unfolding dvd_def by blast
+  {assume "n \<noteq> m*h" hence "k* n \<noteq> k* (m*h)" using kz by simp
+    with h have False by (simp add: mult_assoc)}
+  hence "n = m * h" by blast
+  thus ?thesis by simp
+qed
+
+theorem zdvd_int: "(x dvd y) = (int x dvd int y)"
+proof -
+  have "\<And>k. int y = int x * k \<Longrightarrow> x dvd y"
+  proof -
+    fix k
+    assume A: "int y = int x * k"
+    then show "x dvd y" proof (cases k)
+      case (1 n) with A have "y = x * n" by (simp add: of_nat_mult [symmetric])
+      then show ?thesis ..
+    next
+      case (2 n) with A have "int y = int x * (- int (Suc n))" by simp
+      also have "\<dots> = - (int x * int (Suc n))" by (simp only: mult_minus_right)
+      also have "\<dots> = - int (x * Suc n)" by (simp only: of_nat_mult [symmetric])
+      finally have "- int (x * Suc n) = int y" ..
+      then show ?thesis by (simp only: negative_eq_positive) auto
+    qed
+  qed
+  then show ?thesis by (auto elim!: dvdE simp only: dvd_triv_left of_nat_mult)
+qed
+
+lemma zdvd1_eq[simp]: "(x::int) dvd 1 = ( \<bar>x\<bar> = 1)"
+proof
+  assume d: "x dvd 1" hence "int (nat \<bar>x\<bar>) dvd int (nat 1)" by simp
+  hence "nat \<bar>x\<bar> dvd 1" by (simp add: zdvd_int)
+  hence "nat \<bar>x\<bar> = 1"  by simp
+  thus "\<bar>x\<bar> = 1" by (cases "x < 0", auto)
+next
+  assume "\<bar>x\<bar>=1"
+  then have "x = 1 \<or> x = -1" by auto
+  then show "x dvd 1" by (auto intro: dvdI)
+qed
+
+lemma zdvd_mult_cancel1: 
+  assumes mp:"m \<noteq>(0::int)" shows "(m * n dvd m) = (\<bar>n\<bar> = 1)"
+proof
+  assume n1: "\<bar>n\<bar> = 1" thus "m * n dvd m" 
+    by (cases "n >0", auto simp add: minus_dvd_iff minus_equation_iff)
+next
+  assume H: "m * n dvd m" hence H2: "m * n dvd m * 1" by simp
+  from zdvd_mult_cancel[OF H2 mp] show "\<bar>n\<bar> = 1" by (simp only: zdvd1_eq)
+qed
+
+lemma int_dvd_iff: "(int m dvd z) = (m dvd nat (abs z))"
+  unfolding zdvd_int by (cases "z \<ge> 0") simp_all
+
+lemma dvd_int_iff: "(z dvd int m) = (nat (abs z) dvd m)"
+  unfolding zdvd_int by (cases "z \<ge> 0") simp_all
+
+lemma nat_dvd_iff: "(nat z dvd m) = (if 0 \<le> z then (z dvd int m) else m = 0)"
+  by (auto simp add: dvd_int_iff)
+
+lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \<le> (n::int)"
+  apply (rule_tac z=n in int_cases)
+  apply (auto simp add: dvd_int_iff)
+  apply (rule_tac z=z in int_cases)
+  apply (auto simp add: dvd_imp_le)
+  done
+
+
 subsection {* Configuration of the code generator *}
 
 code_datatype Pls Min Bit0 Bit1 "number_of \<Colon> int \<Rightarrow> int"
--- a/src/HOL/IntDiv.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/IntDiv.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -1024,139 +1024,16 @@
 lemmas zdvd_iff_zmod_eq_0_number_of [simp] =
   dvd_eq_mod_eq_0 [of "number_of x::int" "number_of y::int", standard]
 
-lemma zdvd_anti_sym:
-    "0 < m ==> 0 < n ==> m dvd n ==> n dvd m ==> m = (n::int)"
-  apply (simp add: dvd_def, auto)
-  apply (simp add: mult_assoc zero_less_mult_iff zmult_eq_1_iff)
-  done
-
-lemma zdvd_dvd_eq: assumes "a \<noteq> 0" and "(a::int) dvd b" and "b dvd a" 
-  shows "\<bar>a\<bar> = \<bar>b\<bar>"
-proof-
-  from `a dvd b` obtain k where k:"b = a*k" unfolding dvd_def by blast 
-  from `b dvd a` obtain k' where k':"a = b*k'" unfolding dvd_def by blast 
-  from k k' have "a = a*k*k'" by simp
-  with mult_cancel_left1[where c="a" and b="k*k'"]
-  have kk':"k*k' = 1" using `a\<noteq>0` by (simp add: mult_assoc)
-  hence "k = 1 \<and> k' = 1 \<or> k = -1 \<and> k' = -1" by (simp add: zmult_eq_1_iff)
-  thus ?thesis using k k' by auto
-qed
-
-lemma zdvd_zdiffD: "k dvd m - n ==> k dvd n ==> k dvd (m::int)"
-  apply (subgoal_tac "m = n + (m - n)")
-   apply (erule ssubst)
-   apply (blast intro: dvd_add, simp)
-  done
-
-lemma zdvd_reduce: "(k dvd n + k * m) = (k dvd (n::int))"
-apply (rule iffI)
- apply (erule_tac [2] dvd_add)
- apply (subgoal_tac "n = (n + k * m) - k * m")
-  apply (erule ssubst)
-  apply (erule dvd_diff)
-  apply(simp_all)
-done
-
 lemma zdvd_zmod: "f dvd m ==> f dvd (n::int) ==> f dvd m mod n"
   by (rule dvd_mod) (* TODO: remove *)
 
 lemma zdvd_zmod_imp_zdvd: "k dvd m mod n ==> k dvd n ==> k dvd (m::int)"
   by (rule dvd_mod_imp_dvd) (* TODO: remove *)
 
-lemma dvd_imp_le_int: "(i::int) ~= 0 ==> d dvd i ==> abs d <= abs i"
-apply(auto simp:abs_if)
-   apply(clarsimp simp:dvd_def mult_less_0_iff)
-  using mult_le_cancel_left_neg[of _ "-1::int"]
-  apply(clarsimp simp:dvd_def mult_less_0_iff)
- apply(clarsimp simp:dvd_def mult_less_0_iff
-         minus_mult_right simp del: mult_minus_right)
-apply(clarsimp simp:dvd_def mult_less_0_iff)
-done
-
-lemma zdvd_not_zless: "0 < m ==> m < n ==> \<not> n dvd (m::int)"
-  apply (auto elim!: dvdE)
-  apply (subgoal_tac "0 < n")
-   prefer 2
-   apply (blast intro: order_less_trans)
-  apply (simp add: zero_less_mult_iff)
-  done
-
 lemma zmult_div_cancel: "(n::int) * (m div n) = m - (m mod n)"
   using zmod_zdiv_equality[where a="m" and b="n"]
   by (simp add: algebra_simps)
 
-lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m"
-apply (subgoal_tac "m mod n = 0")
- apply (simp add: zmult_div_cancel)
-apply (simp only: dvd_eq_mod_eq_0)
-done
-
-lemma zdvd_mult_cancel: assumes d:"k * m dvd k * n" and kz:"k \<noteq> (0::int)"
-  shows "m dvd n"
-proof-
-  from d obtain h where h: "k*n = k*m * h" unfolding dvd_def by blast
-  {assume "n \<noteq> m*h" hence "k* n \<noteq> k* (m*h)" using kz by simp
-    with h have False by (simp add: mult_assoc)}
-  hence "n = m * h" by blast
-  thus ?thesis by simp
-qed
-
-theorem zdvd_int: "(x dvd y) = (int x dvd int y)"
-proof -
-  have "\<And>k. int y = int x * k \<Longrightarrow> x dvd y"
-  proof -
-    fix k
-    assume A: "int y = int x * k"
-    then show "x dvd y" proof (cases k)
-      case (1 n) with A have "y = x * n" by (simp add: zmult_int)
-      then show ?thesis ..
-    next
-      case (2 n) with A have "int y = int x * (- int (Suc n))" by simp
-      also have "\<dots> = - (int x * int (Suc n))" by (simp only: mult_minus_right)
-      also have "\<dots> = - int (x * Suc n)" by (simp only: zmult_int)
-      finally have "- int (x * Suc n) = int y" ..
-      then show ?thesis by (simp only: negative_eq_positive) auto
-    qed
-  qed
-  then show ?thesis by (auto elim!: dvdE simp only: dvd_triv_left int_mult)
-qed
-
-lemma zdvd1_eq[simp]: "(x::int) dvd 1 = ( \<bar>x\<bar> = 1)"
-proof
-  assume d: "x dvd 1" hence "int (nat \<bar>x\<bar>) dvd int (nat 1)" by simp
-  hence "nat \<bar>x\<bar> dvd 1" by (simp add: zdvd_int)
-  hence "nat \<bar>x\<bar> = 1"  by simp
-  thus "\<bar>x\<bar> = 1" by (cases "x < 0", auto)
-next
-  assume "\<bar>x\<bar>=1" thus "x dvd 1" 
-    by(cases "x < 0",simp_all add: minus_equation_iff dvd_eq_mod_eq_0)
-qed
-lemma zdvd_mult_cancel1: 
-  assumes mp:"m \<noteq>(0::int)" shows "(m * n dvd m) = (\<bar>n\<bar> = 1)"
-proof
-  assume n1: "\<bar>n\<bar> = 1" thus "m * n dvd m" 
-    by (cases "n >0", auto simp add: minus_dvd_iff minus_equation_iff)
-next
-  assume H: "m * n dvd m" hence H2: "m * n dvd m * 1" by simp
-  from zdvd_mult_cancel[OF H2 mp] show "\<bar>n\<bar> = 1" by (simp only: zdvd1_eq)
-qed
-
-lemma int_dvd_iff: "(int m dvd z) = (m dvd nat (abs z))"
-  unfolding zdvd_int by (cases "z \<ge> 0") simp_all
-
-lemma dvd_int_iff: "(z dvd int m) = (nat (abs z) dvd m)"
-  unfolding zdvd_int by (cases "z \<ge> 0") simp_all
-
-lemma nat_dvd_iff: "(nat z dvd m) = (if 0 \<le> z then (z dvd int m) else m = 0)"
-  by (auto simp add: dvd_int_iff)
-
-lemma zdvd_imp_le: "[| z dvd n; 0 < n |] ==> z \<le> (n::int)"
-  apply (rule_tac z=n in int_cases)
-  apply (auto simp add: dvd_int_iff)
-  apply (rule_tac z=z in int_cases)
-  apply (auto simp add: dvd_imp_le)
-  done
-
 lemma zpower_zmod: "((x::int) mod m)^y mod m = x^y mod m"
 apply (induct "y", auto)
 apply (rule zmod_zmult1_eq [THEN trans])
@@ -1182,6 +1059,12 @@
 lemma abs_div: "(y::int) dvd x \<Longrightarrow> abs (x div y) = abs x div abs y"
 by (unfold dvd_def, cases "y=0", auto simp add: abs_mult)
 
+lemma zdvd_mult_div_cancel:"(n::int) dvd m \<Longrightarrow> n * (m div n) = m"
+apply (subgoal_tac "m mod n = 0")
+ apply (simp add: zmult_div_cancel)
+apply (simp only: dvd_eq_mod_eq_0)
+done
+
 text{*Suggested by Matthias Daum*}
 lemma int_power_div_base:
      "\<lbrakk>0 < m; 0 < k\<rbrakk> \<Longrightarrow> k ^ m div k = (k::int) ^ (m - Suc 0)"
@@ -1349,6 +1232,43 @@
 declare dvd_eq_mod_eq_0_number_of [simp]
 
 
+subsection {* Transfer setup *}
+
+lemma transfer_nat_int_functions:
+    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
+    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)"
+  by (auto simp add: nat_div_distrib nat_mod_distrib)
+
+lemma transfer_nat_int_function_closures:
+    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0"
+    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0"
+  apply (cases "y = 0")
+  apply (auto simp add: pos_imp_zdiv_nonneg_iff)
+  apply (cases "y = 0")
+  apply auto
+done
+
+declare TransferMorphism_nat_int[transfer add return:
+  transfer_nat_int_functions
+  transfer_nat_int_function_closures
+]
+
+lemma transfer_int_nat_functions:
+    "(int x) div (int y) = int (x div y)"
+    "(int x) mod (int y) = int (x mod y)"
+  by (auto simp add: zdiv_int zmod_int)
+
+lemma transfer_int_nat_function_closures:
+    "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)"
+    "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)"
+  by (simp_all only: is_nat_def transfer_nat_int_function_closures)
+
+declare TransferMorphism_int_nat[transfer add return:
+  transfer_int_nat_functions
+  transfer_int_nat_function_closures
+]
+
+
 subsection {* Code generation *}
 
 definition pdivmod :: "int \<Rightarrow> int \<Rightarrow> int \<times> int" where
--- a/src/HOL/IsaMakefile	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/IsaMakefile	Thu Oct 29 18:17:26 2009 +0100
@@ -223,7 +223,6 @@
   Tools/sat_funcs.ML \
   Tools/sat_solver.ML \
   Tools/split_rule.ML \
-  Tools/transfer.ML \
   Tools/typecopy.ML \
   Tools/typedef_codegen.ML \
   Tools/typedef.ML \
@@ -255,6 +254,7 @@
   Main.thy \
   Map.thy \
   Nat_Numeral.thy \
+  Nat_Transfer.thy \
   Presburger.thy \
   Predicate_Compile.thy \
   Quickcheck.thy \
@@ -276,6 +276,7 @@
   Tools/Groebner_Basis/misc.ML \
   Tools/Groebner_Basis/normalizer.ML \
   Tools/Groebner_Basis/normalizer_data.ML \
+  Tools/choice_specification.ML \
   Tools/int_arith.ML \
   Tools/list_code.ML \
   Tools/meson.ML \
@@ -299,7 +300,6 @@
   Tools/Qelim/presburger.ML \
   Tools/Qelim/qelim.ML \
   Tools/recdef.ML \
-  Tools/choice_specification.ML \
   Tools/res_atp.ML \
   Tools/res_axioms.ML \
   Tools/res_blacklist.ML \
@@ -308,6 +308,7 @@
   Tools/res_reconstruct.ML \
   Tools/string_code.ML \
   Tools/string_syntax.ML \
+  Tools/transfer.ML \
   Tools/TFL/casesplit.ML \
   Tools/TFL/dcterm.ML \
   Tools/TFL/post.ML \
@@ -335,7 +336,6 @@
   Log.thy \
   Lubs.thy \
   MacLaurin.thy \
-  Nat_Transfer.thy \
   NthRoot.thy \
   PReal.thy \
   Parity.thy \
--- a/src/HOL/List.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/List.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -3587,8 +3587,8 @@
 by (blast intro: listrel.intros)
 
 lemma listrel_Cons:
-     "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})";
-by (auto simp add: set_Cons_def intro: listrel.intros) 
+     "listrel r `` {x#xs} = set_Cons (r``{x}) (listrel r `` {xs})"
+by (auto simp add: set_Cons_def intro: listrel.intros)
 
 
 subsection {* Size function *}
@@ -3615,6 +3615,45 @@
 by (induct xs) force+
 
 
+subsection {* Transfer *}
+
+definition
+  embed_list :: "nat list \<Rightarrow> int list"
+where
+  "embed_list l = map int l"
+
+definition
+  nat_list :: "int list \<Rightarrow> bool"
+where
+  "nat_list l = nat_set (set l)"
+
+definition
+  return_list :: "int list \<Rightarrow> nat list"
+where
+  "return_list l = map nat l"
+
+lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow>
+    embed_list (return_list l) = l"
+  unfolding embed_list_def return_list_def nat_list_def nat_set_def
+  apply (induct l)
+  apply auto
+done
+
+lemma transfer_nat_int_list_functions:
+  "l @ m = return_list (embed_list l @ embed_list m)"
+  "[] = return_list []"
+  unfolding return_list_def embed_list_def
+  apply auto
+  apply (induct l, auto)
+  apply (induct m, auto)
+done
+
+(*
+lemma transfer_nat_int_fold1: "fold f l x =
+    fold (%x. f (nat x)) (embed_list l) x";
+*)
+
+
 subsection {* Code generator *}
 
 subsubsection {* Setup *}
@@ -4017,5 +4056,4 @@
   "list_ex P [i..j] = (~ all_from_to_int (%x. ~P x) i j)"
 by(simp add: all_from_to_int_iff_ball list_ex_iff)
 
-
 end
--- a/src/HOL/Nat_Transfer.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Nat_Transfer.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -1,15 +1,26 @@
 
 (* Authors: Jeremy Avigad and Amine Chaieb *)
 
-header {* Sets up transfer from nats to ints and back. *}
+header {* Generic transfer machinery;  specific transfer from nats to ints and back. *}
 
 theory Nat_Transfer
-imports Main Parity
+imports Nat_Numeral
+uses ("Tools/transfer.ML")
 begin
 
+subsection {* Generic transfer machinery *}
+
+definition TransferMorphism:: "('b \<Rightarrow> 'a) \<Rightarrow> 'b set \<Rightarrow> bool"
+  where "TransferMorphism a B \<longleftrightarrow> True"
+
+use "Tools/transfer.ML"
+
+setup Transfer.setup
+
+
 subsection {* Set up transfer from nat to int *}
 
-(* set up transfer direction *)
+text {* set up transfer direction *}
 
 lemma TransferMorphism_nat_int: "TransferMorphism nat (op <= (0::int))"
   by (simp add: TransferMorphism_def)
@@ -20,7 +31,7 @@
   labels: natint
 ]
 
-(* basic functions and relations *)
+text {* basic functions and relations *}
 
 lemma transfer_nat_int_numerals:
     "(0::nat) = nat 0"
@@ -43,31 +54,20 @@
     "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) * (nat y) = nat (x * y)"
     "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) - (nat y) = nat (tsub x y)"
     "(x::int) >= 0 \<Longrightarrow> (nat x)^n = nat (x^n)"
-    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) div (nat y) = nat (x div y)"
-    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> (nat x) mod (nat y) = nat (x mod y)"
   by (auto simp add: eq_nat_nat_iff nat_mult_distrib
-      nat_power_eq nat_div_distrib nat_mod_distrib tsub_def)
+      nat_power_eq tsub_def)
 
 lemma transfer_nat_int_function_closures:
     "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x + y >= 0"
     "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x * y >= 0"
     "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> tsub x y >= 0"
     "(x::int) >= 0 \<Longrightarrow> x^n >= 0"
-    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x div y >= 0"
-    "(x::int) >= 0 \<Longrightarrow> y >= 0 \<Longrightarrow> x mod y >= 0"
     "(0::int) >= 0"
     "(1::int) >= 0"
     "(2::int) >= 0"
     "(3::int) >= 0"
     "int z >= 0"
   apply (auto simp add: zero_le_mult_iff tsub_def)
-  apply (case_tac "y = 0")
-  apply auto
-  apply (subst pos_imp_zdiv_nonneg_iff, auto)
-  apply (case_tac "y = 0")
-  apply force
-  apply (rule pos_mod_sign)
-  apply arith
 done
 
 lemma transfer_nat_int_relations:
@@ -89,7 +89,21 @@
 ]
 
 
-(* first-order quantifiers *)
+text {* first-order quantifiers *}
+
+lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))"
+  by (simp split add: split_nat)
+
+lemma ex_nat: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> P (nat x))"
+proof
+  assume "\<exists>x. P x"
+  then obtain x where "P x" ..
+  then have "int x \<ge> 0 \<and> P (nat (int x))" by simp
+  then show "\<exists>x\<ge>0. P (nat x)" ..
+next
+  assume "\<exists>x\<ge>0. P (nat x)"
+  then show "\<exists>x. P x" by auto
+qed
 
 lemma transfer_nat_int_quantifiers:
     "(ALL (x::nat). P x) = (ALL (x::int). x >= 0 \<longrightarrow> P (nat x))"
@@ -110,7 +124,7 @@
   cong: all_cong ex_cong]
 
 
-(* if *)
+text {* if *}
 
 lemma nat_if_cong: "(if P then (nat x) else (nat y)) =
     nat (if P then x else y)"
@@ -119,7 +133,7 @@
 declare TransferMorphism_nat_int [transfer add return: nat_if_cong]
 
 
-(* operations with sets *)
+text {* operations with sets *}
 
 definition
   nat_set :: "int set \<Rightarrow> bool"
@@ -132,8 +146,6 @@
     "A Un B = nat ` (int ` A Un int ` B)"
     "A Int B = nat ` (int ` A Int int ` B)"
     "{x. P x} = nat ` {x. x >= 0 & P(nat x)}"
-    "{..n} = nat ` {0..int n}"
-    "{m..n} = nat ` {int m..int n}"  (* need all variants of these! *)
   apply (rule card_image [symmetric])
   apply (auto simp add: inj_on_def image_def)
   apply (rule_tac x = "int x" in bexI)
@@ -144,17 +156,12 @@
   apply auto
   apply (rule_tac x = "int x" in exI)
   apply auto
-  apply (rule_tac x = "int x" in bexI)
-  apply auto
-  apply (rule_tac x = "int x" in bexI)
-  apply auto
 done
 
 lemma transfer_nat_int_set_function_closures:
     "nat_set {}"
     "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
     "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
-    "x >= 0 \<Longrightarrow> nat_set {x..y}"
     "nat_set {x. x >= 0 & P x}"
     "nat_set (int ` C)"
     "nat_set A \<Longrightarrow> x : A \<Longrightarrow> x >= 0" (* does it hurt to turn this on? *)
@@ -167,7 +174,6 @@
     "(A = B) = (int ` A = int ` B)"
     "(A < B) = (int ` A < int ` B)"
     "(A <= B) = (int ` A <= int ` B)"
-
   apply (rule iffI)
   apply (erule finite_imageI)
   apply (erule finite_imageD)
@@ -194,7 +200,7 @@
 ]
 
 
-(* setsum and setprod *)
+text {* setsum and setprod *}
 
 (* this handles the case where the *domain* of f is nat *)
 lemma transfer_nat_int_sum_prod:
@@ -262,52 +268,10 @@
     transfer_nat_int_sum_prod_closure
   cong: transfer_nat_int_sum_prod_cong]
 
-(* lists *)
-
-definition
-  embed_list :: "nat list \<Rightarrow> int list"
-where
-  "embed_list l = map int l";
-
-definition
-  nat_list :: "int list \<Rightarrow> bool"
-where
-  "nat_list l = nat_set (set l)";
-
-definition
-  return_list :: "int list \<Rightarrow> nat list"
-where
-  "return_list l = map nat l";
-
-thm nat_0_le;
-
-lemma transfer_nat_int_list_return_embed: "nat_list l \<longrightarrow>
-    embed_list (return_list l) = l";
-  unfolding embed_list_def return_list_def nat_list_def nat_set_def
-  apply (induct l);
-  apply auto;
-done;
-
-lemma transfer_nat_int_list_functions:
-  "l @ m = return_list (embed_list l @ embed_list m)"
-  "[] = return_list []";
-  unfolding return_list_def embed_list_def;
-  apply auto;
-  apply (induct l, auto);
-  apply (induct m, auto);
-done;
-
-(*
-lemma transfer_nat_int_fold1: "fold f l x =
-    fold (%x. f (nat x)) (embed_list l) x";
-*)
-
-
-
 
 subsection {* Set up transfer from int to nat *}
 
-(* set up transfer direction *)
+text {* set up transfer direction *}
 
 lemma TransferMorphism_int_nat: "TransferMorphism int (UNIV :: nat set)"
   by (simp add: TransferMorphism_def)
@@ -319,7 +283,11 @@
 ]
 
 
-(* basic functions and relations *)
+text {* basic functions and relations *}
+
+lemma UNIV_apply:
+  "UNIV x = True"
+  by (simp add: top_fun_eq top_bool_eq)
 
 definition
   is_nat :: "int \<Rightarrow> bool"
@@ -338,17 +306,13 @@
     "(int x) * (int y) = int (x * y)"
     "tsub (int x) (int y) = int (x - y)"
     "(int x)^n = int (x^n)"
-    "(int x) div (int y) = int (x div y)"
-    "(int x) mod (int y) = int (x mod y)"
-  by (auto simp add: int_mult tsub_def int_power zdiv_int zmod_int)
+  by (auto simp add: int_mult tsub_def int_power)
 
 lemma transfer_int_nat_function_closures:
     "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x + y)"
     "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x * y)"
     "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (tsub x y)"
     "is_nat x \<Longrightarrow> is_nat (x^n)"
-    "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x div y)"
-    "is_nat x \<Longrightarrow> is_nat y \<Longrightarrow> is_nat (x mod y)"
     "is_nat 0"
     "is_nat 1"
     "is_nat 2"
@@ -361,12 +325,7 @@
     "(int x < int y) = (x < y)"
     "(int x <= int y) = (x <= y)"
     "(int x dvd int y) = (x dvd y)"
-    "(even (int x)) = (even x)"
-  by (auto simp add: zdvd_int even_nat_def)
-
-lemma UNIV_apply:
-  "UNIV x = True"
-  by (simp add: top_fun_eq top_bool_eq)
+  by (auto simp add: zdvd_int)
 
 declare TransferMorphism_int_nat[transfer add return:
   transfer_int_nat_numerals
@@ -377,7 +336,7 @@
 ]
 
 
-(* first-order quantifiers *)
+text {* first-order quantifiers *}
 
 lemma transfer_int_nat_quantifiers:
     "(ALL (x::int) >= 0. P x) = (ALL (x::nat). P (int x))"
@@ -392,7 +351,7 @@
   return: transfer_int_nat_quantifiers]
 
 
-(* if *)
+text {* if *}
 
 lemma int_if_cong: "(if P then (int x) else (int y)) =
     int (if P then x else y)"
@@ -402,7 +361,7 @@
 
 
 
-(* operations with sets *)
+text {* operations with sets *}
 
 lemma transfer_int_nat_set_functions:
     "nat_set A \<Longrightarrow> card A = card (nat ` A)"
@@ -410,7 +369,6 @@
     "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Un B = int ` (nat ` A Un nat ` B)"
     "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> A Int B = int ` (nat ` A Int nat ` B)"
     "{x. x >= 0 & P x} = int ` {x. P(int x)}"
-    "is_nat m \<Longrightarrow> is_nat n \<Longrightarrow> {m..n} = int ` {nat m..nat n}"
        (* need all variants of these! *)
   by (simp_all only: is_nat_def transfer_nat_int_set_functions
           transfer_nat_int_set_function_closures
@@ -421,7 +379,6 @@
     "nat_set {}"
     "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Un B)"
     "nat_set A \<Longrightarrow> nat_set B \<Longrightarrow> nat_set (A Int B)"
-    "is_nat x \<Longrightarrow> nat_set {x..y}"
     "nat_set {x. x >= 0 & P x}"
     "nat_set (int ` C)"
     "nat_set A \<Longrightarrow> x : A \<Longrightarrow> is_nat x"
@@ -454,7 +411,7 @@
 ]
 
 
-(* setsum and setprod *)
+text {* setsum and setprod *}
 
 (* this handles the case where the *domain* of f is int *)
 lemma transfer_int_nat_sum_prod:
--- a/src/HOL/Parity.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Parity.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -28,6 +28,13 @@
 
 end
 
+lemma transfer_int_nat_relations:
+  "even (int x) \<longleftrightarrow> even x"
+  by (simp add: even_nat_def)
+
+declare TransferMorphism_int_nat[transfer add return:
+  transfer_int_nat_relations
+]
 
 lemma even_zero_int[simp]: "even (0::int)" by presburger
 
@@ -310,6 +317,8 @@
 
 subsection {* General Lemmas About Division *}
 
+(*FIXME move to Divides.thy*)
+
 lemma Suc_times_mod_eq: "1<k ==> Suc (k * m) mod k = 1" 
 apply (induct "m")
 apply (simp_all add: mod_Suc)
--- a/src/HOL/Presburger.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Presburger.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -385,20 +385,6 @@
 
 text {* \bigskip Theorems for transforming predicates on nat to predicates on @{text int}*}
 
-lemma all_nat: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x\<ge>0. P (nat x))"
-  by (simp split add: split_nat)
-
-lemma ex_nat: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. 0 \<le> x \<and> P (nat x))"
-proof
-  assume "\<exists>x. P x"
-  then obtain x where "P x" ..
-  then have "int x \<ge> 0 \<and> P (nat (int x))" by simp
-  then show "\<exists>x\<ge>0. P (nat x)" ..
-next
-  assume "\<exists>x\<ge>0. P (nat x)"
-  then show "\<exists>x. P x" by auto
-qed
-
 lemma zdiff_int_split: "P (int (x - y)) =
   ((y \<le> x \<longrightarrow> P (int x - int y)) \<and> (x < y \<longrightarrow> P 0))"
   by (case_tac "y \<le> x", simp_all add: zdiff_int)
--- a/src/HOL/Ring_and_Field.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Ring_and_Field.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -767,6 +767,8 @@
 
 end
 
+class ordered_semiring_1 = ordered_semiring + semiring_1
+
 class ordered_semiring_strict = semiring + comm_monoid_add + ordered_cancel_ab_semigroup_add +
   assumes mult_strict_left_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> c * a < c * b"
   assumes mult_strict_right_mono: "a < b \<Longrightarrow> 0 < c \<Longrightarrow> a * c < b * c"
@@ -884,6 +886,8 @@
 
 end
 
+class ordered_semiring_1_strict = ordered_semiring_strict + semiring_1
+
 class mult_mono1 = times + zero + ord +
   assumes mult_mono1: "a \<le> b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c * a \<le> c * b"
 
@@ -2025,15 +2029,15 @@
 
 lemma mult_right_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1
     ==> x * y <= x"
-by (auto simp add: mult_compare_simps);
+by (auto simp add: mult_compare_simps)
 
 lemma mult_left_le_one_le: "0 <= (x::'a::ordered_idom) ==> 0 <= y ==> y <= 1
     ==> y * x <= x"
-by (auto simp add: mult_compare_simps);
+by (auto simp add: mult_compare_simps)
 
 lemma mult_imp_div_pos_le: "0 < (y::'a::ordered_field) ==> x <= z * y ==>
-    x / y <= z";
-by (subst pos_divide_le_eq, assumption+);
+    x / y <= z"
+by (subst pos_divide_le_eq, assumption+)
 
 lemma mult_imp_le_div_pos: "0 < (y::'a::ordered_field) ==> z * y <= x ==>
     z <= x / y"
@@ -2060,8 +2064,8 @@
 lemma frac_less: "(0::'a::ordered_field) <= x ==> 
     x < y ==> 0 < w ==> w <= z  ==> x / z < y / w"
   apply (rule mult_imp_div_pos_less)
-  apply simp;
-  apply (subst times_divide_eq_left);
+  apply simp
+  apply (subst times_divide_eq_left)
   apply (rule mult_imp_less_div_pos, assumption)
   apply (erule mult_less_le_imp_less)
   apply simp_all
@@ -2071,7 +2075,7 @@
     x <= y ==> 0 < w ==> w < z  ==> x / z < y / w"
   apply (rule mult_imp_div_pos_less)
   apply simp_all
-  apply (subst times_divide_eq_left);
+  apply (subst times_divide_eq_left)
   apply (rule mult_imp_less_div_pos, assumption)
   apply (erule mult_le_less_imp_less)
   apply simp_all
--- a/src/HOL/SetInterval.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/SetInterval.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -9,7 +9,7 @@
 header {* Set intervals *}
 
 theory SetInterval
-imports Int
+imports Int Nat_Transfer
 begin
 
 context ord
@@ -1150,4 +1150,41 @@
   "\<Prod>i\<le>n. t" == "CONST setprod (\<lambda>i. t) {..n}"
   "\<Prod>i<n. t" == "CONST setprod (\<lambda>i. t) {..<n}"
 
+subsection {* Transfer setup *}
+
+lemma transfer_nat_int_set_functions:
+    "{..n} = nat ` {0..int n}"
+    "{m..n} = nat ` {int m..int n}"  (* need all variants of these! *)
+  apply (auto simp add: image_def)
+  apply (rule_tac x = "int x" in bexI)
+  apply auto
+  apply (rule_tac x = "int x" in bexI)
+  apply auto
+  done
+
+lemma transfer_nat_int_set_function_closures:
+    "x >= 0 \<Longrightarrow> nat_set {x..y}"
+  by (simp add: nat_set_def)
+
+declare TransferMorphism_nat_int[transfer add
+  return: transfer_nat_int_set_functions
+    transfer_nat_int_set_function_closures
+]
+
+lemma transfer_int_nat_set_functions:
+    "is_nat m \<Longrightarrow> is_nat n \<Longrightarrow> {m..n} = int ` {nat m..nat n}"
+  by (simp only: is_nat_def transfer_nat_int_set_functions
+    transfer_nat_int_set_function_closures
+    transfer_nat_int_set_return_embed nat_0_le
+    cong: transfer_nat_int_set_cong)
+
+lemma transfer_int_nat_set_function_closures:
+    "is_nat x \<Longrightarrow> nat_set {x..y}"
+  by (simp only: transfer_nat_int_set_function_closures is_nat_def)
+
+declare TransferMorphism_int_nat[transfer add
+  return: transfer_int_nat_set_functions
+    transfer_int_nat_set_function_closures
+]
+
 end
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile.ML	Thu Oct 29 18:17:26 2009 +0100
@@ -156,10 +156,35 @@
 local structure P = OuterParse
 in
 
+(*val parse_argmode' = P.nat >> rpair NONE || P.$$$ "(" |-- P.enum1 "," --| P.$$$ ")"*)
+datatype raw_argmode = Argmode of string | Argmode_Tuple of string list
+
+val parse_argmode' =
+  ((Args.$$$ "i" || Args.$$$ "o") >> Argmode) ||
+  (Args.$$$ "(" |-- P.enum1 "," (Args.$$$ "i" || Args.$$$ "o") --| Args.$$$ ")" >> Argmode_Tuple)
+
+fun mk_numeral_mode ss = flat (map_index (fn (i, s) => if s = "i" then [i + 1] else []) ss)
+
+val parse_smode' = P.$$$ "[" |-- P.enum1 "," parse_argmode' --| P.$$$ "]"
+  >> (fn m => flat (map_index
+    (fn (i, Argmode s) => if s = "i" then [(i + 1, NONE)] else []
+      | (i, Argmode_Tuple ss) => [(i + 1, SOME (mk_numeral_mode ss))]) m))
+
+val parse_smode = (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
+  >> map (rpair (NONE : int list option))
+
+fun gen_parse_mode smode_parser =
+  (Scan.optional
+    ((P.enum "=>" ((Args.$$$ "X" >> K NONE) || (smode_parser >> SOME))) --| Args.$$$ "==>") [])
+    -- smode_parser
+
+val parse_mode = gen_parse_mode parse_smode
+
+val parse_mode' = gen_parse_mode parse_smode'
+
 val opt_modes =
   Scan.optional (P.$$$ "(" |-- Args.$$$ "mode" |-- P.$$$ ":" |--
-   P.enum1 "," (P.$$$ "[" |-- P.enum "," P.nat --| P.$$$ "]")
-  --| P.$$$ ")" >> SOME) NONE
+    P.enum1 "," (parse_mode || parse_mode') --| P.$$$ ")" >> SOME) NONE
 
 val scan_params =
   let
@@ -170,8 +195,7 @@
 
 val _ = OuterSyntax.local_theory_to_proof "code_pred"
   "prove equations for predicate specified by intro/elim rules"
-  OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >>
-    code_pred_cmd)
+  OuterKeyword.thy_goal (opt_modes -- scan_params -- P.term_group >> code_pred_cmd)
 
 end
 
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML	Thu Oct 29 18:17:26 2009 +0100
@@ -9,6 +9,29 @@
 structure Predicate_Compile_Aux =
 struct
 
+
+(* mode *)
+
+type smode = (int * int list option) list
+type mode = smode option list * smode
+datatype tmode = Mode of mode * smode * tmode option list;
+
+fun string_of_smode js =
+    commas (map
+      (fn (i, is) =>
+        string_of_int i ^ (case is of NONE => ""
+    | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js)
+
+fun string_of_mode (iss, is) = space_implode " -> " (map
+  (fn NONE => "X"
+    | SOME js => enclose "[" "]" (string_of_smode js))
+       (iss @ [SOME is]));
+
+fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
+  "predmode: " ^ (string_of_mode predmode) ^ 
+  (if null param_modes then "" else
+    "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
+
 (* general syntactic functions *)
 
 (*Like dest_conj, but flattens conjunctions however nested*)
@@ -136,7 +159,7 @@
 (* Different options for compiler *)
 
 datatype options = Options of {  
-  expected_modes : (string * int list list) option,
+  expected_modes : (string * mode list) option,
   show_steps : bool,
   show_proof_trace : bool,
   show_intermediate_results : bool,
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Oct 29 18:17:26 2009 +0100
@@ -9,24 +9,20 @@
   val setup: theory -> theory
   val code_pred: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
   val code_pred_cmd: Predicate_Compile_Aux.options -> string -> Proof.context -> Proof.state
-  type smode = (int * int list option) list
-  type mode = smode option list * smode
-  datatype tmode = Mode of mode * smode * tmode option list;
   val register_predicate : (string * thm list * thm * int) -> theory -> theory
   val register_intros : string * thm list -> theory -> theory
   val is_registered : theory -> string -> bool
-  val predfun_intro_of: theory -> string -> mode -> thm
-  val predfun_elim_of: theory -> string -> mode -> thm
-  val predfun_name_of: theory -> string -> mode -> string
+  val predfun_intro_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
+  val predfun_elim_of: theory -> string -> Predicate_Compile_Aux.mode -> thm
+  val predfun_name_of: theory -> string -> Predicate_Compile_Aux.mode -> string
   val all_preds_of : theory -> string list
-  val modes_of: theory -> string -> mode list
-  val depth_limited_modes_of: theory -> string -> mode list
-  val depth_limited_function_name_of : theory -> string -> mode -> string
-  val generator_modes_of: theory -> string -> mode list
-  val generator_name_of : theory -> string -> mode -> string
-  val all_modes_of : theory -> (string * mode list) list
-  val all_generator_modes_of : theory -> (string * mode list) list
-  val string_of_mode : mode -> string
+  val modes_of: theory -> string -> Predicate_Compile_Aux.mode list
+  val depth_limited_modes_of: theory -> string -> Predicate_Compile_Aux.mode list
+  val depth_limited_function_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string
+  val generator_modes_of: theory -> string -> Predicate_Compile_Aux.mode list
+  val generator_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string
+  val all_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list
+  val all_generator_modes_of : theory -> (string * Predicate_Compile_Aux.mode list) list
   val intros_of: theory -> string -> thm list
   val nparams_of: theory -> string -> int
   val add_intro: thm -> theory -> theory
@@ -67,8 +63,6 @@
 
 (* debug stuff *)
 
-fun tracing s = (if ! Toplevel.debug then Output.tracing s else ());
-
 fun print_tac s = Seq.single;
 
 fun print_tac' options s = 
@@ -140,9 +134,6 @@
 type mode = arg_mode list
 type tmode = Mode of mode * 
 *)
-type smode = (int * int list option) list
-type mode = smode option list * smode;
-datatype tmode = Mode of mode * smode * tmode option list;
 
 fun gen_split_smode (mk_tuple, strip_tuple) smode ts =
   let
@@ -165,32 +156,16 @@
         (split_smode' smode (i+1) ts)
   in split_smode' smode 1 ts end
 
-val split_smode = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple)   
-val split_smodeT = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT)
+fun split_smode smode ts = gen_split_smode (HOLogic.mk_tuple, HOLogic.strip_tuple) smode ts
+fun split_smodeT smode ts = gen_split_smode (HOLogic.mk_tupleT, HOLogic.strip_tupleT) smode ts
 
 fun gen_split_mode split_smode (iss, is) ts =
   let
     val (t1, t2) = chop (length iss) ts 
   in (t1, split_smode is t2) end
 
-val split_mode = gen_split_mode split_smode
-val split_modeT = gen_split_mode split_smodeT
-
-fun string_of_smode js =
-    commas (map
-      (fn (i, is) =>
-        string_of_int i ^ (case is of NONE => ""
-    | SOME is => "p" ^ enclose "[" "]" (commas (map string_of_int is)))) js)
-
-fun string_of_mode (iss, is) = space_implode " -> " (map
-  (fn NONE => "X"
-    | SOME js => enclose "[" "]" (string_of_smode js))
-       (iss @ [SOME is]));
-
-fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
-  "predmode: " ^ (string_of_mode predmode) ^ 
-  (if null param_modes then "" else
-    "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
+fun split_mode (iss, is) ts = gen_split_mode split_smode (iss, is) ts
+fun split_modeT (iss, is) ts = gen_split_mode split_smodeT (iss, is) ts
 
 datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term
   | Generator of (string * typ);
@@ -333,7 +308,7 @@
 
 fun print_modes options modes =
   if show_modes options then
-    Output.tracing ("Inferred modes:\n" ^
+    tracing ("Inferred modes:\n" ^
       cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
         string_of_mode ms)) modes))
   else ()
@@ -344,7 +319,7 @@
       ^ (string_of_entry pred mode entry)  
     fun print_pred (pred, modes) =
       "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
-    val _ = Output.tracing (cat_lines (map print_pred pred_mode_table))
+    val _ = tracing (cat_lines (map print_pred pred_mode_table))
   in () end;
 
 fun string_of_prem thy (Prem (ts, p)) =
@@ -423,10 +398,10 @@
   case expected_modes options of
     SOME (s, ms) => (case AList.lookup (op =) modes s of
       SOME modes =>
-        if not (eq_set (op =) (map (map (rpair NONE)) ms, map snd modes)) then
+        if not (eq_set (op =) (ms, modes)) then
           error ("expected modes were not inferred:\n"
-          ^ "inferred modes for " ^ s ^ ": "
-          ^ commas (map ((enclose "[" "]") o string_of_smode o snd) modes))
+          ^ "inferred modes for " ^ s ^ ": " ^ commas (map string_of_mode modes)
+          ^ "\n expected modes for " ^ s ^ ": " ^ commas (map string_of_mode ms))
         else ()
       | NONE => ())
   | NONE => ()
@@ -661,7 +636,7 @@
    fun cons_intro gr =
      case try (Graph.get_node gr) name of
        SOME pred_data => Graph.map_node name (map_pred_data
-         (apfst (fn (intros, elim, nparams) => (thm::intros, elim, nparams)))) gr
+         (apfst (fn (intros, elim, nparams) => (intros @ [thm], elim, nparams)))) gr
      | NONE =>
        let
          val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
@@ -1052,9 +1027,9 @@
 fun print_failed_mode options thy modes p m rs i =
   if show_mode_inference options then
     let
-      val _ = Output.tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
+      val _ = tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
       p ^ " violates mode " ^ string_of_mode m)
-      val _ = Output.tracing (string_of_clause thy p (nth rs i))
+      val _ = tracing (string_of_clause thy p (nth rs i))
     in () end
   else ()
 
@@ -1191,6 +1166,28 @@
     (t, names)
   end;
 
+structure Comp_Mod =
+struct
+
+datatype comp_modifiers = Comp_Modifiers of
+{
+  const_name_of : theory -> string -> Predicate_Compile_Aux.mode -> string,
+  funT_of : compilation_funs -> mode -> typ -> typ,
+  additional_arguments : string list -> term list,
+  wrap_compilation : compilation_funs -> string -> typ -> mode -> term list -> term -> term,
+  transform_additional_arguments : indprem -> term list -> term list
+}
+
+fun dest_comp_modifiers (Comp_Modifiers c) = c
+
+val const_name_of = #const_name_of o dest_comp_modifiers
+val funT_of = #funT_of o dest_comp_modifiers
+val additional_arguments = #additional_arguments o dest_comp_modifiers
+val wrap_compilation = #wrap_compilation o dest_comp_modifiers
+val transform_additional_arguments = #transform_additional_arguments o dest_comp_modifiers
+
+end;
+
 fun compile_arg compilation_modifiers compfuns additional_arguments thy param_vs iss arg = 
   let
     fun map_params (t as Free (f, T)) =
@@ -1198,7 +1195,7 @@
         case (the (AList.lookup (op =) (param_vs ~~ iss) f)) of
           SOME is =>
             let
-              val T' = #funT_of compilation_modifiers compfuns ([], is) T
+              val T' = Comp_Mod.funT_of compilation_modifiers compfuns ([], is) T
             in fst (mk_Eval_of additional_arguments ((Free (f, T'), T), SOME is) []) end
         | NONE => t
       else t
@@ -1248,9 +1245,9 @@
      val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
      val f' =
        case f of
-         Const (name, T) => Const (#const_name_of compilation_modifiers thy name mode,
-           #funT_of compilation_modifiers compfuns mode T)
-       | Free (name, T) => Free (name, #funT_of compilation_modifiers compfuns mode T)
+         Const (name, T) => Const (Comp_Mod.const_name_of compilation_modifiers thy name mode,
+           Comp_Mod.funT_of compilation_modifiers compfuns mode T)
+       | Free (name, T) => Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T)
        | _ => error ("PredicateCompiler: illegal parameter term")
    in
      list_comb (f', params' @ args')
@@ -1262,13 +1259,13 @@
        let
          val params' = map (compile_param compilation_modifiers compfuns thy) (ms ~~ params)
            (*val mk_fun_of = if depth_limited then mk_depth_limited_fun_of else mk_fun_of*)
-         val name' = #const_name_of compilation_modifiers thy name mode
-         val T' = #funT_of compilation_modifiers compfuns mode T
+         val name' = Comp_Mod.const_name_of compilation_modifiers thy name mode
+         val T' = Comp_Mod.funT_of compilation_modifiers compfuns mode T
        in
          (list_comb (Const (name', T'), params' @ inargs @ additional_arguments))
        end
   | (Free (name, T), params) =>
-    list_comb (Free (name, #funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments)
+    list_comb (Free (name, Comp_Mod.funT_of compilation_modifiers compfuns mode T), params @ inargs @ additional_arguments)
 
 fun compile_clause compilation_modifiers compfuns thy all_vs param_vs additional_arguments (iss, is) inp (ts, moded_ps) =
   let
@@ -1302,7 +1299,7 @@
             val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
               out_ts' ((names', map (rpair []) vs))
             val additional_arguments' =
-              #transform_additional_arguments compilation_modifiers p additional_arguments
+              Comp_Mod.transform_additional_arguments compilation_modifiers p additional_arguments
             val (compiled_clause, rest) = case p of
                Prem (us, t) =>
                  let
@@ -1356,7 +1353,7 @@
     val (Ts1, Ts2) = chop (length (fst mode)) (binder_types T)
     val (Us1, Us2) = split_smodeT (snd mode) Ts2
     val Ts1' =
-      map2 (fn NONE => I | SOME is => #funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1
+      map2 (fn NONE => I | SOME is => Comp_Mod.funT_of compilation_modifiers compfuns ([], is)) (fst mode) Ts1
     fun mk_input_term (i, NONE) =
         [Free (Name.variant (all_vs @ param_vs) ("x" ^ string_of_int i), nth Ts2 (i - 1))]
       | mk_input_term (i, SOME pis) = case HOLogic.strip_tupleT (nth Ts2 (i - 1)) of
@@ -1370,17 +1367,17 @@
                else [HOLogic.mk_tuple (map Free (vnames ~~ map (fn j => nth Ts (j - 1)) pis))] end
     val in_ts = maps mk_input_term (snd mode)
     val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
-    val additional_arguments = #additional_arguments compilation_modifiers (all_vs @ param_vs)
+    val additional_arguments = Comp_Mod.additional_arguments compilation_modifiers (all_vs @ param_vs)
     val cl_ts =
       map (compile_clause compilation_modifiers compfuns
         thy all_vs param_vs additional_arguments mode (HOLogic.mk_tuple in_ts)) moded_cls;
-    val compilation = #wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
+    val compilation = Comp_Mod.wrap_compilation compilation_modifiers compfuns s T mode additional_arguments
       (if null cl_ts then
         mk_bot compfuns (HOLogic.mk_tupleT Us2)
       else foldr1 (mk_sup compfuns) cl_ts)
     val fun_const =
-      Const (#const_name_of compilation_modifiers thy s mode,
-        #funT_of compilation_modifiers compfuns mode T)
+      Const (Comp_Mod.const_name_of compilation_modifiers thy s mode,
+        Comp_Mod.funT_of compilation_modifiers compfuns mode T)
   in
     HOLogic.mk_Trueprop
       (HOLogic.mk_eq (list_comb (fun_const, params @ in_ts @ additional_arguments), compilation))
@@ -2139,31 +2136,47 @@
 
 (** main function of predicate compiler **)
 
+datatype steps = Steps of
+  {
+  compile_preds : theory -> string list -> string list -> (string * typ) list
+    -> (moded_clause list) pred_mode_table -> term pred_mode_table,
+  create_definitions: (string * typ) list -> string * mode list -> theory -> theory,
+  infer_modes : options -> theory -> (string * mode list) list -> (string * mode list) list
+    -> string list -> (string * (term list * indprem list) list) list
+    -> moded_clause list pred_mode_table,
+  prove : options -> theory -> (string * (term list * indprem list) list) list
+    -> (string * typ) list -> (string * mode list) list
+    -> moded_clause list pred_mode_table -> term pred_mode_table -> thm pred_mode_table,
+  are_not_defined : theory -> string list -> bool,
+  qname : bstring
+  }
+
+
 fun add_equations_of steps options prednames thy =
   let
+    fun dest_steps (Steps s) = s
     val _ = print_step options ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
-    val _ = tracing (commas (map (Display.string_of_thm_global thy) (maps (intros_of thy) prednames)))
       (*val _ = check_intros_elim_match thy prednames*)
       (*val _ = map (check_format_of_intro_rule thy) (maps (intros_of thy) prednames)*)
     val (preds, nparams, all_vs, param_vs, extra_modes, clauses, all_modes) =
       prepare_intrs thy prednames (maps (intros_of thy) prednames)
     val _ = print_step options "Infering modes..."
-    val moded_clauses = #infer_modes steps options thy extra_modes all_modes param_vs clauses 
+    val moded_clauses = #infer_modes (dest_steps steps) options thy extra_modes all_modes param_vs clauses 
     val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
     val _ = check_expected_modes options modes
     val _ = print_modes options modes
       (*val _ = print_moded_clauses thy moded_clauses*)
     val _ = print_step options "Defining executable functions..."
-    val thy' = fold (#create_definitions steps preds) modes thy
+    val thy' = fold (#create_definitions (dest_steps steps) preds) modes thy
       |> Theory.checkpoint
     val _ = print_step options "Compiling equations..."
     val compiled_terms =
-      (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
+      #compile_preds (dest_steps steps) thy' all_vs param_vs preds moded_clauses
     val _ = print_compiled_terms options thy' compiled_terms
     val _ = print_step options "Proving equations..."
-    val result_thms = #prove steps options thy' clauses preds (extra_modes @ modes)
+    val result_thms = #prove (dest_steps steps) options thy' clauses preds (extra_modes @ modes)
       moded_clauses compiled_terms
-    val qname = #qname steps
+    val qname = #qname (dest_steps steps)
     val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
       (fn thm => Context.mapping (Code.add_eqn thm) I))))
     val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
@@ -2181,7 +2194,7 @@
         SOME v => (G, v)
       | NONE => (Graph.new_node (key, value_of key) G, value_of key)
     val (G'', visited') = fold (extend' value_of edges_of) (subtract (op =) visited (edges_of (key, v)))
-      (G', key :: visited) 
+      (G', key :: visited)
   in
     (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
   end;
@@ -2190,6 +2203,7 @@
   
 fun gen_add_equations steps options names thy =
   let
+    fun dest_steps (Steps s) = s
     val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
       |> Theory.checkpoint;
     fun strong_conn_of gr keys =
@@ -2197,24 +2211,25 @@
     val scc = strong_conn_of (PredData.get thy') names
     val thy'' = fold_rev
       (fn preds => fn thy =>
-        if #are_not_defined steps thy preds then
+        if #are_not_defined (dest_steps steps) thy preds then
           add_equations_of steps options preds thy else thy)
       scc thy' |> Theory.checkpoint
   in thy'' end
 
 (* different instantiantions of the predicate compiler *)
 
-val predicate_comp_modifiers =
-  {const_name_of = predfun_name_of,
-  funT_of = funT_of,
+val predicate_comp_modifiers = Comp_Mod.Comp_Modifiers
+  {const_name_of = predfun_name_of : (theory -> string -> mode -> string),
+  funT_of = funT_of : (compilation_funs -> mode -> typ -> typ),
   additional_arguments = K [],
-  wrap_compilation = K (K (K (K (K I)))),
-  transform_additional_arguments = K I
+  wrap_compilation = K (K (K (K (K I))))
+   : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+  transform_additional_arguments = K I : (indprem -> term list -> term list)
   }
 
-val depth_limited_comp_modifiers =
+val depth_limited_comp_modifiers = Comp_Mod.Comp_Modifiers
   {const_name_of = depth_limited_function_name_of,
-  funT_of = depth_limited_funT_of,
+  funT_of = depth_limited_funT_of : (compilation_funs -> mode -> typ -> typ),
   additional_arguments = fn names =>
     let
       val [depth_name, polarity_name] = Name.variant_list names ["depth", "polarity"]
@@ -2245,38 +2260,38 @@
     in [polarity', depth'] end
   }
 
-val rpred_comp_modifiers =
+val rpred_comp_modifiers = Comp_Mod.Comp_Modifiers
   {const_name_of = generator_name_of,
-  funT_of = K generator_funT_of,
+  funT_of = K generator_funT_of : (compilation_funs -> mode -> typ -> typ),
   additional_arguments = fn names => [Free (Name.variant names "size", @{typ code_numeral})],
-  wrap_compilation = K (K (K (K (K I)))),
-  transform_additional_arguments = K I
+  wrap_compilation = K (K (K (K (K I))))
+    : (compilation_funs -> string -> typ -> mode -> term list -> term -> term),
+  transform_additional_arguments = K I : (indprem -> term list -> term list)
   }
 
-
 val add_equations = gen_add_equations
-  {infer_modes = infer_modes,
+  (Steps {infer_modes = infer_modes,
   create_definitions = create_definitions,
   compile_preds = compile_preds predicate_comp_modifiers PredicateCompFuns.compfuns,
   prove = prove,
   are_not_defined = fn thy => forall (null o modes_of thy),
-  qname = "equation"}
+  qname = "equation"})
 
 val add_depth_limited_equations = gen_add_equations
-  {infer_modes = infer_modes,
+  (Steps {infer_modes = infer_modes,
   create_definitions = create_definitions_of_depth_limited_functions,
   compile_preds = compile_preds depth_limited_comp_modifiers PredicateCompFuns.compfuns,
   prove = prove_by_skip,
   are_not_defined = fn thy => forall (null o depth_limited_modes_of thy),
-  qname = "depth_limited_equation"}
+  qname = "depth_limited_equation"})
 
 val add_quickcheck_equations = gen_add_equations
-  {infer_modes = infer_modes_with_generator,
+  (Steps {infer_modes = infer_modes_with_generator,
   create_definitions = rpred_create_definitions,
   compile_preds = compile_preds rpred_comp_modifiers RandomPredCompFuns.compfuns,
   prove = prove_by_skip,
   are_not_defined = fn thy => forall (null o rpred_modes_of thy),
-  qname = "rpred_equation"}
+  qname = "rpred_equation"})
 
 (** user interface **)
 
@@ -2307,7 +2322,7 @@
         (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
       |> LocalTheory.checkpoint
     val thy' = ProofContext.theory_of lthy'
-    val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
+    val preds = Graph.all_succs (PredData.get thy') [const] |> filter_out (has_elim thy')
     fun mk_cases const =
       let
         val T = Sign.the_const_type thy const
--- a/src/HOL/Tools/transfer.ML	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/Tools/transfer.ML	Thu Oct 29 18:17:26 2009 +0100
@@ -14,8 +14,15 @@
 structure Transfer : TRANSFER =
 struct
 
-type entry = {inj : thm list, emb : thm list, ret : thm list, cong : thm list,
-  guess : bool, hints : string list};
+type entry = { inj : thm list, emb : thm list, ret : thm list, cong : thm list,
+  guess : bool, hints : string list };
+
+fun merge_entry ({ inj = inj1, emb = emb1, ret = ret1, cong = cong1, guess = guess1, hints = hints1 } : entry,
+  { inj = inj2, emb = emb2, ret = ret2, cong = cong2, guess = guess2, hints = hints2 } : entry) =
+    { inj = merge Thm.eq_thm (inj1, inj2), emb = merge Thm.eq_thm (emb1, emb2),
+      ret = merge Thm.eq_thm (ret1, ret2), cong = merge Thm.eq_thm (cong1, cong2),
+      guess = guess1 andalso guess2, hints = merge (op =) (hints1, hints2) };
+
 type data = simpset * (thm * entry) list;
 
 structure Data = GenericDataFun
@@ -24,7 +31,7 @@
   val empty = (HOL_ss, []);
   val extend  = I;
   fun merge _ ((ss1, e1), (ss2, e2)) =
-    (merge_ss (ss1, ss2), AList.merge Thm.eq_thm (K true) (e1, e2));
+    (merge_ss (ss1, ss2), AList.join Thm.eq_thm (K merge_entry) (e1, e2));
 );
 
 val get = Data.get o Context.Proof;
--- a/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile_Alternative_Defs.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -45,7 +45,7 @@
     unfolding mem_def[symmetric, of _ a2]
     apply auto
     unfolding mem_def
-    apply auto
+    apply fastsimp
     done
 qed
 
--- a/src/HOL/ex/Predicate_Compile_ex.thy	Thu Oct 29 17:58:26 2009 +0100
+++ b/src/HOL/ex/Predicate_Compile_ex.thy	Thu Oct 29 18:17:26 2009 +0100
@@ -1,12 +1,12 @@
 theory Predicate_Compile_ex
-imports Main Predicate_Compile_Alternative_Defs
+imports "../Main" Predicate_Compile_Alternative_Defs
 begin
 
 subsection {* Basic predicates *}
 
 inductive False' :: "bool"
 
-code_pred (mode: []) False' .
+code_pred (mode : []) False' .
 code_pred [depth_limited] False' .
 code_pred [rpred] False' .
 
@@ -17,7 +17,7 @@
 definition EmptySet' :: "'a \<Rightarrow> bool"
 where "EmptySet' = {}"
 
-code_pred (mode: [], [1]) [inductify, show_intermediate_results] EmptySet' .
+code_pred (mode: [], [1]) [inductify] EmptySet' .
 
 inductive EmptyRel :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
 
@@ -26,7 +26,13 @@
 inductive EmptyClosure :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
 for r :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
 
-code_pred (mode: [], [1], [2], [1, 2])EmptyClosure .
+code_pred
+  (mode: [] ==> [], [] ==> [1], [] ==> [2], [] ==> [1, 2],
+         [1] ==> [], [1] ==> [1], [1] ==> [2], [1] ==> [1, 2],
+         [2] ==> [], [2] ==> [1], [2] ==> [2], [2] ==> [1, 2],
+         [1, 2] ==> [], [1, 2] ==> [1], [1, 2] ==> [2], [1, 2] ==> [1, 2])
+  EmptyClosure .
+
 thm EmptyClosure.equation
 (* TODO: inductive package is broken!
 inductive False'' :: "bool"
@@ -60,8 +66,88 @@
 where
   "zerozero (0, 0)"
 
-code_pred zerozero .
-code_pred [rpred, show_compilation] zerozero .
+code_pred (mode: [i], [(i, o)], [(o, i)], [o]) zerozero .
+code_pred [rpred] zerozero .
+
+subsection {* Alternative Rules *}
+
+datatype char = C | D | E | F | G | H
+
+inductive is_C_or_D
+where
+  "(x = C) \<or> (x = D) ==> is_C_or_D x"
+
+code_pred (mode: [1]) is_C_or_D .
+thm is_C_or_D.equation
+
+inductive is_D_or_E
+where
+  "(x = D) \<or> (x = E) ==> is_D_or_E x"
+
+lemma [code_pred_intros]:
+  "is_D_or_E D"
+by (auto intro: is_D_or_E.intros)
+
+lemma [code_pred_intros]:
+  "is_D_or_E E"
+by (auto intro: is_D_or_E.intros)
+
+code_pred (mode: [], [1]) is_D_or_E
+proof -
+  case is_D_or_E
+  from this(1) show thesis
+  proof
+    fix x
+    assume x: "a1 = x"
+    assume "x = D \<or> x = E"
+    from this show thesis
+    proof
+      assume "x = D" from this x is_D_or_E(2) show thesis by simp
+    next
+      assume "x = E" from this x is_D_or_E(3) show thesis by simp
+    qed
+  qed
+qed
+
+thm is_D_or_E.equation
+
+inductive is_F_or_G
+where
+  "x = F \<or> x = G ==> is_F_or_G x"
+
+lemma [code_pred_intros]:
+  "is_F_or_G F"
+by (auto intro: is_F_or_G.intros)
+
+lemma [code_pred_intros]:
+  "is_F_or_G G"
+by (auto intro: is_F_or_G.intros)
+
+inductive is_FGH
+where
+  "is_F_or_G x ==> is_FGH x"
+| "is_FGH H"
+
+text {* Compilation of is_FGH requires elimination rule for is_F_or_G *}
+
+code_pred (mode: [], [1]) is_FGH
+proof -
+  case is_F_or_G
+  from this(1) show thesis
+  proof
+    fix x
+    assume x: "a1 = x"
+    assume "x = F \<or> x = G"
+    from this show thesis
+    proof
+      assume "x = F"
+      from this x is_F_or_G(2) show thesis by simp
+    next
+      assume "x = G"
+      from this x is_F_or_G(3) show thesis by simp
+    qed
+  qed
+qed
 
 subsection {* Preprocessor Inlining  *}
 
@@ -123,7 +209,7 @@
 
 definition odd' where "odd' x == \<not> even x"
 
-code_pred [inductify] odd' .
+code_pred (mode: [1]) [inductify] odd' .
 code_pred [inductify, depth_limited] odd' .
 code_pred [inductify, rpred] odd' .
 
@@ -135,7 +221,7 @@
 where
   "n mod 2 = 0 \<Longrightarrow> is_even n"
 
-code_pred is_even .
+code_pred (mode: [1]) is_even .
 
 subsection {* append predicate *}
 
@@ -172,10 +258,19 @@
 
 lemmas [code_pred_intros] = append2_Nil append2.intros(2)
 
-code_pred append2
+code_pred (mode: [1, 2], [3], [2, 3], [1, 3], [1, 2, 3]) append2
 proof -
   case append2
-  from append2.cases[OF append2(1)] append2(2-3) show thesis by blast
+  from append2(1) show thesis
+  proof
+    fix xs
+    assume "a1 = []" "a2 = xs" "a3 = xs"
+    from this append2(2) show thesis by simp
+  next
+    fix xs ys zs x
+    assume "a1 = x # xs" "a2 = ys" "a3 = x # zs" "append2 xs ys zs"
+    from this append2(3) show thesis by fastsimp
+  qed
 qed
 
 inductive tupled_append :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
@@ -183,7 +278,7 @@
   "tupled_append ([], xs, xs)"
 | "tupled_append (xs, ys, zs) \<Longrightarrow> tupled_append (x # xs, ys, x # zs)"
 
-code_pred tupled_append .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append .
 code_pred [rpred] tupled_append .
 thm tupled_append.equation
 (*
@@ -197,7 +292,7 @@
 | "[| ys = fst (xa, y); x # zs = snd (xa, y);
  tupled_append' (xs, ys, zs) |] ==> tupled_append' (x # xs, xa, y)"
 
-code_pred tupled_append' .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) tupled_append' .
 thm tupled_append'.equation
 
 inductive tupled_append'' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
@@ -205,9 +300,7 @@
   "tupled_append'' ([], xs, xs)"
 | "ys = fst yszs ==> x # zs = snd yszs ==> tupled_append'' (xs, ys, zs) \<Longrightarrow> tupled_append'' (x # xs, yszs)"
 
-thm tupled_append''.cases
-
-code_pred [inductify] tupled_append'' .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] tupled_append'' .
 thm tupled_append''.equation
 
 inductive tupled_append''' :: "'a list \<times> 'a list \<times> 'a list \<Rightarrow> bool"
@@ -215,7 +308,7 @@
   "tupled_append''' ([], xs, xs)"
 | "yszs = (ys, zs) ==> tupled_append''' (xs, yszs) \<Longrightarrow> tupled_append''' (x # xs, ys, x # zs)"
 
-code_pred [inductify] tupled_append''' .
+code_pred (mode: [(i,i,o)], [(i,o,i)], [(o,i,i)], [(o,o,i)], [i]) [inductify] tupled_append''' .
 thm tupled_append'''.equation
 
 subsection {* map_ofP predicate *}
@@ -237,7 +330,7 @@
 | "P x ==> filter1 P xs ys ==> filter1 P (x#xs) (x#ys)"
 | "\<not> P x ==> filter1 P xs ys ==> filter1 P (x#xs) ys"
 
-code_pred (mode: [1], [1, 2]) filter1 .
+code_pred (mode: [1] ==> [1], [1] ==> [1, 2]) filter1 .
 code_pred [depth_limited] filter1 .
 code_pred [rpred] filter1 .
 
@@ -260,7 +353,7 @@
 where
   "List.filter P xs = ys ==> filter3 P xs ys"
 
-code_pred filter3 .
+code_pred (mode: [] ==> [1], [] ==> [1, 2], [1] ==> [1], [1] ==> [1, 2]) filter3 .
 code_pred [depth_limited] filter3 .
 thm filter3.depth_limited_equation
 
@@ -268,7 +361,7 @@
 where
   "List.filter P xs = ys ==> filter4 P xs ys"
 
-code_pred filter4 .
+code_pred (mode: [1, 2], [1, 2, 3]) filter4 .
 code_pred [depth_limited] filter4 .
 code_pred [rpred] filter4 .
 
@@ -288,7 +381,7 @@
   "tupled_rev ([], [])"
 | "tupled_rev (xs, xs') \<Longrightarrow> tupled_append (xs', [x], ys) \<Longrightarrow> tupled_rev (x#xs, ys)"
 
-code_pred tupled_rev .
+code_pred (mode: [(i, o)], [(o, i)], [i]) tupled_rev .
 thm tupled_rev.equation
 
 subsection {* partition predicate *}
@@ -299,7 +392,7 @@
   | "f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) (x # ys) zs"
   | "\<not> f x \<Longrightarrow> partition f xs ys zs \<Longrightarrow> partition f (x # xs) ys (x # zs)"
 
-code_pred (mode: [1], [2, 3], [1, 2], [1, 3], [1, 2, 3]) partition .
+code_pred (mode: [1] ==> [1], [1] ==> [2, 3], [1] ==> [1, 2], [1] ==> [1, 3], [1] ==> [1, 2, 3]) partition .
 code_pred [depth_limited] partition .
 code_pred [rpred] partition .
 
@@ -314,7 +407,7 @@
   | "f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, x # ys, zs)"
   | "\<not> f x \<Longrightarrow> tupled_partition f (xs, ys, zs) \<Longrightarrow> tupled_partition f (x # xs, ys, x # zs)"
 
-code_pred tupled_partition .
+code_pred (mode: [i] ==> [i], [i] ==> [(i, i, o)], [i] ==> [(i, o, i)], [i] ==> [(o, i, i)], [i] ==> [(i, o, o)]) tupled_partition .
 
 thm tupled_partition.equation
 
@@ -325,7 +418,7 @@
 
 subsection {* transitive predicate *}
 
-code_pred tranclp
+code_pred (mode: [1] ==> [1, 2], [1] ==> [1], [2] ==> [1, 2], [2] ==> [2], [] ==> [1, 2], [] ==> [1], [] ==> [2], [] ==> []) tranclp
 proof -
   case tranclp
   from this converse_tranclpE[OF this(1)] show thesis by metis
@@ -658,6 +751,8 @@
 | "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
 | "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
 
+code_pred (mode: [], [1]) S\<^isub>4p .
+
 subsection {* Lambda *}
 
 datatype type =
@@ -708,4 +803,10 @@
   | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
 
+code_pred (mode: [1, 2], [1, 2, 3]) typing .
+thm typing.equation
+
+code_pred (mode: [1], [1, 2]) beta .
+thm beta.equation
+
 end
\ No newline at end of file
--- a/src/HOL/ex/predicate_compile.ML	Thu Oct 29 17:58:26 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2182 +0,0 @@
-(* Author: Lukas Bulwahn, TU Muenchen
-
-(Prototype of) A compiler from predicates specified by intro/elim rules
-to equations.
-*)
-
-signature PREDICATE_COMPILE =
-sig
-  type mode = int list option list * int list
-  (*val add_equations_of: bool -> string list -> theory -> theory *)
-  val register_predicate : (thm list * thm * int) -> theory -> theory
-  val is_registered : theory -> string -> bool
- (* val fetch_pred_data : theory -> string -> (thm list * thm * int)  *)
-  val predfun_intro_of: theory -> string -> mode -> thm
-  val predfun_elim_of: theory -> string -> mode -> thm
-  val strip_intro_concl: int -> term -> term * (term list * term list)
-  val predfun_name_of: theory -> string -> mode -> string
-  val all_preds_of : theory -> string list
-  val modes_of: theory -> string -> mode list
-  val string_of_mode : mode -> string
-  val intros_of: theory -> string -> thm list
-  val nparams_of: theory -> string -> int
-  val add_intro: thm -> theory -> theory
-  val set_elim: thm -> theory -> theory
-  val setup: theory -> theory
-  val code_pred: string -> Proof.context -> Proof.state
-  val code_pred_cmd: string -> Proof.context -> Proof.state
-  val print_stored_rules: theory -> unit
-  val print_all_modes: theory -> unit
-  val do_proofs: bool Unsynchronized.ref
-  val mk_casesrule : Proof.context -> int -> thm list -> term
-  val analyze_compr: theory -> term -> term
-  val eval_ref: (unit -> term Predicate.pred) option Unsynchronized.ref
-  val add_equations : string list -> theory -> theory
-  val code_pred_intros_attrib : attribute
-  (* used by Quickcheck_Generator *) 
-  (*val funT_of : mode -> typ -> typ
-  val mk_if_pred : term -> term
-  val mk_Eval : term * term -> term*)
-  val mk_tupleT : typ list -> typ
-(*  val mk_predT :  typ -> typ *)
-  (* temporary for testing of the compilation *)
-  datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
-    GeneratorPrem of term list * term | Generator of (string * typ);
-  val prepare_intrs: theory -> string list ->
-    (string * typ) list * int * string list * string list * (string * mode list) list *
-    (string * (term list * indprem list) list) list * (string * (int option list * int)) list
-  datatype compilation_funs = CompilationFuns of {
-    mk_predT : typ -> typ,
-    dest_predT : typ -> typ,
-    mk_bot : typ -> term,
-    mk_single : term -> term,
-    mk_bind : term * term -> term,
-    mk_sup : term * term -> term,
-    mk_if : term -> term,
-    mk_not : term -> term,
-    mk_map : typ -> typ -> term -> term -> term,
-    lift_pred : term -> term
-  };  
-  datatype tmode = Mode of mode * int list * tmode option list;
-  type moded_clause = term list * (indprem * tmode) list
-  type 'a pred_mode_table = (string * (mode * 'a) list) list
-  val infer_modes : bool -> theory -> (string * (int list option list * int list) list) list
-    -> (string * (int option list * int)) list -> string list
-    -> (string * (term list * indprem list) list) list
-    -> (moded_clause list) pred_mode_table
-  val infer_modes_with_generator : theory -> (string * (int list option list * int list) list) list
-    -> (string * (int option list * int)) list -> string list
-    -> (string * (term list * indprem list) list) list
-    -> (moded_clause list) pred_mode_table  
-  (*val compile_preds : theory -> compilation_funs -> string list -> string list
-    -> (string * typ) list -> (moded_clause list) pred_mode_table -> term pred_mode_table
-  val rpred_create_definitions :(string * typ) list -> string * mode list
-    -> theory -> theory 
-  val split_smode : int list -> term list -> (term list * term list) *)
-  val print_moded_clauses :
-    theory -> (moded_clause list) pred_mode_table -> unit
-  val print_compiled_terms : theory -> term pred_mode_table -> unit
-  (*val rpred_prove_preds : theory -> term pred_mode_table -> thm pred_mode_table*)
-  val rpred_compfuns : compilation_funs
-  val dest_funT : typ -> typ * typ
- (* val depending_preds_of : theory -> thm list -> string list *)
-  val add_quickcheck_equations : string list -> theory -> theory
-  val add_sizelim_equations : string list -> theory -> theory
-  val is_inductive_predicate : theory -> string -> bool
-  val terms_vs : term list -> string list
-  val subsets : int -> int -> int list list
-  val check_mode_clause : bool -> theory -> string list ->
-    (string * mode list) list -> (string * mode list) list -> mode -> (term list * indprem list)
-      -> (term list * (indprem * tmode) list) option
-  val string_of_moded_prem : theory -> (indprem * tmode) -> string
-  val all_modes_of : theory -> (string * mode list) list
-  val all_generator_modes_of : theory -> (string * mode list) list
-  val compile_clause : compilation_funs -> term option -> (term list -> term) ->
-    theory -> string list -> string list -> mode -> term -> moded_clause -> term
-  val preprocess_intro : theory -> thm -> thm
-  val is_constrt : theory -> term -> bool
-  val is_predT : typ -> bool
-  val guess_nparams : typ -> int
-end;
-
-structure Predicate_Compile : PREDICATE_COMPILE =
-struct
-
-(** auxiliary **)
-
-(* debug stuff *)
-
-fun tracing s = (if ! Toplevel.debug then tracing s else ());
-
-fun print_tac s = Seq.single; (* (if ! Toplevel.debug then Tactical.print_tac s else Seq.single); *)
-fun debug_tac msg = Seq.single; (* (fn st => (tracing msg; Seq.single st)); *)
-
-val do_proofs = Unsynchronized.ref true;
-
-fun mycheat_tac thy i st =
-  (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) i) st
-
-fun remove_last_goal thy st =
-  (Tactic.rtac (Skip_Proof.make_thm thy (Var (("A", 0), propT))) (nprems_of st)) st
-
-(* reference to preprocessing of InductiveSet package *)
-
-val ind_set_codegen_preproc = Inductive_Set.codegen_preproc;
-
-(** fundamentals **)
-
-(* syntactic operations *)
-
-fun mk_eq (x, xs) =
-  let fun mk_eqs _ [] = []
-        | mk_eqs a (b::cs) =
-            HOLogic.mk_eq (Free (a, fastype_of b), b) :: mk_eqs a cs
-  in mk_eqs x xs end;
-
-fun mk_tupleT [] = HOLogic.unitT
-  | mk_tupleT Ts = foldr1 HOLogic.mk_prodT Ts;
-
-fun dest_tupleT (Type (@{type_name Product_Type.unit}, [])) = []
-  | dest_tupleT (Type (@{type_name "*"}, [T1, T2])) = T1 :: (dest_tupleT T2)
-  | dest_tupleT t = [t]
-
-fun mk_tuple [] = HOLogic.unit
-  | mk_tuple ts = foldr1 HOLogic.mk_prod ts;
-
-fun dest_tuple (Const (@{const_name Product_Type.Unity}, _)) = []
-  | dest_tuple (Const (@{const_name Pair}, _) $ t1 $ t2) = t1 :: (dest_tuple t2)
-  | dest_tuple t = [t]
-
-fun mk_scomp (t, u) =
-  let
-    val T = fastype_of t
-    val U = fastype_of u
-    val [A] = binder_types T
-    val D = body_type U 
-  in 
-    Const (@{const_name "scomp"}, T --> U --> A --> D) $ t $ u
-  end;
-
-fun dest_funT (Type ("fun",[S, T])) = (S, T)
-  | dest_funT T = raise TYPE ("dest_funT", [T], [])
- 
-fun mk_fun_comp (t, u) =
-  let
-    val (_, B) = dest_funT (fastype_of t)
-    val (C, A) = dest_funT (fastype_of u)
-  in
-    Const(@{const_name "Fun.comp"}, (A --> B) --> (C --> A) --> C --> B) $ t $ u
-  end;
-
-fun dest_randomT (Type ("fun", [@{typ Random.seed},
-  Type ("*", [Type ("*", [T, @{typ "unit => Code_Evaluation.term"}]) ,@{typ Random.seed}])])) = T
-  | dest_randomT T = raise TYPE ("dest_randomT", [T], [])
-
-(* destruction of intro rules *)
-
-(* FIXME: look for other place where this functionality was used before *)
-fun strip_intro_concl nparams intro = let
-  val _ $ u = Logic.strip_imp_concl intro
-  val (pred, all_args) = strip_comb u
-  val (params, args) = chop nparams all_args
-in (pred, (params, args)) end
-
-(** data structures **)
-
-type smode = int list;
-type mode = smode option list * smode;
-datatype tmode = Mode of mode * int list * tmode option list;
-
-fun split_smode is ts =
-  let
-    fun split_smode' _ _ [] = ([], [])
-      | split_smode' is i (t::ts) = (if i mem is then apfst else apsnd) (cons t)
-          (split_smode' is (i+1) ts)
-  in split_smode' is 1 ts end
-
-fun split_mode (iss, is) ts =
-  let
-    val (t1, t2) = chop (length iss) ts 
-  in (t1, split_smode is t2) end
-
-fun string_of_mode (iss, is) = space_implode " -> " (map
-  (fn NONE => "X"
-    | SOME js => enclose "[" "]" (commas (map string_of_int js)))
-       (iss @ [SOME is]));
-
-fun string_of_tmode (Mode (predmode, termmode, param_modes)) =
-  "predmode: " ^ (string_of_mode predmode) ^ 
-  (if null param_modes then "" else
-    "; " ^ "params: " ^ commas (map (the_default "NONE" o Option.map string_of_tmode) param_modes))
-    
-datatype indprem = Prem of term list * term | Negprem of term list * term | Sidecond of term |
-  GeneratorPrem of term list * term | Generator of (string * typ);
-
-type moded_clause = term list * (indprem * tmode) list
-type 'a pred_mode_table = (string * (mode * 'a) list) list
-
-datatype predfun_data = PredfunData of {
-  name : string,
-  definition : thm,
-  intro : thm,
-  elim : thm
-};
-
-fun rep_predfun_data (PredfunData data) = data;
-fun mk_predfun_data (name, definition, intro, elim) =
-  PredfunData {name = name, definition = definition, intro = intro, elim = elim}
-
-datatype function_data = FunctionData of {
-  name : string,
-  equation : thm option (* is not used at all? *)
-};
-
-fun rep_function_data (FunctionData data) = data;
-fun mk_function_data (name, equation) =
-  FunctionData {name = name, equation = equation}
-
-datatype pred_data = PredData of {
-  intros : thm list,
-  elim : thm option,
-  nparams : int,
-  functions : (mode * predfun_data) list,
-  generators : (mode * function_data) list,
-  sizelim_functions : (mode * function_data) list 
-};
-
-fun rep_pred_data (PredData data) = data;
-fun mk_pred_data ((intros, elim, nparams), (functions, generators, sizelim_functions)) =
-  PredData {intros = intros, elim = elim, nparams = nparams,
-    functions = functions, generators = generators, sizelim_functions = sizelim_functions}
-fun map_pred_data f (PredData {intros, elim, nparams, functions, generators, sizelim_functions}) =
-  mk_pred_data (f ((intros, elim, nparams), (functions, generators, sizelim_functions)))
-  
-fun eq_option eq (NONE, NONE) = true
-  | eq_option eq (SOME x, SOME y) = eq (x, y)
-  | eq_option eq _ = false
-  
-fun eq_pred_data (PredData d1, PredData d2) = 
-  eq_list (Thm.eq_thm) (#intros d1, #intros d2) andalso
-  eq_option (Thm.eq_thm) (#elim d1, #elim d2) andalso
-  #nparams d1 = #nparams d2
-  
-structure PredData = TheoryDataFun
-(
-  type T = pred_data Graph.T;
-  val empty = Graph.empty;
-  val copy = I;
-  val extend = I;
-  fun merge _ = Graph.merge eq_pred_data;
-);
-
-(* queries *)
-
-fun lookup_pred_data thy name =
-  Option.map rep_pred_data (try (Graph.get_node (PredData.get thy)) name)
-
-fun the_pred_data thy name = case lookup_pred_data thy name
- of NONE => error ("No such predicate " ^ quote name)  
-  | SOME data => data;
-
-val is_registered = is_some oo lookup_pred_data 
-
-val all_preds_of = Graph.keys o PredData.get
-
-val intros_of = #intros oo the_pred_data
-
-fun the_elim_of thy name = case #elim (the_pred_data thy name)
- of NONE => error ("No elimination rule for predicate " ^ quote name)
-  | SOME thm => thm 
-  
-val has_elim = is_some o #elim oo the_pred_data;
-
-val nparams_of = #nparams oo the_pred_data
-
-val modes_of = (map fst) o #functions oo the_pred_data
-
-fun all_modes_of thy = map (fn name => (name, modes_of thy name)) (all_preds_of thy) 
-
-val is_compiled = not o null o #functions oo the_pred_data
-
-fun lookup_predfun_data thy name mode =
-  Option.map rep_predfun_data (AList.lookup (op =)
-  (#functions (the_pred_data thy name)) mode)
-
-fun the_predfun_data thy name mode = case lookup_predfun_data thy name mode
-  of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
-   | SOME data => data;
-
-val predfun_name_of = #name ooo the_predfun_data
-
-val predfun_definition_of = #definition ooo the_predfun_data
-
-val predfun_intro_of = #intro ooo the_predfun_data
-
-val predfun_elim_of = #elim ooo the_predfun_data
-
-fun lookup_generator_data thy name mode = 
-  Option.map rep_function_data (AList.lookup (op =)
-  (#generators (the_pred_data thy name)) mode)
-  
-fun the_generator_data thy name mode = case lookup_generator_data thy name mode
-  of NONE => error ("No generator defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name)
-   | SOME data => data
-
-val generator_name_of = #name ooo the_generator_data
-
-val generator_modes_of = (map fst) o #generators oo the_pred_data
-
-fun all_generator_modes_of thy =
-  map (fn name => (name, generator_modes_of thy name)) (all_preds_of thy) 
-
-fun lookup_sizelim_function_data thy name mode =
-  Option.map rep_function_data (AList.lookup (op =)
-  (#sizelim_functions (the_pred_data thy name)) mode)
-
-fun the_sizelim_function_data thy name mode = case lookup_sizelim_function_data thy name mode
-  of NONE => error ("No size-limited function defined for mode " ^ string_of_mode mode
-    ^ " of predicate " ^ name)
-   | SOME data => data
-
-val sizelim_function_name_of = #name ooo the_sizelim_function_data
-
-(*val generator_modes_of = (map fst) o #generators oo the_pred_data*)
-     
-(* diagnostic display functions *)
-
-fun print_modes modes = tracing ("Inferred modes:\n" ^
-  cat_lines (map (fn (s, ms) => s ^ ": " ^ commas (map
-    string_of_mode ms)) modes));
-
-fun print_pred_mode_table string_of_entry thy pred_mode_table =
-  let
-    fun print_mode pred (mode, entry) =  "mode : " ^ (string_of_mode mode)
-      ^ (string_of_entry pred mode entry)  
-    fun print_pred (pred, modes) =
-      "predicate " ^ pred ^ ": " ^ cat_lines (map (print_mode pred) modes)
-    val _ = tracing (cat_lines (map print_pred pred_mode_table))
-  in () end;
-
-fun string_of_moded_prem thy (Prem (ts, p), tmode) =
-    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
-    "(" ^ (string_of_tmode tmode) ^ ")"
-  | string_of_moded_prem thy (GeneratorPrem (ts, p), Mode (predmode, is, _)) =
-    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
-    "(generator_mode: " ^ (string_of_mode predmode) ^ ")"
-  | string_of_moded_prem thy (Generator (v, T), _) =
-    "Generator for " ^ v ^ " of Type " ^ (Syntax.string_of_typ_global thy T)
-  | string_of_moded_prem thy (Negprem (ts, p), Mode (_, is, _)) =
-    (Syntax.string_of_term_global thy (list_comb (p, ts))) ^
-    "(negative mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"
-  | string_of_moded_prem thy (Sidecond t, Mode (_, is, _)) =
-    (Syntax.string_of_term_global thy t) ^
-    "(sidecond mode: " ^ (space_implode ", " (map string_of_int is)) ^ ")"    
-  | string_of_moded_prem _ _ = error "string_of_moded_prem: unimplemented"
-     
-fun print_moded_clauses thy =
-  let        
-    fun string_of_clause pred mode clauses =
-      cat_lines (map (fn (ts, prems) => (space_implode " --> "
-        (map (string_of_moded_prem thy) prems)) ^ " --> " ^ pred ^ " "
-        ^ (space_implode " " (map (Syntax.string_of_term_global thy) ts))) clauses)
-  in print_pred_mode_table string_of_clause thy end;
-
-fun print_compiled_terms thy =
-  print_pred_mode_table (fn _ => fn _ => Syntax.string_of_term_global thy) thy
-    
-fun print_stored_rules thy =
-  let
-    val preds = (Graph.keys o PredData.get) thy
-    fun print pred () = let
-      val _ = writeln ("predicate: " ^ pred)
-      val _ = writeln ("number of parameters: " ^ string_of_int (nparams_of thy pred))
-      val _ = writeln ("introrules: ")
-      val _ = fold (fn thm => fn u => writeln (Display.string_of_thm_global thy thm))
-        (rev (intros_of thy pred)) ()
-    in
-      if (has_elim thy pred) then
-        writeln ("elimrule: " ^ Display.string_of_thm_global thy (the_elim_of thy pred))
-      else
-        writeln ("no elimrule defined")
-    end
-  in
-    fold print preds ()
-  end;
-
-fun print_all_modes thy =
-  let
-    val _ = writeln ("Inferred modes:")
-    fun print (pred, modes) u =
-      let
-        val _ = writeln ("predicate: " ^ pred)
-        val _ = writeln ("modes: " ^ (commas (map string_of_mode modes)))
-      in u end  
-  in
-    fold print (all_modes_of thy) ()
-  end
-  
-(** preprocessing rules **)  
-
-fun imp_prems_conv cv ct =
-  case Thm.term_of ct of
-    Const ("==>", _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct
-  | _ => Conv.all_conv ct
-
-fun Trueprop_conv cv ct =
-  case Thm.term_of ct of
-    Const ("Trueprop", _) $ _ => Conv.arg_conv cv ct  
-  | _ => error "Trueprop_conv"
-
-fun preprocess_intro thy rule =
-  Conv.fconv_rule
-    (imp_prems_conv
-      (Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq})))))
-    (Thm.transfer thy rule)
-
-fun preprocess_elim thy nparams elimrule =
-  let
-    fun replace_eqs (Const ("Trueprop", _) $ (Const ("op =", T) $ lhs $ rhs)) =
-       HOLogic.mk_Trueprop (Const (@{const_name Predicate.eq}, T) $ lhs $ rhs)
-     | replace_eqs t = t
-    val prems = Thm.prems_of elimrule
-    val nargs = length (snd (strip_comb (HOLogic.dest_Trueprop (hd prems)))) - nparams
-    fun preprocess_case t =
-     let
-       val params = Logic.strip_params t
-       val (assums1, assums2) = chop nargs (Logic.strip_assums_hyp t)
-       val assums_hyp' = assums1 @ (map replace_eqs assums2)
-     in
-       list_all (params, Logic.list_implies (assums_hyp', Logic.strip_assums_concl t))
-     end 
-    val cases' = map preprocess_case (tl prems)
-    val elimrule' = Logic.list_implies ((hd prems) :: cases', Thm.concl_of elimrule)
-  in
-    Thm.equal_elim
-      (Thm.symmetric (Conv.implies_concl_conv (MetaSimplifier.rewrite true [@{thm eq_is_eq}])
-         (cterm_of thy elimrule')))
-      elimrule
-  end;
-
-(* special case: predicate with no introduction rule *)
-fun noclause thy predname elim = let
-  val T = (Logic.unvarifyT o Sign.the_const_type thy) predname
-  val Ts = binder_types T
-  val names = Name.variant_list []
-        (map (fn i => "x" ^ (string_of_int i)) (1 upto (length Ts)))
-  val vs = map2 (curry Free) names Ts
-  val clausehd = HOLogic.mk_Trueprop (list_comb (Const (predname, T), vs))
-  val intro_t = Logic.mk_implies (@{prop False}, clausehd)
-  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))
-  val elim_t = Logic.list_implies ([clausehd, Logic.mk_implies (@{prop False}, P)], P)
-  val intro = Goal.prove (ProofContext.init thy) names [] intro_t
-        (fn {...} => etac @{thm FalseE} 1)
-  val elim = Goal.prove (ProofContext.init thy) ("P" :: names) [] elim_t
-        (fn {...} => etac elim 1) 
-in
-  ([intro], elim)
-end
-
-fun fetch_pred_data thy name =
-  case try (Inductive.the_inductive (ProofContext.init thy)) name of
-    SOME (info as (_, result)) => 
-      let
-        fun is_intro_of intro =
-          let
-            val (const, _) = strip_comb (HOLogic.dest_Trueprop (concl_of intro))
-          in (fst (dest_Const const) = name) end;      
-        val intros = ind_set_codegen_preproc thy ((map (preprocess_intro thy))
-          (filter is_intro_of (#intrs result)))
-        val pre_elim = nth (#elims result) (find_index (fn s => s = name) (#names (fst info)))
-        val nparams = length (Inductive.params_of (#raw_induct result))
-        val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
-        val (intros, elim) = if null intros then noclause thy name elim else (intros, elim)
-      in
-        mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
-      end                                                                    
-  | NONE => error ("No such predicate: " ^ quote name)
-  
-(* updaters *)
-
-fun apfst3 f (x, y, z) =  (f x, y, z)
-fun apsnd3 f (x, y, z) =  (x, f y, z)
-fun aptrd3 f (x, y, z) =  (x, y, f z)
-
-fun add_predfun name mode data =
-  let
-    val add = (apsnd o apfst3 o cons) (mode, mk_predfun_data data)
-  in PredData.map (Graph.map_node name (map_pred_data add)) end
-
-fun is_inductive_predicate thy name =
-  is_some (try (Inductive.the_inductive (ProofContext.init thy)) name)
-
-fun depending_preds_of thy (key, value) =
-  let
-    val intros = (#intros o rep_pred_data) value
-  in
-    fold Term.add_const_names (map Thm.prop_of intros) []
-      |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate thy c orelse is_registered thy c))
-  end;
-    
-    
-(* code dependency graph *)    
-(*
-fun dependencies_of thy name =
-  let
-    val (intros, elim, nparams) = fetch_pred_data thy name 
-    val data = mk_pred_data ((intros, SOME elim, nparams), ([], [], []))
-    val keys = depending_preds_of thy intros
-  in
-    (data, keys)
-  end;
-*)
-(* guessing number of parameters *)
-fun find_indexes pred xs =
-  let
-    fun find is n [] = is
-      | find is n (x :: xs) = find (if pred x then (n :: is) else is) (n + 1) xs;
-  in rev (find [] 0 xs) end;
-
-fun is_predT (T as Type("fun", [_, _])) = (snd (strip_type T) = HOLogic.boolT)
-  | is_predT _ = false
-  
-fun guess_nparams T =
-  let
-    val argTs = binder_types T
-    val nparams = fold Integer.max
-      (map (fn x => x + 1) (find_indexes is_predT argTs)) 0
-  in nparams end;
-
-fun add_intro thm thy = let
-   val (name, T) = dest_Const (fst (strip_intro_concl 0 (prop_of thm)))
-   fun cons_intro gr =
-     case try (Graph.get_node gr) name of
-       SOME pred_data => Graph.map_node name (map_pred_data
-         (apfst (fn (intro, elim, nparams) => (thm::intro, elim, nparams)))) gr
-     | NONE =>
-       let
-         val nparams = the_default (guess_nparams T)  (try (#nparams o rep_pred_data o (fetch_pred_data thy)) name)
-       in Graph.new_node (name, mk_pred_data (([thm], NONE, nparams), ([], [], []))) gr end;
-  in PredData.map cons_intro thy end
-
-fun set_elim thm = let
-    val (name, _) = dest_Const (fst 
-      (strip_comb (HOLogic.dest_Trueprop (hd (prems_of thm)))))
-    fun set (intros, _, nparams) = (intros, SOME thm, nparams)  
-  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
-
-fun set_nparams name nparams = let
-    fun set (intros, elim, _ ) = (intros, elim, nparams) 
-  in PredData.map (Graph.map_node name (map_pred_data (apfst set))) end
-    
-fun register_predicate (pre_intros, pre_elim, nparams) thy = let
-    val (name, _) = dest_Const (fst (strip_intro_concl nparams (prop_of (hd pre_intros))))
-    (* preprocessing *)
-    val intros = ind_set_codegen_preproc thy (map (preprocess_intro thy) pre_intros)
-    val elim = singleton (ind_set_codegen_preproc thy) (preprocess_elim thy nparams pre_elim)
-  in
-    PredData.map
-      (Graph.new_node (name, mk_pred_data ((intros, SOME elim, nparams), ([], [], [])))) thy
-  end
-
-fun set_generator_name pred mode name = 
-  let
-    val set = (apsnd o apsnd3 o cons) (mode, mk_function_data (name, NONE))
-  in
-    PredData.map (Graph.map_node pred (map_pred_data set))
-  end
-
-fun set_sizelim_function_name pred mode name = 
-  let
-    val set = (apsnd o aptrd3 o cons) (mode, mk_function_data (name, NONE))
-  in
-    PredData.map (Graph.map_node pred (map_pred_data set))
-  end
-
-(** data structures for generic compilation for different monads **)
-
-(* maybe rename functions more generic:
-  mk_predT -> mk_monadT; dest_predT -> dest_monadT
-  mk_single -> mk_return (?)
-*)
-datatype compilation_funs = CompilationFuns of {
-  mk_predT : typ -> typ,
-  dest_predT : typ -> typ,
-  mk_bot : typ -> term,
-  mk_single : term -> term,
-  mk_bind : term * term -> term,
-  mk_sup : term * term -> term,
-  mk_if : term -> term,
-  mk_not : term -> term,
-(*  funT_of : mode -> typ -> typ, *)
-(*  mk_fun_of : theory -> (string * typ) -> mode -> term, *) 
-  mk_map : typ -> typ -> term -> term -> term,
-  lift_pred : term -> term
-};
-
-fun mk_predT (CompilationFuns funs) = #mk_predT funs
-fun dest_predT (CompilationFuns funs) = #dest_predT funs
-fun mk_bot (CompilationFuns funs) = #mk_bot funs
-fun mk_single (CompilationFuns funs) = #mk_single funs
-fun mk_bind (CompilationFuns funs) = #mk_bind funs
-fun mk_sup (CompilationFuns funs) = #mk_sup funs
-fun mk_if (CompilationFuns funs) = #mk_if funs
-fun mk_not (CompilationFuns funs) = #mk_not funs
-(*fun funT_of (CompilationFuns funs) = #funT_of funs*)
-(*fun mk_fun_of (CompilationFuns funs) = #mk_fun_of funs*)
-fun mk_map (CompilationFuns funs) = #mk_map funs
-fun lift_pred (CompilationFuns funs) = #lift_pred funs
-
-fun funT_of compfuns (iss, is) T =
-  let
-    val Ts = binder_types T
-    val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
-    val paramTs' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss paramTs 
-  in
-    (paramTs' @ inargTs) ---> (mk_predT compfuns (mk_tupleT outargTs))
-  end;
-
-fun sizelim_funT_of compfuns (iss, is) T =
-  let
-    val Ts = binder_types T
-    val (paramTs, (inargTs, outargTs)) = split_mode (iss, is) Ts
-    val paramTs' = map2 (fn SOME is => sizelim_funT_of compfuns ([], is) | NONE => I) iss paramTs 
-  in
-    (paramTs' @ inargTs @ [@{typ "code_numeral"}]) ---> (mk_predT compfuns (mk_tupleT outargTs))
-  end;  
-
-fun mk_fun_of compfuns thy (name, T) mode = 
-  Const (predfun_name_of thy name mode, funT_of compfuns mode T)
-
-fun mk_sizelim_fun_of compfuns thy (name, T) mode =
-  Const (sizelim_function_name_of thy name mode, sizelim_funT_of compfuns mode T)
-  
-fun mk_generator_of compfuns thy (name, T) mode = 
-  Const (generator_name_of thy name mode, sizelim_funT_of compfuns mode T)
-
-
-structure PredicateCompFuns =
-struct
-
-fun mk_predT T = Type (@{type_name "Predicate.pred"}, [T])
-
-fun dest_predT (Type (@{type_name "Predicate.pred"}, [T])) = T
-  | dest_predT T = raise TYPE ("dest_predT", [T], []);
-
-fun mk_bot T = Const (@{const_name Orderings.bot}, mk_predT T);
-
-fun mk_single t =
-  let val T = fastype_of t
-  in Const(@{const_name Predicate.single}, T --> mk_predT T) $ t end;
-
-fun mk_bind (x, f) =
-  let val T as Type ("fun", [_, U]) = fastype_of f
-  in
-    Const (@{const_name Predicate.bind}, fastype_of x --> T --> U) $ x $ f
-  end;
-
-val mk_sup = HOLogic.mk_binop @{const_name sup};
-
-fun mk_if cond = Const (@{const_name Predicate.if_pred},
-  HOLogic.boolT --> mk_predT HOLogic.unitT) $ cond;
-
-fun mk_not t = let val T = mk_predT HOLogic.unitT
-  in Const (@{const_name Predicate.not_pred}, T --> T) $ t end
-
-fun mk_Enum f =
-  let val T as Type ("fun", [T', _]) = fastype_of f
-  in
-    Const (@{const_name Predicate.Pred}, T --> mk_predT T') $ f    
-  end;
-
-fun mk_Eval (f, x) =
-  let
-    val T = fastype_of x
-  in
-    Const (@{const_name Predicate.eval}, mk_predT T --> T --> HOLogic.boolT) $ f $ x
-  end;
-
-fun mk_map T1 T2 tf tp = Const (@{const_name Predicate.map},
-  (T1 --> T2) --> mk_predT T1 --> mk_predT T2) $ tf $ tp;
-
-val lift_pred = I
-
-val compfuns = CompilationFuns {mk_predT = mk_predT, dest_predT = dest_predT, mk_bot = mk_bot,
-  mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
-  mk_map = mk_map, lift_pred = lift_pred};
-
-end;
-
-(* termify_code:
-val termT = Type ("Code_Evaluation.term", []);
-fun termifyT T = HOLogic.mk_prodT (T, HOLogic.unitT --> termT)
-*)
-(*
-fun lift_random random =
-  let
-    val T = dest_randomT (fastype_of random)
-  in
-    mk_scomp (random,
-      mk_fun_comp (HOLogic.pair_const (PredicateCompFuns.mk_predT T) @{typ Random.seed},
-        mk_fun_comp (Const (@{const_name Predicate.single}, T --> (PredicateCompFuns.mk_predT T)),
-          Const (@{const_name "fst"}, HOLogic.mk_prodT (T, @{typ "unit => term"}) --> T)))) 
-  end;
-*)
- 
-structure RPredCompFuns =
-struct
-
-fun mk_rpredT T =
-  @{typ "Random.seed"} --> HOLogic.mk_prodT (PredicateCompFuns.mk_predT T, @{typ "Random.seed"})
-
-fun dest_rpredT (Type ("fun", [_,
-  Type (@{type_name "*"}, [Type (@{type_name "Predicate.pred"}, [T]), _])])) = T
-  | dest_rpredT T = raise TYPE ("dest_rpredT", [T], []); 
-
-fun mk_bot T = Const(@{const_name RPred.bot}, mk_rpredT T)
-
-fun mk_single t =
-  let
-    val T = fastype_of t
-  in
-    Const (@{const_name RPred.return}, T --> mk_rpredT T) $ t
-  end;
-
-fun mk_bind (x, f) =
-  let
-    val T as (Type ("fun", [_, U])) = fastype_of f
-  in
-    Const (@{const_name RPred.bind}, fastype_of x --> T --> U) $ x $ f
-  end
-
-val mk_sup = HOLogic.mk_binop @{const_name RPred.supp}
-
-fun mk_if cond = Const (@{const_name RPred.if_rpred},
-  HOLogic.boolT --> mk_rpredT HOLogic.unitT) $ cond;
-
-fun mk_not t = error "Negation is not defined for RPred"
-
-fun mk_map t = error "FIXME" (*FIXME*)
-
-fun lift_pred t =
-  let
-    val T = PredicateCompFuns.dest_predT (fastype_of t)
-    val lift_predT = PredicateCompFuns.mk_predT T --> mk_rpredT T 
-  in
-    Const (@{const_name "RPred.lift_pred"}, lift_predT) $ t  
-  end;
-
-val compfuns = CompilationFuns {mk_predT = mk_rpredT, dest_predT = dest_rpredT, mk_bot = mk_bot,
-    mk_single = mk_single, mk_bind = mk_bind, mk_sup = mk_sup, mk_if = mk_if, mk_not = mk_not,
-    mk_map = mk_map, lift_pred = lift_pred};
-
-end;
-(* for external use with interactive mode *)
-val rpred_compfuns = RPredCompFuns.compfuns;
-
-fun lift_random random =
-  let
-    val T = dest_randomT (fastype_of random)
-  in
-    Const (@{const_name lift_random}, (@{typ Random.seed} -->
-      HOLogic.mk_prodT (HOLogic.mk_prodT (T, @{typ "unit => term"}), @{typ Random.seed})) --> 
-      RPredCompFuns.mk_rpredT T) $ random
-  end;
- 
-(* Mode analysis *)
-
-(*** check if a term contains only constructor functions ***)
-fun is_constrt thy =
-  let
-    val cnstrs = flat (maps
-      (map (fn (_, (Tname, _, cs)) => map (apsnd (rpair Tname o length)) cs) o #descr o snd)
-      (Symtab.dest (Datatype.get_all thy)));
-    fun check t = (case strip_comb t of
-        (Free _, []) => true
-      | (Const (s, T), ts) => (case (AList.lookup (op =) cnstrs s, body_type T) of
-            (SOME (i, Tname), Type (Tname', _)) => length ts = i andalso Tname = Tname' andalso forall check ts
-          | _ => false)
-      | _ => false)
-  in check end;
-
-(*** check if a type is an equality type (i.e. doesn't contain fun)
-  FIXME this is only an approximation ***)
-fun is_eqT (Type (s, Ts)) = s <> "fun" andalso forall is_eqT Ts
-  | is_eqT _ = true;
-
-fun term_vs tm = fold_aterms (fn Free (x, T) => cons x | _ => I) tm [];
-val terms_vs = distinct (op =) o maps term_vs;
-
-(** collect all Frees in a term (with duplicates!) **)
-fun term_vTs tm =
-  fold_aterms (fn Free xT => cons xT | _ => I) tm [];
-
-(*FIXME this function should not be named merge... make it local instead*)
-fun merge xs [] = xs
-  | merge [] ys = ys
-  | merge (x::xs) (y::ys) = if length x >= length y then x::merge xs (y::ys)
-      else y::merge (x::xs) ys;
-
-fun subsets i j = if i <= j then
-       let val is = subsets (i+1) j
-       in merge (map (fn ks => i::ks) is) is end
-     else [[]];
-     
-(* FIXME: should be in library - map_prod *)
-fun cprod ([], ys) = []
-  | cprod (x :: xs, ys) = map (pair x) ys @ cprod (xs, ys);
-
-fun cprods xss = List.foldr (map op :: o cprod) [[]] xss;
-
-
-  
-(*TODO: cleanup function and put together with modes_of_term *)
-(*
-fun modes_of_param default modes t = let
-    val (vs, t') = strip_abs t
-    val b = length vs
-    fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
-        let
-          val (args1, args2) =
-            if length args < length iss then
-              error ("Too few arguments for inductive predicate " ^ name)
-            else chop (length iss) args;
-          val k = length args2;
-          val perm = map (fn i => (find_index_eq (Bound (b - i)) args2) + 1)
-            (1 upto b)  
-          val partial_mode = (1 upto k) \\ perm
-        in
-          if not (partial_mode subset is) then [] else
-          let
-            val is' = 
-            (fold_index (fn (i, j) => if j mem is then cons (i + 1) else I) perm [])
-            |> fold (fn i => if i > k then cons (i - k + b) else I) is
-              
-           val res = map (fn x => Mode (m, is', x)) (cprods (map
-            (fn (NONE, _) => [NONE]
-              | (SOME js, arg) => map SOME (filter
-                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
-                    (iss ~~ args1)))
-          in res end
-        end)) (AList.lookup op = modes name)
-  in case strip_comb t' of
-    (Const (name, _), args) => the_default default (mk_modes name args)
-    | (Var ((name, _), _), args) => the (mk_modes name args)
-    | (Free (name, _), args) => the (mk_modes name args)
-    | _ => default end
-  
-and
-*)
-fun modes_of_term modes t =
-  let
-    val ks = 1 upto length (binder_types (fastype_of t));
-    val default = [Mode (([], ks), ks, [])];
-    fun mk_modes name args = Option.map (maps (fn (m as (iss, is)) =>
-        let
-          val (args1, args2) =
-            if length args < length iss then
-              error ("Too few arguments for inductive predicate " ^ name)
-            else chop (length iss) args;
-          val k = length args2;
-          val prfx = 1 upto k
-        in
-          if not (is_prefix op = prfx is) then [] else
-          let val is' = map (fn i => i - k) (List.drop (is, k))
-          in map (fn x => Mode (m, is', x)) (cprods (map
-            (fn (NONE, _) => [NONE]
-              | (SOME js, arg) => map SOME (filter
-                  (fn Mode (_, js', _) => js=js') (modes_of_term modes arg)))
-                    (iss ~~ args1)))
-          end
-        end)) (AList.lookup op = modes name)
-
-  in
-    case strip_comb (Envir.eta_contract t) of
-      (Const (name, _), args) => the_default default (mk_modes name args)
-    | (Var ((name, _), _), args) => the (mk_modes name args)
-    | (Free (name, _), args) => the (mk_modes name args)
-    | (Abs _, []) => error "Abs at param position" (* modes_of_param default modes t *)
-    | _ => default
-  end
-  
-fun select_mode_prem thy modes vs ps =
-  find_first (is_some o snd) (ps ~~ map
-    (fn Prem (us, t) => find_first (fn Mode (_, is, _) =>
-          let
-            val (in_ts, out_ts) = split_smode is us;
-            val (out_ts', in_ts') = List.partition (is_constrt thy) out_ts;
-            val vTs = maps term_vTs out_ts';
-            val dupTs = map snd (duplicates (op =) vTs) @
-              map_filter (AList.lookup (op =) vTs) vs;
-          in
-            subset (op =) (terms_vs (in_ts @ in_ts'), vs) andalso
-            forall (is_eqT o fastype_of) in_ts' andalso
-            subset (op =) (term_vs t, vs) andalso
-            forall is_eqT dupTs
-          end)
-            (modes_of_term modes t handle Option =>
-               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
-      | Negprem (us, t) => find_first (fn Mode (_, is, _) =>
-            length us = length is andalso
-            subset (op =) (terms_vs us, vs) andalso
-            subset (op =) (term_vs t, vs)
-            (modes_of_term modes t handle Option =>
-               error ("Bad predicate: " ^ Syntax.string_of_term_global thy t))
-      | Sidecond t => if subset (op =) (term_vs t, vs) then SOME (Mode (([], []), [], []))
-          else NONE
-      ) ps);
-
-fun fold_prem f (Prem (args, _)) = fold f args
-  | fold_prem f (Negprem (args, _)) = fold f args
-  | fold_prem f (Sidecond t) = f t
-
-fun all_subsets [] = [[]]
-  | all_subsets (x::xs) = let val xss' = all_subsets xs in xss' @ (map (cons x) xss') end
-
-fun generator vTs v = 
-  let
-    val T = the (AList.lookup (op =) vTs v)
-  in
-    (Generator (v, T), Mode (([], []), [], []))
-  end;
-
-fun gen_prem (Prem (us, t)) = GeneratorPrem (us, t) 
-  | gen_prem _ = error "gen_prem : invalid input for gen_prem"
-
-fun param_gen_prem param_vs (p as Prem (us, t as Free (v, _))) =
-  if member (op =) param_vs v then
-    GeneratorPrem (us, t)
-  else p  
-  | param_gen_prem param_vs p = p
-  
-fun check_mode_clause with_generator thy param_vs modes gen_modes (iss, is) (ts, ps) =
-  let
-    val modes' = modes @ map_filter
-      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-        (param_vs ~~ iss);
-    val gen_modes' = gen_modes @ map_filter
-      (fn (_, NONE) => NONE | (v, SOME js) => SOME (v, [([], js)]))
-        (param_vs ~~ iss);  
-    val vTs = distinct (op =) ((fold o fold_prem) Term.add_frees ps (fold Term.add_frees ts []))
-    val prem_vs = distinct (op =) ((fold o fold_prem) Term.add_free_names ps [])
-    fun check_mode_prems acc_ps vs [] = SOME (acc_ps, vs)
-      | check_mode_prems acc_ps vs ps = (case select_mode_prem thy modes' vs ps of
-          NONE =>
-            (if with_generator then
-              (case select_mode_prem thy gen_modes' vs ps of
-                  SOME (p, SOME mode) => check_mode_prems ((gen_prem p, mode) :: acc_ps) 
-                  (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
-                  (filter_out (equal p) ps)
-                | NONE =>
-                  let 
-                    val all_generator_vs = all_subsets (prem_vs \\ vs) |> sort (int_ord o (pairself length))
-                  in
-                    case (find_first (fn generator_vs => is_some
-                      (select_mode_prem thy modes' (union (op =) vs generator_vs) ps)) all_generator_vs) of
-                      SOME generator_vs => check_mode_prems ((map (generator vTs) generator_vs) @ acc_ps)
-                        (union (op =) vs generator_vs) ps
-                    | NONE => NONE
-                  end)
-            else
-              NONE)
-        | SOME (p, SOME mode) => check_mode_prems ((if with_generator then param_gen_prem param_vs p else p, mode) :: acc_ps) 
-            (case p of Prem (us, _) => union (op =) vs (terms_vs us) | _ => vs)
-            (filter_out (equal p) ps))
-    val (in_ts, in_ts') = List.partition (is_constrt thy) (fst (split_smode is ts));
-    val in_vs = terms_vs in_ts;
-    val concl_vs = terms_vs ts
-  in
-    if forall is_eqT (map snd (duplicates (op =) (maps term_vTs in_ts))) andalso
-    forall (is_eqT o fastype_of) in_ts' then
-      case check_mode_prems [] (union (op =) param_vs in_vs) ps of
-         NONE => NONE
-       | SOME (acc_ps, vs) =>
-         if with_generator then
-           SOME (ts, (rev acc_ps) @ (map (generator vTs) (concl_vs \\ vs))) 
-         else
-           if subset (op =) (concl_vs, vs) then SOME (ts, rev acc_ps) else NONE
-    else NONE
-  end;
-
-fun check_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
-  let val SOME rs = AList.lookup (op =) preds p
-  in (p, filter (fn m => case find_index
-    (is_none o check_mode_clause with_generator thy param_vs modes gen_modes m) rs of
-      ~1 => true
-    | i => (tracing ("Clause " ^ string_of_int (i + 1) ^ " of " ^
-      p ^ " violates mode " ^ string_of_mode m); false)) ms)
-  end;
-
-fun get_modes_pred with_generator thy param_vs preds modes gen_modes (p, ms) =
-  let
-    val SOME rs = AList.lookup (op =) preds p 
-  in
-    (p, map (fn m =>
-      (m, map (the o check_mode_clause with_generator thy param_vs modes gen_modes m) rs)) ms)
-  end;
-  
-fun fixp f (x : (string * mode list) list) =
-  let val y = f x
-  in if x = y then x else fixp f y end;
-
-fun modes_of_arities arities =
-  (map (fn (s, (ks, k)) => (s, cprod (cprods (map
-            (fn NONE => [NONE]
-              | SOME k' => map SOME (subsets 1 k')) ks),
-            subsets 1 k))) arities)
-  
-fun infer_modes with_generator thy extra_modes arities param_vs preds =
-  let
-    val modes =
-      fixp (fn modes =>
-        map (check_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes)
-          (modes_of_arities arities)
-  in
-    map (get_modes_pred with_generator thy param_vs preds (modes @ extra_modes) []) modes
-  end;
-
-fun remove_from rem [] = []
-  | remove_from rem ((k, vs) :: xs) =
-    (case AList.lookup (op =) rem k of
-      NONE => (k, vs)
-    | SOME vs' => (k, vs \\ vs'))
-    :: remove_from rem xs
-    
-fun infer_modes_with_generator thy extra_modes arities param_vs preds =
-  let
-    val prednames = map fst preds
-    val extra_modes = all_modes_of thy
-    val gen_modes = all_generator_modes_of thy
-      |> filter_out (fn (name, _) => member (op =) prednames name)
-    val starting_modes = remove_from extra_modes (modes_of_arities arities) 
-    val modes =
-      fixp (fn modes =>
-        map (check_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes)
-         starting_modes 
-  in
-    map (get_modes_pred true thy param_vs preds extra_modes (gen_modes @ modes)) modes
-  end;
-
-(* term construction *)
-
-fun mk_v (names, vs) s T = (case AList.lookup (op =) vs s of
-      NONE => (Free (s, T), (names, (s, [])::vs))
-    | SOME xs =>
-        let
-          val s' = Name.variant names s;
-          val v = Free (s', T)
-        in
-          (v, (s'::names, AList.update (op =) (s, v::xs) vs))
-        end);
-
-fun distinct_v (Free (s, T)) nvs = mk_v nvs s T
-  | distinct_v (t $ u) nvs =
-      let
-        val (t', nvs') = distinct_v t nvs;
-        val (u', nvs'') = distinct_v u nvs';
-      in (t' $ u', nvs'') end
-  | distinct_v x nvs = (x, nvs);
-
-fun compile_match thy compfuns eqs eqs' out_ts success_t =
-  let
-    val eqs'' = maps mk_eq eqs @ eqs'
-    val names = fold Term.add_free_names (success_t :: eqs'' @ out_ts) [];
-    val name = Name.variant names "x";
-    val name' = Name.variant (name :: names) "y";
-    val T = mk_tupleT (map fastype_of out_ts);
-    val U = fastype_of success_t;
-    val U' = dest_predT compfuns U;
-    val v = Free (name, T);
-    val v' = Free (name', T);
-  in
-    lambda v (fst (Datatype.make_case
-      (ProofContext.init thy) false [] v
-      [(mk_tuple out_ts,
-        if null eqs'' then success_t
-        else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $
-          foldr1 HOLogic.mk_conj eqs'' $ success_t $
-            mk_bot compfuns U'),
-       (v', mk_bot compfuns U')]))
-  end;
-
-(*FIXME function can be removed*)
-fun mk_funcomp f t =
-  let
-    val names = Term.add_free_names t [];
-    val Ts = binder_types (fastype_of t);
-    val vs = map Free
-      (Name.variant_list names (replicate (length Ts) "x") ~~ Ts)
-  in
-    fold_rev lambda vs (f (list_comb (t, vs)))
-  end;
-(*
-fun compile_param_ext thy compfuns modes (NONE, t) = t
-  | compile_param_ext thy compfuns modes (m as SOME (Mode ((iss, is'), is, ms)), t) =
-      let
-        val (vs, u) = strip_abs t
-        val (ivs, ovs) = split_mode is vs    
-        val (f, args) = strip_comb u
-        val (params, args') = chop (length ms) args
-        val (inargs, outargs) = split_mode is' args'
-        val b = length vs
-        val perm = map (fn i => (find_index_eq (Bound (b - i)) args') + 1) (1 upto b)
-        val outp_perm =
-          snd (split_mode is perm)
-          |> map (fn i => i - length (filter (fn x => x < i) is'))
-        val names = [] -- TODO
-        val out_names = Name.variant_list names (replicate (length outargs) "x")
-        val f' = case f of
-            Const (name, T) =>
-              if AList.defined op = modes name then
-                mk_predfun_of thy compfuns (name, T) (iss, is')
-              else error "compile param: Not an inductive predicate with correct mode"
-          | Free (name, T) => Free (name, param_funT_of compfuns T (SOME is'))
-        val outTs = dest_tupleT (dest_predT compfuns (body_type (fastype_of f')))
-        val out_vs = map Free (out_names ~~ outTs)
-        val params' = map (compile_param thy modes) (ms ~~ params)
-        val f_app = list_comb (f', params' @ inargs)
-        val single_t = (mk_single compfuns (mk_tuple (map (fn i => nth out_vs (i - 1)) outp_perm)))
-        val match_t = compile_match thy compfuns [] [] out_vs single_t
-      in list_abs (ivs,
-        mk_bind compfuns (f_app, match_t))
-      end
-  | compile_param_ext _ _ _ _ = error "compile params"
-*)
-
-fun compile_param size thy compfuns (NONE, t) = t
-  | compile_param size thy compfuns (m as SOME (Mode ((iss, is'), is, ms)), t) =
-   let
-     val (f, args) = strip_comb (Envir.eta_contract t)
-     val (params, args') = chop (length ms) args
-     val params' = map (compile_param size thy compfuns) (ms ~~ params)
-     val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
-     val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of
-     val f' =
-       case f of
-         Const (name, T) =>
-           mk_fun_of compfuns thy (name, T) (iss, is')
-       | Free (name, T) => Free (name, funT_of compfuns (iss, is') T)
-       | _ => error ("PredicateCompiler: illegal parameter term")
-   in list_comb (f', params' @ args') end
-   
-fun compile_expr size thy ((Mode (mode, is, ms)), t) =
-  case strip_comb t of
-    (Const (name, T), params) =>
-       let
-         val params' = map (compile_param size thy PredicateCompFuns.compfuns) (ms ~~ params)
-         val mk_fun_of = case size of NONE => mk_fun_of | SOME _ => mk_sizelim_fun_of
-       in
-         list_comb (mk_fun_of PredicateCompFuns.compfuns thy (name, T) mode, params')
-       end
-  | (Free (name, T), args) =>
-       let 
-         val funT_of = case size of NONE => funT_of | SOME _ => sizelim_funT_of 
-       in
-         list_comb (Free (name, funT_of PredicateCompFuns.compfuns ([], is) T), args)
-       end;
-       
-fun compile_gen_expr size thy compfuns ((Mode (mode, is, ms)), t) =
-  case strip_comb t of
-    (Const (name, T), params) =>
-      let
-        val params' = map (compile_param size thy compfuns) (ms ~~ params)
-      in
-        list_comb (mk_generator_of compfuns thy (name, T) mode, params')
-      end
-    | (Free (name, T), args) =>
-      list_comb (Free (name, sizelim_funT_of RPredCompFuns.compfuns ([], is) T), args)
-          
-(** specific rpred functions -- move them to the correct place in this file *)
-
-(* uncommented termify code; causes more trouble than expected at first *) 
-(*
-fun mk_valtermify_term (t as Const (c, T)) = HOLogic.mk_prod (t, Abs ("u", HOLogic.unitT, HOLogic.reflect_term t))
-  | mk_valtermify_term (Free (x, T)) = Free (x, termifyT T) 
-  | mk_valtermify_term (t1 $ t2) =
-    let
-      val T = fastype_of t1
-      val (T1, T2) = dest_funT T
-      val t1' = mk_valtermify_term t1
-      val t2' = mk_valtermify_term t2
-    in
-      Const ("Code_Evaluation.valapp", termifyT T --> termifyT T1 --> termifyT T2) $ t1' $ t2'
-    end
-  | mk_valtermify_term _ = error "Not a valid term for mk_valtermify_term"
-*)
-
-fun compile_clause compfuns size final_term thy all_vs param_vs (iss, is) inp (ts, moded_ps) =
-  let
-    fun check_constrt t (names, eqs) =
-      if is_constrt thy t then (t, (names, eqs)) else
-        let
-          val s = Name.variant names "x";
-          val v = Free (s, fastype_of t)
-        in (v, (s::names, HOLogic.mk_eq (v, t)::eqs)) end;
-
-    val (in_ts, out_ts) = split_smode is ts;
-    val (in_ts', (all_vs', eqs)) =
-      fold_map check_constrt in_ts (all_vs, []);
-
-    fun compile_prems out_ts' vs names [] =
-          let
-            val (out_ts'', (names', eqs')) =
-              fold_map check_constrt out_ts' (names, []);
-            val (out_ts''', (names'', constr_vs)) = fold_map distinct_v
-              out_ts'' (names', map (rpair []) vs);
-          in
-          (* termify code:
-            compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
-              (mk_single compfuns (mk_tuple (map mk_valtermify_term out_ts)))
-           *)
-            compile_match thy compfuns constr_vs (eqs @ eqs') out_ts'''
-              (final_term out_ts)
-          end
-      | compile_prems out_ts vs names ((p, mode as Mode ((_, is), _, _)) :: ps) =
-          let
-            val vs' = distinct (op =) (flat (vs :: map term_vs out_ts));
-            val (out_ts', (names', eqs)) =
-              fold_map check_constrt out_ts (names, [])
-            val (out_ts'', (names'', constr_vs')) = fold_map distinct_v
-              out_ts' ((names', map (rpair []) vs))
-            val (compiled_clause, rest) = case p of
-               Prem (us, t) =>
-                 let
-                   val (in_ts, out_ts''') = split_smode is us;
-                   val args = case size of
-                     NONE => in_ts
-                   | SOME size_t => in_ts @ [size_t]
-                   val u = lift_pred compfuns
-                     (list_comb (compile_expr size thy (mode, t), args))                     
-                   val rest = compile_prems out_ts''' vs' names'' ps
-                 in
-                   (u, rest)
-                 end
-             | Negprem (us, t) =>
-                 let
-                   val (in_ts, out_ts''') = split_smode is us
-                   val u = lift_pred compfuns
-                     (mk_not PredicateCompFuns.compfuns (list_comb (compile_expr NONE thy (mode, t), in_ts)))
-                   val rest = compile_prems out_ts''' vs' names'' ps
-                 in
-                   (u, rest)
-                 end
-             | Sidecond t =>
-                 let
-                   val rest = compile_prems [] vs' names'' ps;
-                 in
-                   (mk_if compfuns t, rest)
-                 end
-             | GeneratorPrem (us, t) =>
-                 let
-                   val (in_ts, out_ts''') = split_smode is us;
-                   val args = case size of
-                     NONE => in_ts
-                   | SOME size_t => in_ts @ [size_t]
-                   val u = list_comb (compile_gen_expr size thy compfuns (mode, t), args)
-                   val rest = compile_prems out_ts''' vs' names'' ps
-                 in
-                   (u, rest)
-                 end
-             | Generator (v, T) =>
-                 let
-                   val u = lift_random (HOLogic.mk_random T @{term "1::code_numeral"})
-                   val rest = compile_prems [Free (v, T)]  vs' names'' ps;
-                 in
-                   (u, rest)
-                 end
-          in
-            compile_match thy compfuns constr_vs' eqs out_ts'' 
-              (mk_bind compfuns (compiled_clause, rest))
-          end
-    val prem_t = compile_prems in_ts' param_vs all_vs' moded_ps;
-  in
-    mk_bind compfuns (mk_single compfuns inp, prem_t)
-  end
-
-fun compile_pred compfuns mk_fun_of use_size thy all_vs param_vs s T mode moded_cls =
-  let
-    val (Ts1, (Us1, Us2)) = split_mode mode (binder_types T)
-    val funT_of = if use_size then sizelim_funT_of else funT_of 
-    val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) (fst mode) Ts1
-    val xnames = Name.variant_list (all_vs @ param_vs)
-      (map (fn i => "x" ^ string_of_int i) (snd mode));
-    val size_name = Name.variant (all_vs @ param_vs @ xnames) "size"
-    (* termify code: val xs = map2 (fn s => fn T => Free (s, termifyT T)) xnames Us1; *)
-    val xs = map2 (fn s => fn T => Free (s, T)) xnames Us1;
-    val params = map2 (fn s => fn T => Free (s, T)) param_vs Ts1'
-    val size = Free (size_name, @{typ "code_numeral"})
-    val decr_size =
-      if use_size then
-        SOME (Const ("HOL.minus_class.minus", @{typ "code_numeral => code_numeral => code_numeral"})
-          $ size $ Const ("HOL.one_class.one", @{typ "Code_Numeral.code_numeral"}))
-      else
-        NONE
-    val cl_ts =
-      map (compile_clause compfuns decr_size (fn out_ts => mk_single compfuns (mk_tuple out_ts))
-        thy all_vs param_vs mode (mk_tuple xs)) moded_cls;
-    val t = foldr1 (mk_sup compfuns) cl_ts
-    val T' = mk_predT compfuns (mk_tupleT Us2)
-    val size_t = Const (@{const_name "If"}, @{typ bool} --> T' --> T' --> T')
-      $ HOLogic.mk_eq (size, @{term "0 :: code_numeral"})
-      $ mk_bot compfuns (dest_predT compfuns T') $ t
-    val fun_const = mk_fun_of compfuns thy (s, T) mode
-    val eq = if use_size then
-      (list_comb (fun_const, params @ xs @ [size]), size_t)
-    else
-      (list_comb (fun_const, params @ xs), t)
-  in
-    HOLogic.mk_Trueprop (HOLogic.mk_eq eq)
-  end;
-  
-(* special setup for simpset *)                  
-val HOL_basic_ss' = HOL_basic_ss setSolver 
-  (mk_solver "all_tac_solver" (fn _ => fn _ => all_tac))
-
-(* Definition of executable functions and their intro and elim rules *)
-
-fun print_arities arities = tracing ("Arities:\n" ^
-  cat_lines (map (fn (s, (ks, k)) => s ^ ": " ^
-    space_implode " -> " (map
-      (fn NONE => "X" | SOME k' => string_of_int k')
-        (ks @ [SOME k]))) arities));
-
-fun mk_Eval_of ((x, T), NONE) names = (x, names)
-  | mk_Eval_of ((x, T), SOME mode) names = let
-  val Ts = binder_types T
-  val argnames = Name.variant_list names
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-  val args = map Free (argnames ~~ Ts)
-  val (inargs, outargs) = split_smode mode args
-  val r = PredicateCompFuns.mk_Eval (list_comb (x, inargs), mk_tuple outargs)
-  val t = fold_rev lambda args r 
-in
-  (t, argnames @ names)
-end;
-
-fun create_intro_elim_rule (mode as (iss, is)) defthm mode_id funT pred thy =
-let
-  val Ts = binder_types (fastype_of pred)
-  val funtrm = Const (mode_id, funT)
-  val argnames = Name.variant_list []
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-  val (Ts1, Ts2) = chop (length iss) Ts;
-  val Ts1' = map2 (fn NONE => I | SOME is => funT_of (PredicateCompFuns.compfuns) ([], is)) iss Ts1
-  val args = map Free (argnames ~~ (Ts1' @ Ts2))
-  val (params, ioargs) = chop (length iss) args
-  val (inargs, outargs) = split_smode is ioargs
-  val param_names = Name.variant_list argnames
-    (map (fn i => "p" ^ string_of_int i) (1 upto (length iss)))
-  val param_vs = map Free (param_names ~~ Ts1)
-  val (params', names) = fold_map mk_Eval_of ((params ~~ Ts1) ~~ iss) []
-  val predpropI = HOLogic.mk_Trueprop (list_comb (pred, param_vs @ ioargs))
-  val predpropE = HOLogic.mk_Trueprop (list_comb (pred, params' @ ioargs))
-  val param_eqs = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (param_vs ~~ params')
-  val funargs = params @ inargs
-  val funpropE = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
-                  if null outargs then Free("y", HOLogic.unitT) else mk_tuple outargs))
-  val funpropI = HOLogic.mk_Trueprop (PredicateCompFuns.mk_Eval (list_comb (funtrm, funargs),
-                   mk_tuple outargs))
-  val introtrm = Logic.list_implies (predpropI :: param_eqs, funpropI)
-  val simprules = [defthm, @{thm eval_pred},
-                   @{thm "split_beta"}, @{thm "fst_conv"}, @{thm "snd_conv"}]
-  val unfolddef_tac = Simplifier.asm_full_simp_tac (HOL_basic_ss addsimps simprules) 1
-  val introthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y"]) [] introtrm (fn {...} => unfolddef_tac)
-  val P = HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT));
-  val elimtrm = Logic.list_implies ([funpropE, Logic.mk_implies (predpropE, P)], P)
-  val elimthm = Goal.prove (ProofContext.init thy) (argnames @ param_names @ ["y", "P"]) [] elimtrm (fn {...} => unfolddef_tac)
-in 
-  (introthm, elimthm)
-end;
-
-fun create_constname_of_mode thy prefix name mode = 
-  let
-    fun string_of_mode mode = if null mode then "0"
-      else space_implode "_" (map string_of_int mode)
-    val HOmode = space_implode "_and_"
-      (fold (fn NONE => I | SOME mode => cons (string_of_mode mode)) (fst mode) [])
-  in
-    (Sign.full_bname thy (prefix ^ (Long_Name.base_name name))) ^
-      (if HOmode = "" then "_" else "_for_" ^ HOmode ^ "_yields_") ^ (string_of_mode (snd mode))
-  end;
-  
-fun create_definitions preds (name, modes) thy =
-  let
-    val compfuns = PredicateCompFuns.compfuns
-    val T = AList.lookup (op =) preds name |> the
-    fun create_definition (mode as (iss, is)) thy = let
-      val mode_cname = create_constname_of_mode thy "" name mode
-      val mode_cbasename = Long_Name.base_name mode_cname
-      val Ts = binder_types T
-      val (Ts1, Ts2) = chop (length iss) Ts
-      val (Us1, Us2) =  split_smode is Ts2
-      val Ts1' = map2 (fn NONE => I | SOME is => funT_of compfuns ([], is)) iss Ts1
-      val funT = (Ts1' @ Us1) ---> (mk_predT compfuns (mk_tupleT Us2))
-      val names = Name.variant_list []
-        (map (fn i => "x" ^ string_of_int i) (1 upto (length Ts)));
-      val xs = map Free (names ~~ (Ts1' @ Ts2));                   
-      val (xparams, xargs) = chop (length iss) xs;
-      val (xins, xouts) = split_smode is xargs 
-      val (xparams', names') = fold_map mk_Eval_of ((xparams ~~ Ts1) ~~ iss) names
-      fun mk_split_lambda [] t = lambda (Free (Name.variant names' "x", HOLogic.unitT)) t
-        | mk_split_lambda [x] t = lambda x t
-        | mk_split_lambda xs t =
-        let
-          fun mk_split_lambda' (x::y::[]) t = HOLogic.mk_split (lambda x (lambda y t))
-            | mk_split_lambda' (x::xs) t = HOLogic.mk_split (lambda x (mk_split_lambda' xs t))
-        in
-          mk_split_lambda' xs t
-        end;
-      val predterm = PredicateCompFuns.mk_Enum (mk_split_lambda xouts
-        (list_comb (Const (name, T), xparams' @ xargs)))
-      val lhs = list_comb (Const (mode_cname, funT), xparams @ xins)
-      val def = Logic.mk_equals (lhs, predterm)
-      val ([definition], thy') = thy |>
-        Sign.add_consts_i [(Binding.name mode_cbasename, funT, NoSyn)] |>
-        PureThy.add_defs false [((Binding.name (mode_cbasename ^ "_def"), def), [])]
-      val (intro, elim) =
-        create_intro_elim_rule mode definition mode_cname funT (Const (name, T)) thy'
-      in thy' |> add_predfun name mode (mode_cname, definition, intro, elim)
-        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "I"), intro) |> snd
-        |> PureThy.store_thm (Binding.name (mode_cbasename ^ "E"), elim)  |> snd
-        |> Theory.checkpoint
-      end;
-  in
-    fold create_definition modes thy
-  end;
-
-fun sizelim_create_definitions preds (name, modes) thy =
-  let
-    val T = AList.lookup (op =) preds name |> the
-    fun create_definition mode thy =
-      let
-        val mode_cname = create_constname_of_mode thy "sizelim_" name mode
-        val funT = sizelim_funT_of PredicateCompFuns.compfuns mode T
-      in
-        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
-        |> set_sizelim_function_name name mode mode_cname 
-      end;
-  in
-    fold create_definition modes thy
-  end;
-    
-fun rpred_create_definitions preds (name, modes) thy =
-  let
-    val T = AList.lookup (op =) preds name |> the
-    fun create_definition mode thy =
-      let
-        val mode_cname = create_constname_of_mode thy "gen_" name mode
-        val funT = sizelim_funT_of RPredCompFuns.compfuns mode T
-      in
-        thy |> Sign.add_consts_i [(Binding.name (Long_Name.base_name mode_cname), funT, NoSyn)]
-        |> set_generator_name name mode mode_cname 
-      end;
-  in
-    fold create_definition modes thy
-  end;
-  
-(* Proving equivalence of term *)
-
-fun is_Type (Type _) = true
-  | is_Type _ = false
-
-(* returns true if t is an application of an datatype constructor *)
-(* which then consequently would be splitted *)
-(* else false *)
-fun is_constructor thy t =
-  if (is_Type (fastype_of t)) then
-    (case Datatype.get_info thy ((fst o dest_Type o fastype_of) t) of
-      NONE => false
-    | SOME info => (let
-      val constr_consts = maps (fn (_, (_, _, constrs)) => map fst constrs) (#descr info)
-      val (c, _) = strip_comb t
-      in (case c of
-        Const (name, _) => name mem_string constr_consts
-        | _ => false) end))
-  else false
-
-(* MAJOR FIXME:  prove_params should be simple
- - different form of introrule for parameters ? *)
-fun prove_param thy (NONE, t) = TRY (rtac @{thm refl} 1)
-  | prove_param thy (m as SOME (Mode (mode, is, ms)), t) =
-  let
-    val  (f, args) = strip_comb (Envir.eta_contract t)
-    val (params, _) = chop (length ms) args
-    val f_tac = case f of
-      Const (name, T) => simp_tac (HOL_basic_ss addsimps 
-         (@{thm eval_pred}::(predfun_definition_of thy name mode)::
-         @{thm "Product_Type.split_conv"}::[])) 1
-    | Free _ => TRY (rtac @{thm refl} 1)
-    | Abs _ => error "prove_param: No valid parameter term"
-  in
-    REPEAT_DETERM (etac @{thm thin_rl} 1)
-    THEN REPEAT_DETERM (rtac @{thm ext} 1)
-    THEN print_tac "prove_param"
-    THEN f_tac
-    THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map (prove_param thy) (ms ~~ params)))
-    THEN (REPEAT_DETERM (atac 1))
-  end
-
-fun prove_expr thy (Mode (mode, is, ms), t, us) (premposition : int) =
-  case strip_comb t of
-    (Const (name, T), args) =>  
-      let
-        val introrule = predfun_intro_of thy name mode
-        val (args1, args2) = chop (length ms) args
-      in
-        rtac @{thm bindI} 1
-        THEN print_tac "before intro rule:"
-        (* for the right assumption in first position *)
-        THEN rotate_tac premposition 1
-        THEN debug_tac (Display.string_of_thm (ProofContext.init thy) introrule)
-        THEN rtac introrule 1
-        THEN print_tac "after intro rule"
-        (* work with parameter arguments *)
-        THEN (atac 1)
-        THEN (print_tac "parameter goal")
-        THEN (EVERY (map (prove_param thy) (ms ~~ args1)))
-        THEN (REPEAT_DETERM (atac 1))
-      end
-  | _ => rtac @{thm bindI} 1 THEN atac 1
-
-fun SOLVED tac st = FILTER (fn st' => nprems_of st' = nprems_of st - 1) tac st; 
-
-fun SOLVEDALL tac st = FILTER (fn st' => nprems_of st' = 0) tac st
-
-fun prove_match thy (out_ts : term list) = let
-  fun get_case_rewrite t =
-    if (is_constructor thy t) then let
-      val case_rewrites = (#case_rewrites (Datatype.the_info thy
-        ((fst o dest_Type o fastype_of) t)))
-      in case_rewrites @ maps get_case_rewrite (snd (strip_comb t)) end
-    else []
-  val simprules = @{thm "unit.cases"} :: @{thm "prod.cases"} :: maps get_case_rewrite out_ts
-(* replace TRY by determining if it necessary - are there equations when calling compile match? *)
-in
-   (* make this simpset better! *)
-  asm_simp_tac (HOL_basic_ss' addsimps simprules) 1
-  THEN print_tac "after prove_match:"
-  THEN (DETERM (TRY (EqSubst.eqsubst_tac (ProofContext.init thy) [0] [@{thm "HOL.if_P"}] 1
-         THEN (REPEAT_DETERM (rtac @{thm conjI} 1 THEN (SOLVED (asm_simp_tac HOL_basic_ss 1))))
-         THEN (SOLVED (asm_simp_tac HOL_basic_ss 1)))))
-  THEN print_tac "after if simplification"
-end;
-
-(* corresponds to compile_fun -- maybe call that also compile_sidecond? *)
-
-fun prove_sidecond thy modes t =
-  let
-    fun preds_of t nameTs = case strip_comb t of 
-      (f as Const (name, T), args) =>
-        if AList.defined (op =) modes name then (name, T) :: nameTs
-          else fold preds_of args nameTs
-      | _ => nameTs
-    val preds = preds_of t []
-    val defs = map
-      (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
-        preds
-  in 
-    (* remove not_False_eq_True when simpset in prove_match is better *)
-    simp_tac (HOL_basic_ss addsimps @{thm not_False_eq_True} :: @{thm eval_pred} :: defs) 1 
-    (* need better control here! *)
-  end
-
-fun prove_clause thy nargs modes (iss, is) (_, clauses) (ts, moded_ps) =
-  let
-    val (in_ts, clause_out_ts) = split_smode is ts;
-    fun prove_prems out_ts [] =
-      (prove_match thy out_ts)
-      THEN asm_simp_tac HOL_basic_ss' 1
-      THEN (rtac (if null clause_out_ts then @{thm singleI_unit} else @{thm singleI}) 1)
-    | prove_prems out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
-      let
-        val premposition = (find_index (equal p) clauses) + nargs
-        val rest_tac = (case p of Prem (us, t) =>
-            let
-              val (_, out_ts''') = split_smode is us
-              val rec_tac = prove_prems out_ts''' ps
-            in
-              print_tac "before clause:"
-              THEN asm_simp_tac HOL_basic_ss 1
-              THEN print_tac "before prove_expr:"
-              THEN prove_expr thy (mode, t, us) premposition
-              THEN print_tac "after prove_expr:"
-              THEN rec_tac
-            end
-          | Negprem (us, t) =>
-            let
-              val (_, out_ts''') = split_smode is us
-              val rec_tac = prove_prems out_ts''' ps
-              val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
-              val (_, params) = strip_comb t
-            in
-              rtac @{thm bindI} 1
-              THEN (if (is_some name) then
-                  simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1
-                  THEN rtac @{thm not_predI} 1
-                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
-                  THEN (REPEAT_DETERM (atac 1))
-                  (* FIXME: work with parameter arguments *)
-                  THEN (EVERY (map (prove_param thy) (param_modes ~~ params)))
-                else
-                  rtac @{thm not_predI'} 1)
-                  THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
-              THEN rec_tac
-            end
-          | Sidecond t =>
-           rtac @{thm bindI} 1
-           THEN rtac @{thm if_predI} 1
-           THEN print_tac "before sidecond:"
-           THEN prove_sidecond thy modes t
-           THEN print_tac "after sidecond:"
-           THEN prove_prems [] ps)
-      in (prove_match thy out_ts)
-          THEN rest_tac
-      end;
-    val prems_tac = prove_prems in_ts moded_ps
-  in
-    rtac @{thm bindI} 1
-    THEN rtac @{thm singleI} 1
-    THEN prems_tac
-  end;
-
-fun select_sup 1 1 = []
-  | select_sup _ 1 = [rtac @{thm supI1}]
-  | select_sup n i = (rtac @{thm supI2})::(select_sup (n - 1) (i - 1));
-
-fun prove_one_direction thy clauses preds modes pred mode moded_clauses =
-  let
-    val T = the (AList.lookup (op =) preds pred)
-    val nargs = length (binder_types T) - nparams_of thy pred
-    val pred_case_rule = the_elim_of thy pred
-  in
-    REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"}))
-    THEN etac (predfun_elim_of thy pred mode) 1
-    THEN etac pred_case_rule 1
-    THEN (EVERY (map
-           (fn i => EVERY' (select_sup (length moded_clauses) i) i) 
-             (1 upto (length moded_clauses))))
-    THEN (EVERY (map2 (prove_clause thy nargs modes mode) clauses moded_clauses))
-    THEN print_tac "proved one direction"
-  end;
-
-(** Proof in the other direction **)
-
-fun prove_match2 thy out_ts = let
-  fun split_term_tac (Free _) = all_tac
-    | split_term_tac t =
-      if (is_constructor thy t) then let
-        val info = Datatype.the_info thy ((fst o dest_Type o fastype_of) t)
-        val num_of_constrs = length (#case_rewrites info)
-        (* special treatment of pairs -- because of fishing *)
-        val split_rules = case (fst o dest_Type o fastype_of) t of
-          "*" => [@{thm prod.split_asm}] 
-          | _ => PureThy.get_thms thy (((fst o dest_Type o fastype_of) t) ^ ".split_asm")
-        val (_, ts) = strip_comb t
-      in
-        (Splitter.split_asm_tac split_rules 1)
-(*        THEN (Simplifier.asm_full_simp_tac HOL_basic_ss 1)
-          THEN (DETERM (TRY (etac @{thm Pair_inject} 1))) *)
-        THEN (REPEAT_DETERM_N (num_of_constrs - 1) (etac @{thm botE} 1 ORELSE etac @{thm botE} 2))
-        THEN (EVERY (map split_term_tac ts))
-      end
-    else all_tac
-  in
-    split_term_tac (mk_tuple out_ts)
-    THEN (DETERM (TRY ((Splitter.split_asm_tac [@{thm "split_if_asm"}] 1) THEN (etac @{thm botE} 2))))
-  end
-
-(* VERY LARGE SIMILIRATIY to function prove_param 
--- join both functions
-*)
-(* TODO: remove function *)
-
-fun prove_param2 thy (NONE, t) = all_tac 
-  | prove_param2 thy (m as SOME (Mode (mode, is, ms)), t) = let
-    val  (f, args) = strip_comb (Envir.eta_contract t)
-    val (params, _) = chop (length ms) args
-    val f_tac = case f of
-        Const (name, T) => full_simp_tac (HOL_basic_ss addsimps 
-           (@{thm eval_pred}::(predfun_definition_of thy name mode)
-           :: @{thm "Product_Type.split_conv"}::[])) 1
-      | Free _ => all_tac
-      | _ => error "prove_param2: illegal parameter term"
-  in  
-    print_tac "before simplification in prove_args:"
-    THEN f_tac
-    THEN print_tac "after simplification in prove_args"
-    THEN (EVERY (map (prove_param2 thy) (ms ~~ params)))
-  end
-
-
-fun prove_expr2 thy (Mode (mode, is, ms), t) = 
-  (case strip_comb t of
-    (Const (name, T), args) =>
-      etac @{thm bindE} 1
-      THEN (REPEAT_DETERM (CHANGED (rewtac @{thm "split_paired_all"})))
-      THEN print_tac "prove_expr2-before"
-      THEN (debug_tac (Syntax.string_of_term_global thy
-        (prop_of (predfun_elim_of thy name mode))))
-      THEN (etac (predfun_elim_of thy name mode) 1)
-      THEN print_tac "prove_expr2"
-      THEN (EVERY (map (prove_param2 thy) (ms ~~ args)))
-      THEN print_tac "finished prove_expr2"      
-    | _ => etac @{thm bindE} 1)
-    
-(* FIXME: what is this for? *)
-(* replace defined by has_mode thy pred *)
-(* TODO: rewrite function *)
-fun prove_sidecond2 thy modes t = let
-  fun preds_of t nameTs = case strip_comb t of 
-    (f as Const (name, T), args) =>
-      if AList.defined (op =) modes name then (name, T) :: nameTs
-        else fold preds_of args nameTs
-    | _ => nameTs
-  val preds = preds_of t []
-  val defs = map
-    (fn (pred, T) => predfun_definition_of thy pred ([], (1 upto (length (binder_types T)))))
-      preds
-  in
-   (* only simplify the one assumption *)
-   full_simp_tac (HOL_basic_ss' addsimps @{thm eval_pred} :: defs) 1 
-   (* need better control here! *)
-   THEN print_tac "after sidecond2 simplification"
-   end
-  
-fun prove_clause2 thy modes pred (iss, is) (ts, ps) i =
-  let
-    val pred_intro_rule = nth (intros_of thy pred) (i - 1)
-    val (in_ts, clause_out_ts) = split_smode is ts;
-    fun prove_prems2 out_ts [] =
-      print_tac "before prove_match2 - last call:"
-      THEN prove_match2 thy out_ts
-      THEN print_tac "after prove_match2 - last call:"
-      THEN (etac @{thm singleE} 1)
-      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
-      THEN (asm_full_simp_tac HOL_basic_ss' 1)
-      THEN (REPEAT_DETERM (etac @{thm Pair_inject} 1))
-      THEN (asm_full_simp_tac HOL_basic_ss' 1)
-      THEN SOLVED (print_tac "state before applying intro rule:"
-      THEN (rtac pred_intro_rule 1)
-      (* How to handle equality correctly? *)
-      THEN (print_tac "state before assumption matching")
-      THEN (REPEAT (atac 1 ORELSE 
-         (CHANGED (asm_full_simp_tac HOL_basic_ss' 1)
-          THEN print_tac "state after simp_tac:"))))
-    | prove_prems2 out_ts ((p, mode as Mode ((iss, is), _, param_modes)) :: ps) =
-      let
-        val rest_tac = (case p of
-          Prem (us, t) =>
-          let
-            val (_, out_ts''') = split_smode is us
-            val rec_tac = prove_prems2 out_ts''' ps
-          in
-            (prove_expr2 thy (mode, t)) THEN rec_tac
-          end
-        | Negprem (us, t) =>
-          let
-            val (_, out_ts''') = split_smode is us
-            val rec_tac = prove_prems2 out_ts''' ps
-            val name = (case strip_comb t of (Const (c, _), _) => SOME c | _ => NONE)
-            val (_, params) = strip_comb t
-          in
-            print_tac "before neg prem 2"
-            THEN etac @{thm bindE} 1
-            THEN (if is_some name then
-                full_simp_tac (HOL_basic_ss addsimps [predfun_definition_of thy (the name) (iss, is)]) 1 
-                THEN etac @{thm not_predE} 1
-                THEN simp_tac (HOL_basic_ss addsimps [@{thm not_False_eq_True}]) 1
-                THEN (EVERY (map (prove_param2 thy) (param_modes ~~ params)))
-              else
-                etac @{thm not_predE'} 1)
-            THEN rec_tac
-          end 
-        | Sidecond t =>
-          etac @{thm bindE} 1
-          THEN etac @{thm if_predE} 1
-          THEN prove_sidecond2 thy modes t 
-          THEN prove_prems2 [] ps)
-      in print_tac "before prove_match2:"
-         THEN prove_match2 thy out_ts
-         THEN print_tac "after prove_match2:"
-         THEN rest_tac
-      end;
-    val prems_tac = prove_prems2 in_ts ps 
-  in
-    print_tac "starting prove_clause2"
-    THEN etac @{thm bindE} 1
-    THEN (etac @{thm singleE'} 1)
-    THEN (TRY (etac @{thm Pair_inject} 1))
-    THEN print_tac "after singleE':"
-    THEN prems_tac
-  end;
- 
-fun prove_other_direction thy modes pred mode moded_clauses =
-  let
-    fun prove_clause clause i =
-      (if i < length moded_clauses then etac @{thm supE} 1 else all_tac)
-      THEN (prove_clause2 thy modes pred mode clause i)
-  in
-    (DETERM (TRY (rtac @{thm unit.induct} 1)))
-     THEN (REPEAT_DETERM (CHANGED (rewtac @{thm split_paired_all})))
-     THEN (rtac (predfun_intro_of thy pred mode) 1)
-     THEN (REPEAT_DETERM (rtac @{thm refl} 2))
-     THEN (EVERY (map2 prove_clause moded_clauses (1 upto (length moded_clauses))))
-  end;
-
-(** proof procedure **)
-
-fun prove_pred thy clauses preds modes pred mode (moded_clauses, compiled_term) =
-  let
-    val ctxt = ProofContext.init thy
-    val clauses = the (AList.lookup (op =) clauses pred)
-  in
-    Goal.prove ctxt (Term.add_free_names compiled_term []) [] compiled_term
-      (if !do_proofs then
-        (fn _ =>
-        rtac @{thm pred_iffI} 1
-        THEN prove_one_direction thy clauses preds modes pred mode moded_clauses
-        THEN print_tac "proved one direction"
-        THEN prove_other_direction thy modes pred mode moded_clauses
-        THEN print_tac "proved other direction")
-       else (fn _ => mycheat_tac thy 1))
-  end;
-
-(* composition of mode inference, definition, compilation and proof *)
-
-(** auxillary combinators for table of preds and modes **)
-
-fun map_preds_modes f preds_modes_table =
-  map (fn (pred, modes) =>
-    (pred, map (fn (mode, value) => (mode, f pred mode value)) modes)) preds_modes_table
-
-fun join_preds_modes table1 table2 =
-  map_preds_modes (fn pred => fn mode => fn value =>
-    (value, the (AList.lookup (op =) (the (AList.lookup (op =) table2 pred)) mode))) table1
-    
-fun maps_modes preds_modes_table =
-  map (fn (pred, modes) =>
-    (pred, map (fn (mode, value) => value) modes)) preds_modes_table  
-    
-fun compile_preds compfuns mk_fun_of use_size thy all_vs param_vs preds moded_clauses =
-  map_preds_modes (fn pred => compile_pred compfuns mk_fun_of use_size thy all_vs param_vs pred
-      (the (AList.lookup (op =) preds pred))) moded_clauses  
-  
-fun prove thy clauses preds modes moded_clauses compiled_terms =
-  map_preds_modes (prove_pred thy clauses preds modes)
-    (join_preds_modes moded_clauses compiled_terms)
-
-fun prove_by_skip thy _ _ _ _ compiled_terms =
-  map_preds_modes (fn pred => fn mode => fn t => Drule.standard (Skip_Proof.make_thm thy t))
-    compiled_terms
-    
-fun prepare_intrs thy prednames =
-  let
-    val intrs = maps (intros_of thy) prednames
-      |> map (Logic.unvarify o prop_of)
-    val nparams = nparams_of thy (hd prednames)
-    val extra_modes = all_modes_of thy |> filter_out (fn (name, _) => member (op =) prednames name)
-    val preds = distinct (op =) (map (dest_Const o fst o (strip_intro_concl nparams)) intrs)
-    val _ $ u = Logic.strip_imp_concl (hd intrs);
-    val params = List.take (snd (strip_comb u), nparams);
-    val param_vs = maps term_vs params
-    val all_vs = terms_vs intrs
-    fun dest_prem t =
-      (case strip_comb t of
-        (v as Free _, ts) => if v mem params then Prem (ts, v) else Sidecond t
-      | (c as Const (@{const_name Not}, _), [t]) => (case dest_prem t of          
-          Prem (ts, t) => Negprem (ts, t)
-        | Negprem _ => error ("Double negation not allowed in premise: " ^ (Syntax.string_of_term_global thy (c $ t))) 
-        | Sidecond t => Sidecond (c $ t))
-      | (c as Const (s, _), ts) =>
-        if is_registered thy s then
-          let val (ts1, ts2) = chop (nparams_of thy s) ts
-          in Prem (ts2, list_comb (c, ts1)) end
-        else Sidecond t
-      | _ => Sidecond t)
-    fun add_clause intr (clauses, arities) =
-    let
-      val _ $ t = Logic.strip_imp_concl intr;
-      val (Const (name, T), ts) = strip_comb t;
-      val (ts1, ts2) = chop nparams ts;
-      val prems = map (dest_prem o HOLogic.dest_Trueprop) (Logic.strip_imp_prems intr);
-      val (Ts, Us) = chop nparams (binder_types T)
-    in
-      (AList.update op = (name, these (AList.lookup op = clauses name) @
-        [(ts2, prems)]) clauses,
-       AList.update op = (name, (map (fn U => (case strip_type U of
-                 (Rs as _ :: _, Type ("bool", [])) => SOME (length Rs)
-               | _ => NONE)) Ts,
-             length Us)) arities)
-    end;
-    val (clauses, arities) = fold add_clause intrs ([], []);
-  in (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) end;
-
-(** main function of predicate compiler **)
-
-fun add_equations_of steps prednames thy =
-  let
-    val _ = tracing ("Starting predicate compiler for predicates " ^ commas prednames ^ "...")
-    val (preds, nparams, all_vs, param_vs, extra_modes, clauses, arities) =
-      prepare_intrs thy prednames
-    val _ = tracing "Infering modes..."
-    val moded_clauses = #infer_modes steps thy extra_modes arities param_vs clauses 
-    val modes = map (fn (p, mps) => (p, map fst mps)) moded_clauses
-    val _ = print_modes modes
-    val _ = print_moded_clauses thy moded_clauses
-    val _ = tracing "Defining executable functions..."
-    val thy' = fold (#create_definitions steps preds) modes thy
-      |> Theory.checkpoint
-    val _ = tracing "Compiling equations..."
-    val compiled_terms =
-      (#compile_preds steps) thy' all_vs param_vs preds moded_clauses
-    val _ = print_compiled_terms thy' compiled_terms
-    val _ = tracing "Proving equations..."
-    val result_thms = #prove steps thy' clauses preds (extra_modes @ modes)
-      moded_clauses compiled_terms
-    val qname = #qname steps
-    (* val attrib = gn thy => Attrib.attribute_i thy Code.add_eqn_attrib *)
-    val attrib = fn thy => Attrib.attribute_i thy (Attrib.internal (K (Thm.declaration_attribute
-      (fn thm => Context.mapping (Code.add_eqn thm) I))))
-    val thy'' = fold (fn (name, result_thms) => fn thy => snd (PureThy.add_thmss
-      [((Binding.qualify true (Long_Name.base_name name) (Binding.name qname), result_thms),
-        [attrib thy ])] thy))
-      (maps_modes result_thms) thy'
-      |> Theory.checkpoint
-  in
-    thy''
-  end
-
-fun extend' value_of edges_of key (G, visited) =
-  let
-    val (G', v) = case try (Graph.get_node G) key of
-        SOME v => (G, v)
-      | NONE => (Graph.new_node (key, value_of key) G, value_of key)
-    val (G'', visited') = fold (extend' value_of edges_of) (edges_of (key, v) \\ visited)
-      (G', key :: visited) 
-  in
-    (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited')
-  end;
-
-fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) 
-  
-fun gen_add_equations steps names thy =
-  let
-    val thy' = PredData.map (fold (extend (fetch_pred_data thy) (depending_preds_of thy)) names) thy
-      |> Theory.checkpoint;
-    fun strong_conn_of gr keys =
-      Graph.strong_conn (Graph.subgraph (member (op =) (Graph.all_succs gr keys)) gr)
-    val scc = strong_conn_of (PredData.get thy') names
-    val thy'' = fold_rev
-      (fn preds => fn thy =>
-        if #are_not_defined steps thy preds then add_equations_of steps preds thy else thy)
-      scc thy' |> Theory.checkpoint
-  in thy'' end
-
-(* different instantiantions of the predicate compiler *)
-
-val add_equations = gen_add_equations
-  {infer_modes = infer_modes false,
-  create_definitions = create_definitions,
-  compile_preds = compile_preds PredicateCompFuns.compfuns mk_fun_of false,
-  prove = prove,
-  are_not_defined = (fn thy => forall (null o modes_of thy)),
-  qname = "equation"}
-
-val add_sizelim_equations = gen_add_equations
-  {infer_modes = infer_modes false,
-  create_definitions = sizelim_create_definitions,
-  compile_preds = compile_preds PredicateCompFuns.compfuns mk_sizelim_fun_of true,
-  prove = prove_by_skip,
-  are_not_defined = (fn thy => fn preds => true), (* TODO *)
-  qname = "sizelim_equation"
-  }
-  
-val add_quickcheck_equations = gen_add_equations
-  {infer_modes = infer_modes_with_generator,
-  create_definitions = rpred_create_definitions,
-  compile_preds = compile_preds RPredCompFuns.compfuns mk_generator_of true,
-  prove = prove_by_skip,
-  are_not_defined = (fn thy => fn preds => true), (* TODO *)
-  qname = "rpred_equation"}
-
-(** user interface **)
-
-(* generation of case rules from user-given introduction rules *)
-
-fun mk_casesrule ctxt nparams introrules =
-  let
-    val intros = map (Logic.unvarify o prop_of) introrules
-    val (pred, (params, args)) = strip_intro_concl nparams (hd intros)
-    val ([propname], ctxt1) = Variable.variant_fixes ["thesis"] ctxt
-    val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT))
-    val (argnames, ctxt2) = Variable.variant_fixes
-      (map (fn i => "a" ^ string_of_int i) (1 upto (length args))) ctxt1
-    val argvs = map2 (curry Free) argnames (map fastype_of args)
-    fun mk_case intro =
-      let
-        val (_, (_, args)) = strip_intro_concl nparams intro
-        val prems = Logic.strip_imp_prems intro
-        val eqprems = map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (argvs ~~ args)
-        val frees = (fold o fold_aterms)
-          (fn t as Free _ =>
-              if member (op aconv) params t then I else insert (op aconv) t
-           | _ => I) (args @ prems) []
-      in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end
-    val assm = HOLogic.mk_Trueprop (list_comb (pred, params @ argvs))
-    val cases = map mk_case intros
-  in Logic.list_implies (assm :: cases, prop) end;
-
-(* code_pred_intro attribute *)
-
-fun attrib f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I);
-
-val code_pred_intros_attrib = attrib add_intro;
-
-local
-
-(* TODO: make TheoryDataFun to GenericDataFun & remove duplication of local theory and theory *)
-(* TODO: must create state to prove multiple cases *)
-fun generic_code_pred prep_const raw_const lthy =
-  let
-    val thy = ProofContext.theory_of lthy
-    val const = prep_const thy raw_const
-    val lthy' = LocalTheory.theory (PredData.map
-        (extend (fetch_pred_data thy) (depending_preds_of thy) const)) lthy
-      |> LocalTheory.checkpoint
-    val thy' = ProofContext.theory_of lthy'
-    val preds = Graph.all_preds (PredData.get thy') [const] |> filter_out (has_elim thy')
-    fun mk_cases const =
-      let
-        val nparams = nparams_of thy' const
-        val intros = intros_of thy' const
-      in mk_casesrule lthy' nparams intros end  
-    val cases_rules = map mk_cases preds
-    val cases =
-      map (fn case_rule => RuleCases.Case {fixes = [],
-        assumes = [("", Logic.strip_imp_prems case_rule)],
-        binds = [], cases = []}) cases_rules
-    val case_env = map2 (fn p => fn c => (Long_Name.base_name p, SOME c)) preds cases
-    val lthy'' = lthy'
-      |> fold Variable.auto_fixes cases_rules 
-      |> ProofContext.add_cases true case_env
-    fun after_qed thms goal_ctxt =
-      let
-        val global_thms = ProofContext.export goal_ctxt
-          (ProofContext.init (ProofContext.theory_of goal_ctxt)) (map the_single thms)
-      in
-        goal_ctxt |> LocalTheory.theory (fold set_elim global_thms #> add_equations [const])
-      end  
-  in
-    Proof.theorem_i NONE after_qed (map (single o (rpair [])) cases_rules) lthy''
-  end;
-
-structure P = OuterParse
-
-in
-
-val code_pred = generic_code_pred (K I);
-val code_pred_cmd = generic_code_pred Code.read_const
-
-val setup = PredData.put (Graph.empty) #>
-  Attrib.setup @{binding code_pred_intros} (Scan.succeed (attrib add_intro))
-    "adding alternative introduction rules for code generation of inductive predicates"
-(*  Attrib.setup @{binding code_ind_cases} (Scan.succeed add_elim_attrib)
-    "adding alternative elimination rules for code generation of inductive predicates";
-    *)
-  (*FIXME name discrepancy in attribs and ML code*)
-  (*FIXME intros should be better named intro*)
-  (*FIXME why distinguished attribute for cases?*)
-
-val _ = OuterSyntax.local_theory_to_proof "code_pred"
-  "prove equations for predicate specified by intro/elim rules"
-  OuterKeyword.thy_goal (P.term_group >> code_pred_cmd)
-
-end
-
-(*FIXME
-- Naming of auxiliary rules necessary?
-- add default code equations P x y z = P_i_i_i x y z
-*)
-
-(* transformation for code generation *)
-
-val eval_ref = Unsynchronized.ref (NONE : (unit -> term Predicate.pred) option);
-
-(*FIXME turn this into an LCF-guarded preprocessor for comprehensions*)
-fun analyze_compr thy t_compr =
-  let
-    val split = case t_compr of (Const (@{const_name Collect}, _) $ t) => t
-      | _ => error ("Not a set comprehension: " ^ Syntax.string_of_term_global thy t_compr);
-    val (body, Ts, fp) = HOLogic.strip_psplits split;
-    val (pred as Const (name, T), all_args) = strip_comb body;
-    val (params, args) = chop (nparams_of thy name) all_args;
-    val user_mode = map_filter I (map_index
-      (fn (i, t) => case t of Bound j => if j < length Ts then NONE
-        else SOME (i+1) | _ => SOME (i+1)) args); (*FIXME dangling bounds should not occur*)
-    val modes = filter (fn Mode (_, is, _) => is = user_mode)
-      (modes_of_term (all_modes_of thy) (list_comb (pred, params)));
-    val m = case modes
-     of [] => error ("No mode possible for comprehension "
-                ^ Syntax.string_of_term_global thy t_compr)
-      | [m] => m
-      | m :: _ :: _ => (warning ("Multiple modes possible for comprehension "
-                ^ Syntax.string_of_term_global thy t_compr); m);
-    val (inargs, outargs) = split_smode user_mode args;
-    val t_pred = list_comb (compile_expr NONE thy (m, list_comb (pred, params)), inargs);
-    val t_eval = if null outargs then t_pred else let
-        val outargs_bounds = map (fn Bound i => i) outargs;
-        val outargsTs = map (nth Ts) outargs_bounds;
-        val T_pred = HOLogic.mk_tupleT outargsTs;
-        val T_compr = HOLogic.mk_ptupleT fp Ts;
-        val arrange_bounds = map_index I outargs_bounds
-          |> sort (prod_ord (K EQUAL) int_ord)
-          |> map fst;
-        val arrange = funpow (length outargs_bounds - 1) HOLogic.mk_split
-          (Term.list_abs (map (pair "") outargsTs,
-            HOLogic.mk_ptuple fp T_compr (map Bound arrange_bounds)))
-      in mk_map PredicateCompFuns.compfuns T_pred T_compr arrange t_pred end
-  in t_eval end;
-
-fun eval thy t_compr =
-  let
-    val t = analyze_compr thy t_compr;
-    val T = dest_predT PredicateCompFuns.compfuns (fastype_of t);
-    val t' = mk_map PredicateCompFuns.compfuns T HOLogic.termT (HOLogic.term_of_const T) t;
-  in (T, Code_ML.eval NONE ("Predicate_Compile.eval_ref", eval_ref) Predicate.map thy t' []) end;
-
-fun values ctxt k t_compr =
-  let
-    val thy = ProofContext.theory_of ctxt;
-    val (T, t) = eval thy t_compr;
-    val setT = HOLogic.mk_setT T;
-    val (ts, _) = Predicate.yieldn k t;
-    val elemsT = HOLogic.mk_set T ts;
-  in if k = ~1 orelse length ts < k then elemsT
-    else Const (@{const_name Lattices.sup}, setT --> setT --> setT) $ elemsT $ t_compr
-  end;
-
-fun values_cmd modes k raw_t state =
-  let
-    val ctxt = Toplevel.context_of state;
-    val t = Syntax.read_term ctxt raw_t;
-    val t' = values ctxt k t;
-    val ty' = Term.type_of t';
-    val ctxt' = Variable.auto_fixes t' ctxt;
-    val p = PrintMode.with_modes modes (fn () =>
-      Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk,
-        Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) ();
-  in Pretty.writeln p end;
-
-local structure P = OuterParse in
-
-val opt_modes = Scan.optional (P.$$$ "(" |-- P.!!! (Scan.repeat1 P.xname --| P.$$$ ")")) [];
-
-val _ = OuterSyntax.improper_command "values" "enumerate and print comprehensions" OuterKeyword.diag
-  (opt_modes -- Scan.optional P.nat ~1 -- P.term
-    >> (fn ((modes, k), t) => Toplevel.no_timing o Toplevel.keep
-        (values_cmd modes k t)));
-
-end;
-
-end;
-