moved Quot package to Attic (still compiles there with "isabelle make")
authorChristian Urban <urbanc@in.tum.de>
Thu, 25 Feb 2010 07:57:17 +0100
changeset 1260 9df6144e281b
parent 1259 db158e995bfc
child 1261 853abc14c5c6
moved Quot package to Attic (still compiles there with "isabelle make")
Attic/FIXME-TODO
Attic/IsaMakefile
Attic/Quot/Examples/AbsRepTest.thy
Attic/Quot/Examples/FSet.thy
Attic/Quot/Examples/FSet2.thy
Attic/Quot/Examples/FSet3.thy
Attic/Quot/Examples/IntEx.thy
Attic/Quot/Examples/IntEx2.thy
Attic/Quot/Examples/LFex.thy
Attic/Quot/Examples/LamEx.thy
Attic/Quot/Examples/LarryDatatype.thy
Attic/Quot/Examples/LarryInt.thy
Attic/Quot/Examples/SigmaEx.thy
Attic/Quot/Examples/Terms.thy
Attic/Quot/Quotient.thy
Attic/Quot/Quotient_List.thy
Attic/Quot/Quotient_Option.thy
Attic/Quot/Quotient_Product.thy
Attic/Quot/Quotient_Sum.thy
Attic/Quot/Quotient_Syntax.thy
Attic/Quot/ROOT.ML
Attic/Quot/quotient_def.ML
Attic/Quot/quotient_info.ML
Attic/Quot/quotient_tacs.ML
Attic/Quot/quotient_term.ML
Attic/Quot/quotient_typ.ML
FIXME-TODO
IsaMakefile
Quot/Examples/AbsRepTest.thy
Quot/Examples/FSet.thy
Quot/Examples/FSet2.thy
Quot/Examples/FSet3.thy
Quot/Examples/IntEx.thy
Quot/Examples/IntEx2.thy
Quot/Examples/LFex.thy
Quot/Examples/LamEx.thy
Quot/Examples/LarryDatatype.thy
Quot/Examples/LarryInt.thy
Quot/Examples/SigmaEx.thy
Quot/Examples/Terms.thy
Quot/Quotient.thy
Quot/Quotient_List.thy
Quot/Quotient_Option.thy
Quot/Quotient_Product.thy
Quot/Quotient_Sum.thy
Quot/Quotient_Syntax.thy
Quot/ROOT.ML
Quot/quotient_def.ML
Quot/quotient_info.ML
Quot/quotient_tacs.ML
Quot/quotient_term.ML
Quot/quotient_typ.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/FIXME-TODO	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,69 @@
+Highest Priority
+================
+
+- give examples for the new quantifier translations in regularization
+  (quotient_term.ML)
+
+
+Higher Priority
+===============
+
+
+- Also, in the interest of making nicer generated documentation, you
+  might want to change all your "section" headings in Quotient.thy to
+  "subsection", and add a "header" statement to the top of the file.
+  Otherwise, each "section" gets its own chapter in the generated pdf,
+  when the rest of HOL has one chapter per theory file (the chapter
+  title comes from the "header" statement).
+
+- If the constant definition gives the wrong definition
+  term, one gets a cryptic message about absrep_fun
+
+- Handle theorems that include Ball/Bex. For this, would 
+  it help if we introduced separate Bex and Ball constants 
+  for quotienting?
+
+- The user should be able to give quotient_respects and 
+  preserves theorems in a more natural form.
+
+Lower Priority
+==============
+
+- accept partial equivalence relations
+
+- think about what happens if things go wrong (like
+  theorem cannot be lifted) / proper diagnostic 
+  messages for the user
+
+- inductions from the datatype package have a strange
+  order of quantifiers in assumptions.
+
+- find clean ways how to write down the "mathematical"
+  procedure for a possible submission (Peter submitted 
+  his work only to TPHOLs 2005...we would have to go
+  maybe for the Journal of Formalised Mathematics)
+
+- add tests for adding theorems to the various thm lists
+
+- Maybe quotient and equiv theorems like the ones for
+  [QuotList, QuotOption, QuotPair...] could be automatically
+  proven?
+
+- Examples: Finite multiset.
+
+- The current syntax of the quotient_definition is
+
+      "qconst :: qty"
+      as "rconst"
+
+  Is it possible to have the more Isabelle-like
+  syntax
+   
+      qconst :: "qty"
+      as "rconst"
+
+  That means "qconst :: qty" is not read as a term, but
+  as two entities.
+
+- Restrict automatic translation to particular quotient types
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/IsaMakefile	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,46 @@
+
+## targets
+
+default: Quot
+images: 
+
+all: Quot
+
+
+## global settings
+
+SRC = $(ISABELLE_HOME)/src
+OUT = $(ISABELLE_OUTPUT)
+LOG = $(OUT)/log
+
+USEDIR = $(ISABELLE_TOOL) usedir -v true -t true ##-D generated
+
+
+## Quot
+
+Quot: $(LOG)/HOL-Quot.gz
+
+$(LOG)/HOL-Quot.gz: Quot/ROOT.ML Quot/*.thy
+	@$(USEDIR) HOL-Nominal Quot
+
+paper: $(LOG)/HOL-Quot-Paper.gz
+
+$(LOG)/HOL-Quot-Paper.gz: Paper/ROOT.ML Paper/document/root.tex Paper/*.thy
+	@$(USEDIR) -D generated HOL Paper
+	$(ISATOOL) document -o pdf  Paper/generated
+	@cp Paper/document.pdf paper.pdf
+
+keywords:
+	mkdir -p tmp
+	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/Pure.gz tmp 
+	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/HOL.gz tmp
+	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/Pure-ProofGeneral.gz tmp
+	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/HOL-Nominal.gz tmp
+	cp $(LOG)/HOL-Nominal-Quot.gz tmp
+	isabelle keywords -k quot tmp/*
+	
+	
+## clean
+
+clean:
+	@rm -f $(LOG)/HOL-Quot.gz
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/AbsRepTest.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,240 @@
+theory AbsRepTest
+imports "../Quotient" "../Quotient_List" "../Quotient_Option" "../Quotient_Sum" "../Quotient_Product" List
+begin
+
+
+(*
+ML_command "ProofContext.debug := false"
+ML_command "ProofContext.verbose := false"
+*)
+
+ML {* open Quotient_Term *}
+
+ML {*
+fun test_funs flag ctxt (rty, qty) =
+  (absrep_fun_chk flag ctxt (rty, qty)
+   |> Syntax.string_of_term ctxt
+   |> writeln;
+   equiv_relation_chk ctxt (rty, qty) 
+   |> Syntax.string_of_term ctxt
+   |> writeln)
+*}
+
+definition
+  erel1 (infixl "\<approx>1" 50)
+where
+  "erel1 \<equiv> \<lambda>xs ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
+
+quotient_type 
+  'a fset = "'a list" / erel1
+  apply(rule equivpI)
+  unfolding erel1_def reflp_def symp_def transp_def
+  by auto
+
+definition
+  erel2 (infixl "\<approx>2" 50)
+where
+  "erel2 \<equiv> \<lambda>(xs::('a * 'a) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
+
+quotient_type 
+  'a foo = "('a * 'a) list" / erel2
+  apply(rule equivpI)
+  unfolding erel2_def reflp_def symp_def transp_def
+  by auto
+
+definition
+  erel3 (infixl "\<approx>3" 50)
+where
+  "erel3 \<equiv> \<lambda>(xs::('a * int) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
+
+quotient_type 
+  'a bar = "('a * int) list" / "erel3"
+  apply(rule equivpI)
+  unfolding erel3_def reflp_def symp_def transp_def
+  by auto
+
+fun
+  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infixl "\<approx>4" 50)
+where
+  "intrel (x, y) (u, v) = (x + v = u + y)"
+
+quotient_type myint = "nat \<times> nat" / intrel
+  by (auto simp add: equivp_def expand_fun_eq)
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "nat \<times> nat"}, 
+      @{typ "myint"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "('a * 'a) list"}, 
+      @{typ "'a foo"})
+*}
+
+ML {*
+test_funs RepF @{context} 
+     (@{typ "(('a * 'a) list * 'b)"}, 
+      @{typ "('a foo * 'b)"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "(('a list) * int) list"}, 
+      @{typ "('a fset) bar"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "('a list)"}, 
+      @{typ "('a fset)"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "('a list) list"}, 
+      @{typ "('a fset) fset"})
+*}
+
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "((nat * nat) list) list"}, 
+      @{typ "((myint) fset) fset"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "(('a * 'a) list) list"}, 
+      @{typ "(('a * 'a) fset) fset"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+      (@{typ "(nat * nat) list"}, 
+       @{typ "myint fset"})
+*}
+
+ML {*
+test_funs AbsF @{context} 
+     (@{typ "('a list) list \<Rightarrow> 'a list"}, 
+      @{typ "('a fset) fset \<Rightarrow> 'a fset"})
+*}
+
+lemma OO_sym_inv:
+  assumes sr: "symp r"
+  and     ss: "symp s"
+  shows "(r OO s) x y = (s OO r) y x"
+  using sr ss
+  unfolding symp_def
+  apply (metis pred_comp.intros pred_compE ss symp_def)
+  done
+
+lemma abs_o_rep:
+  assumes a: "Quotient r absf repf"
+  shows "absf o repf = id"
+  apply(rule ext)
+  apply(simp add: Quotient_abs_rep[OF a])
+  done
+
+lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
+  apply (rule eq_reflection)
+  apply auto
+  done
+
+lemma map_rel_cong: "b \<approx>1 ba \<Longrightarrow> map f b \<approx>1 map f ba"
+  unfolding erel1_def
+  apply(simp only: set_map set_in_eq)
+  done
+
+lemma quotient_compose_list_gen_pre:
+  assumes a: "equivp r2"
+  and b: "Quotient r2 abs2 rep2"
+  shows  "(list_rel r2 OOO op \<approx>1) r s =
+          ((list_rel r2 OOO op \<approx>1) r r \<and> (list_rel r2 OOO op \<approx>1) s s \<and>
+           abs_fset (map abs2 r) = abs_fset (map abs2 s))"
+  apply rule
+  apply rule
+  apply rule
+  apply (rule list_rel_refl)
+  apply (metis equivp_def a)
+  apply rule
+  apply (rule equivp_reflp[OF fset_equivp])
+  apply (rule list_rel_refl)
+  apply (metis equivp_def a)
+  apply(rule)
+  apply rule
+  apply (rule list_rel_refl)
+  apply (metis equivp_def a)
+  apply rule
+  apply (rule equivp_reflp[OF fset_equivp])
+  apply (rule list_rel_refl)
+  apply (metis equivp_def a)
+  apply (subgoal_tac "map abs2 r \<approx>1 map abs2 s")
+  apply (metis Quotient_rel[OF Quotient_fset])
+  apply (auto)[1]
+  apply (subgoal_tac "map abs2 r = map abs2 b")
+  prefer 2
+  apply (metis Quotient_rel[OF list_quotient[OF b]])
+  apply (subgoal_tac "map abs2 s = map abs2 ba")
+  prefer 2
+  apply (metis Quotient_rel[OF list_quotient[OF b]])
+  apply (simp add: map_rel_cong)
+  apply rule
+  apply (rule rep_abs_rsp[of "list_rel r2" "map abs2"])
+  apply (rule list_quotient)
+  apply (rule b)
+  apply (rule list_rel_refl)
+  apply (metis equivp_def a)
+  apply rule
+  prefer 2
+  apply (rule rep_abs_rsp_left[of "list_rel r2" "map abs2"])
+  apply (rule list_quotient)
+  apply (rule b)
+  apply (rule list_rel_refl)
+  apply (metis equivp_def a)
+  apply (erule conjE)+
+  apply (subgoal_tac "map abs2 r \<approx>1 map abs2 s")
+  apply (rule map_rel_cong)
+  apply (assumption)
+  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp a b)
+  done
+
+lemma quotient_compose_list_gen:
+  assumes a: "Quotient r2 abs2 rep2"
+  and     b: "equivp r2" (* reflp is not enough *)
+  shows  "Quotient ((list_rel r2) OOO (op \<approx>1))
+               (abs_fset \<circ> (map abs2)) ((map rep2) \<circ> rep_fset)"
+  unfolding Quotient_def comp_def
+  apply (rule)+
+  apply (simp add: abs_o_rep[OF a] id_simps Quotient_abs_rep[OF Quotient_fset])
+  apply (rule)
+  apply (rule)
+  apply (rule)
+  apply (rule list_rel_refl)
+  apply (metis b equivp_def)
+  apply (rule)
+  apply (rule equivp_reflp[OF fset_equivp])
+  apply (rule list_rel_refl)
+  apply (metis b equivp_def)
+  apply rule
+  apply rule
+  apply(rule quotient_compose_list_gen_pre[OF b a])
+  done
+
+(* This is the general statement but the types of abs2 and rep2
+   are wrong as can be seen in following exanples *)
+lemma quotient_compose_general:
+  assumes a2: "Quotient r1 abs1 rep1"
+  and         "Quotient r2 abs2 rep2"
+  shows  "Quotient ((list_rel r2) OOO r1)
+               (abs1 \<circ> (map abs2)) ((map rep2) \<circ> rep1)"
+sorry
+
+thm quotient_compose_list_gen[OF Quotient_fset fset_equivp]
+thm quotient_compose_general[OF Quotient_fset]
+(* Doesn't work: *)
+(* thm quotient_compose_general[OF Quotient_fset Quotient_fset] *)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/FSet.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,433 @@
+theory FSet
+imports "../Quotient" "../Quotient_List" "../Quotient_Product" List
+begin
+
+inductive
+  list_eq (infix "\<approx>" 50)
+where
+  "a#b#xs \<approx> b#a#xs"
+| "[] \<approx> []"
+| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
+| "a#a#xs \<approx> a#xs"
+| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
+| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
+
+lemma list_eq_refl:
+  shows "xs \<approx> xs"
+  by (induct xs) (auto intro: list_eq.intros)
+
+lemma equivp_list_eq:
+  shows "equivp list_eq"
+  unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
+  apply(auto intro: list_eq.intros list_eq_refl)
+  done
+
+quotient_type 
+  'a fset = "'a list" / "list_eq"
+  by (rule equivp_list_eq)
+
+quotient_definition
+   "EMPTY :: 'a fset"
+is
+  "[]::'a list"
+
+quotient_definition
+   "INSERT :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+is
+  "op #"
+
+quotient_definition
+   "FUNION :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+is
+  "op @"
+
+fun
+  card1 :: "'a list \<Rightarrow> nat"
+where
+  card1_nil: "(card1 []) = 0"
+| card1_cons: "(card1 (x # xs)) = (if (x mem xs) then (card1 xs) else (Suc (card1 xs)))"
+
+quotient_definition
+   "CARD :: 'a fset \<Rightarrow> nat"
+is
+  "card1"
+
+quotient_definition
+  "fconcat :: ('a fset) fset \<Rightarrow> 'a fset"
+is
+  "concat"
+
+term concat
+term fconcat
+
+text {*
+ Maybe make_const_def should require a theorem that says that the particular lifted function
+ respects the relation. With it such a definition would be impossible:
+ make_const_def CARD @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
+*}
+
+lemma card1_0:
+  fixes a :: "'a list"
+  shows "(card1 a = 0) = (a = [])"
+  by (induct a) auto
+
+lemma not_mem_card1:
+  fixes x :: "'a"
+  fixes xs :: "'a list"
+  shows "(~(x mem xs)) = (card1 (x # xs) = Suc (card1 xs))"
+  by auto
+
+lemma mem_cons:
+  fixes x :: "'a"
+  fixes xs :: "'a list"
+  assumes a : "x mem xs"
+  shows "x # xs \<approx> xs"
+  using a by (induct xs) (auto intro: list_eq.intros )
+
+lemma card1_suc:
+  fixes xs :: "'a list"
+  fixes n :: "nat"
+  assumes c: "card1 xs = Suc n"
+  shows "\<exists>a ys. ~(a mem ys) \<and> xs \<approx> (a # ys)"
+  using c
+apply(induct xs)
+apply (metis Suc_neq_Zero card1_0)
+apply (metis FSet.card1_cons list_eq.intros(6) list_eq_refl mem_cons)
+done
+
+definition
+  rsp_fold
+where
+  "rsp_fold f = ((!u v. (f u v = f v u)) \<and> (!u v w. ((f u (f v w) = f (f u v) w))))"
+
+primrec
+  fold1
+where
+  "fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
+| "fold1 f g z (a # A) =
+     (if rsp_fold f
+     then (
+       if (a mem A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
+     ) else z)"
+
+lemma fs1_strong_cases:
+  fixes X :: "'a list"
+  shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a mem Y) \<and> (X \<approx> a # Y)))"
+  apply (induct X)
+  apply (simp)
+  apply (metis List.member.simps(1) list_eq.intros(6) list_eq_refl mem_cons)
+  done
+
+quotient_definition
+   "IN :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
+is
+  "op mem"
+
+quotient_definition
+   "FOLD :: ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
+is
+  "fold1"
+
+quotient_definition
+  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
+is
+  "map"
+
+lemma mem_rsp:
+  fixes z
+  assumes a: "x \<approx> y"
+  shows "(z mem x) = (z mem y)"
+  using a by induct auto
+
+lemma ho_memb_rsp[quot_respect]:
+  "(op = ===> (op \<approx> ===> op =)) (op mem) (op mem)"
+  by (simp add: mem_rsp)
+
+lemma card1_rsp:
+  fixes a b :: "'a list"
+  assumes e: "a \<approx> b"
+  shows "card1 a = card1 b"
+  using e by induct (simp_all add: mem_rsp)
+
+lemma ho_card1_rsp[quot_respect]: 
+  "(op \<approx> ===> op =) card1 card1"
+  by (simp add: card1_rsp)
+
+lemma cons_rsp:
+  fixes z
+  assumes a: "xs \<approx> ys"
+  shows "(z # xs) \<approx> (z # ys)"
+  using a by (rule list_eq.intros(5))
+
+lemma ho_cons_rsp[quot_respect]:
+  "(op = ===> op \<approx> ===> op \<approx>) op # op #"
+  by (simp add: cons_rsp)
+
+lemma append_rsp_aux1:
+  assumes a : "l2 \<approx> r2 "
+  shows "(h @ l2) \<approx> (h @ r2)"
+using a
+apply(induct h)
+apply(auto intro: list_eq.intros(5))
+done
+
+lemma append_rsp_aux2:
+  assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "
+  shows "(l1 @ l2) \<approx> (r1 @ r2)"
+using a
+apply(induct arbitrary: l2 r2)
+apply(simp_all)
+apply(blast intro: list_eq.intros append_rsp_aux1)+
+done
+
+lemma append_rsp[quot_respect]:
+  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
+  by (auto simp add: append_rsp_aux2)
+
+lemma map_rsp:
+  assumes a: "a \<approx> b"
+  shows "map f a \<approx> map f b"
+  using a
+  apply (induct)
+  apply(auto intro: list_eq.intros)
+  done
+
+lemma ho_map_rsp[quot_respect]:
+  "(op = ===> op \<approx> ===> op \<approx>) map map"
+  by (simp add: map_rsp)
+
+lemma map_append:
+  "(map f (a @ b)) \<approx> (map f a) @ (map f b)"
+ by simp (rule list_eq_refl)
+
+lemma ho_fold_rsp[quot_respect]:
+  "(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
+  apply (auto)
+  apply (case_tac "rsp_fold x")
+  prefer 2
+  apply (erule_tac list_eq.induct)
+  apply (simp_all)
+  apply (erule_tac list_eq.induct)
+  apply (simp_all)
+  apply (auto simp add: mem_rsp rsp_fold_def)
+done
+
+lemma list_equiv_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>"
+by (auto intro: list_eq.intros)
+
+lemma "IN x EMPTY = False"
+apply(lifting member.simps(1))
+done
+
+lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
+apply (lifting member.simps(2))
+done
+
+lemma "INSERT a (INSERT a x) = INSERT a x"
+apply (lifting list_eq.intros(4))
+done
+
+lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"
+apply (lifting list_eq.intros(5))
+done
+
+lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
+apply (lifting card1_suc)
+done
+
+lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
+apply (lifting not_mem_card1)
+done
+
+lemma "FOLD f g (z::'b) (INSERT a x) =
+  (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
+apply(lifting fold1.simps(2))
+done
+
+lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
+apply (lifting map_append)
+done
+
+lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"
+apply (lifting append_assoc)
+done
+
+
+lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
+apply(lifting list.induct)
+done
+
+lemma list_induct_part:
+  assumes a: "P (x :: 'a list) ([] :: 'c list)"
+  assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
+  shows "P x l"
+  apply (rule_tac P="P x" in list.induct)
+  apply (rule a)
+  apply (rule b)
+  apply (assumption)
+  done
+
+lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
+apply (lifting list_induct_part)
+done
+
+lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
+apply (lifting list_induct_part)
+done
+
+lemma "P (x :: 'a fset) ([] :: 'c list) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (e # t)) \<Longrightarrow> P x l"
+apply (lifting list_induct_part)
+done
+
+quotient_type 'a fset2 = "'a list" / "list_eq"
+  by (rule equivp_list_eq)
+
+quotient_definition
+   "EMPTY2 :: 'a fset2"
+is
+  "[]::'a list"
+
+quotient_definition
+   "INSERT2 :: 'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"
+is
+  "op #"
+
+lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
+apply (lifting list_induct_part)
+done
+
+lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT2 e t)) \<Longrightarrow> P x l"
+apply (lifting list_induct_part)
+done
+
+quotient_definition
+  "fset_rec :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
+is
+  "list_rec"
+
+quotient_definition
+  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
+is
+  "list_case"
+
+(* Probably not true without additional assumptions about the function *)
+lemma list_rec_rsp[quot_respect]:
+  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
+  apply (auto)
+  apply (erule_tac list_eq.induct)
+  apply (simp_all)
+  sorry
+
+lemma list_case_rsp[quot_respect]:
+  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
+  apply (auto)
+  sorry
+
+lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
+apply (lifting list.recs(2))
+done
+
+lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"
+apply (lifting list.cases(2))
+done
+
+lemma ttt: "((op @) x ((op #) e [])) = (((op #) e x))"
+sorry
+
+lemma "(FUNION x (INSERT e EMPTY)) = ((INSERT e x))"
+apply (lifting ttt)
+done
+
+
+lemma ttt2: "(\<lambda>e. ((op @) x ((op #) e []))) = (\<lambda>e. ((op #) e x))"
+sorry
+
+lemma "(\<lambda>e. (FUNION x (INSERT e EMPTY))) = (\<lambda>e. (INSERT e x))"
+apply(lifting ttt2)
+apply(regularize)
+apply(rule impI)
+apply(simp)
+apply(rule allI)
+apply(rule list_eq_refl)
+done
+
+lemma ttt3: "(\<lambda>x. ((op @) x (e # []))) = (op #) e"
+sorry
+
+lemma "(\<lambda>x. (FUNION x (INSERT e EMPTY))) = INSERT e"
+apply(lifting ttt3)
+apply(regularize)
+apply(auto simp add: cons_rsp)
+done
+lemma hard: "(\<lambda>P. \<lambda>Q. P (Q (x::'a list))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a list)))"
+sorry
+
+lemma eq_imp_rel:
+  shows "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b"
+  by (simp add: equivp_reflp)
+
+
+lemma hard_lift: "(\<lambda>P. \<lambda>Q. P (Q (x::'a fset))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a fset)))"
+apply(lifting hard)
+apply(regularize)
+apply(rule fun_rel_id_asm)
+apply(subst babs_simp)
+apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
+apply(rule fun_rel_id_asm)
+apply(rule impI)
+apply(rule mp[OF eq_imp_rel[OF fset_equivp]])
+apply(drule fun_cong)
+apply(drule fun_cong)
+apply(assumption)
+done
+
+lemma test: "All (\<lambda>(x::'a list, y). x = y)"
+sorry
+
+lemma  "All (\<lambda>(x::'a fset, y). x = y)"
+apply(lifting test)
+done
+
+lemma test2: "Ex (\<lambda>(x::'a list, y). x = y)"
+sorry
+
+lemma  "Ex (\<lambda>(x::'a fset, y). x = y)"
+apply(lifting test2)
+done
+
+lemma test3: "All (\<lambda> (x :: 'a list, y, z). x = y \<and> y = z)"
+sorry
+
+lemma "All (\<lambda> (x :: 'a fset, y, z). x = y \<and> y = z)"
+apply(lifting test3)
+done
+
+lemma test4: "\<forall> (x :: 'a list, y, z) \<in> Respects(
+  prod_rel (op \<approx>) (prod_rel (op \<approx>) (op \<approx>))
+). x = y \<and> y = z"
+sorry
+
+lemma "All (\<lambda> (x :: 'a fset, y, z). x = y \<and> y = z)"
+apply (lifting test4)
+sorry
+
+lemma test5: "\<forall> (x :: 'a list \<Rightarrow> 'a list, y) \<in> Respects(
+  prod_rel (op \<approx> ===> op \<approx>) (op \<approx> ===> op \<approx>)
+). (op \<approx> ===> op \<approx>) x y"
+sorry
+
+lemma "All (\<lambda> (x :: 'a fset \<Rightarrow> 'a fset, y). x = y)"
+apply (lifting test5)
+done
+
+lemma test6: "\<forall> (x :: 'a list \<Rightarrow> 'a list, y, z) \<in> Respects(
+  prod_rel (op \<approx> ===> op \<approx>) (prod_rel (op \<approx> ===> op \<approx>) (op \<approx> ===> op \<approx>))
+). (op \<approx> ===> op \<approx>) x y \<and> (op \<approx> ===> op \<approx>) y z"
+sorry
+
+lemma "All (\<lambda> (x :: 'a fset \<Rightarrow> 'a fset, y, z). x = y \<and> y = z)"
+apply (lifting test6)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/FSet2.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,128 @@
+theory FSet2
+imports "../Quotient" "../Quotient_List" List
+begin
+
+inductive
+  list_eq (infix "\<approx>" 50)
+where
+  "a#b#xs \<approx> b#a#xs"
+| "[] \<approx> []"
+| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
+| "a#a#xs \<approx> a#xs"
+| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
+| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
+
+lemma list_eq_refl:
+  shows "xs \<approx> xs"
+by (induct xs) (auto intro: list_eq.intros)
+
+lemma equivp_list_eq:
+  shows "equivp list_eq"
+unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
+by (auto intro: list_eq.intros list_eq_refl)
+
+quotient_type
+  'a fset = "'a list" / "list_eq"
+  by (rule equivp_list_eq)
+
+quotient_definition
+  fempty ("{||}")
+where
+  "fempty :: 'a fset"
+is
+  "[]"
+
+quotient_definition
+  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+is
+  "(op #)"
+
+lemma finsert_rsp[quot_respect]:
+  shows "(op = ===> op \<approx> ===> op \<approx>) (op #) (op #)"
+by (auto intro: list_eq.intros)
+
+quotient_definition
+  funion ("_ \<union>f _" [65,66] 65)
+where
+  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+is
+  "(op @)"
+
+lemma append_rsp_aux1:
+  assumes a : "l2 \<approx> r2 "
+  shows "(h @ l2) \<approx> (h @ r2)"
+using a
+apply(induct h)
+apply(auto intro: list_eq.intros(5))
+done
+
+lemma append_rsp_aux2:
+  assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "
+  shows "(l1 @ l2) \<approx> (r1 @ r2)"
+using a
+apply(induct arbitrary: l2 r2)
+apply(simp_all)
+apply(blast intro: list_eq.intros append_rsp_aux1)+
+done
+
+lemma append_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
+  by (auto simp add: append_rsp_aux2)
+
+
+quotient_definition
+  fmem ("_ \<in>f _" [50, 51] 50)
+where
+  "fmem :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
+is
+  "(op mem)"
+
+lemma memb_rsp_aux:
+  assumes a: "x \<approx> y"
+  shows "(z mem x) = (z mem y)"
+  using a by induct auto
+
+lemma memb_rsp[quot_respect]:
+  shows "(op = ===> (op \<approx> ===> op =)) (op mem) (op mem)"
+  by (simp add: memb_rsp_aux)
+
+definition
+  fnot_mem :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ \<notin>f _" [50, 51] 50)
+where
+  "x \<notin>f S \<equiv> \<not>(x \<in>f S)"
+
+definition
+  "inter_list" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
+where
+  "inter_list X Y \<equiv> [x \<leftarrow> X. x\<in>set Y]"
+
+quotient_definition
+  finter ("_ \<inter>f _" [70, 71] 70)
+where
+  "finter::'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+is
+  "inter_list"
+
+no_syntax
+  "@Finset"     :: "args => 'a fset"                       ("{|(_)|}")
+syntax
+  "@Finfset"     :: "args => 'a fset"                       ("{|(_)|}")
+translations
+  "{|x, xs|}"     == "CONST finsert x {|xs|}"
+  "{|x|}"         == "CONST finsert x {||}"
+
+
+subsection {* Empty sets *}
+
+lemma test:
+  shows "\<not>(x # xs \<approx> [])"
+sorry
+
+lemma finsert_not_empty[simp]: 
+  shows "finsert x S \<noteq> {||}"
+  by (lifting test)
+
+
+
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/FSet3.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,717 @@
+theory FSet3
+imports "../Quotient" "../Quotient_List" List
+begin
+
+ML {*
+structure QuotientRules = Named_Thms
+  (val name = "quot_thm"
+   val description = "Quotient theorems.")
+*}
+
+ML {*
+open QuotientRules
+*}
+
+fun
+  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
+where
+  "list_eq xs ys = (\<forall>x. x \<in> set xs \<longleftrightarrow> x \<in> set ys)"
+
+lemma list_eq_equivp:
+  shows "equivp list_eq"
+unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
+by auto
+
+(* FIXME-TODO: because of beta-reduction, one cannot give the *)
+(* FIXME-TODO: relation as a term or abbreviation             *) 
+quotient_type 
+  'a fset = "'a list" / "list_eq"
+by (rule list_eq_equivp)
+
+
+section {* empty fset, finsert and membership *} 
+
+quotient_definition
+  fempty  ("{||}")
+where
+  "fempty :: 'a fset"
+is "[]::'a list"
+
+quotient_definition
+  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" 
+is "op #"
+
+syntax
+  "@Finset"     :: "args => 'a fset"  ("{|(_)|}")
+
+translations
+  "{|x, xs|}" == "CONST finsert x {|xs|}"
+  "{|x|}"     == "CONST finsert x {||}"
+
+definition 
+  memb :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
+where
+  "memb x xs \<equiv> x \<in> set xs"
+
+quotient_definition
+  fin ("_ |\<in>| _" [50, 51] 50)
+where
+  "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
+is "memb"
+
+abbreviation
+  fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ |\<notin>| _" [50, 51] 50)
+where
+  "a |\<notin>| S \<equiv> \<not>(a |\<in>| S)"
+
+lemma memb_rsp[quot_respect]: 
+  shows "(op = ===> op \<approx> ===> op =) memb memb"
+by (auto simp add: memb_def)
+
+lemma nil_rsp[quot_respect]:
+  shows "[] \<approx> []"
+by simp
+
+lemma cons_rsp[quot_respect]: 
+  shows "(op = ===> op \<approx> ===> op \<approx>) op # op #"
+by simp
+
+
+section {* Augmenting a set -- @{const finsert} *}
+
+text {* raw section *}
+
+lemma nil_not_cons:
+  shows "\<not>[] \<approx> x # xs"
+  by auto
+
+lemma memb_cons_iff:
+  shows "memb x (y # xs) = (x = y \<or> memb x xs)"
+  by (induct xs) (auto simp add: memb_def)
+
+lemma memb_consI1:
+  shows "memb x (x # xs)"
+  by (simp add: memb_def)
+
+lemma memb_consI2:
+  shows "memb x xs \<Longrightarrow> memb x (y # xs)"
+  by (simp add: memb_def)
+
+lemma memb_absorb:
+  shows "memb x xs \<Longrightarrow> x # xs \<approx> xs"
+  by (induct xs) (auto simp add: memb_def id_simps)
+
+text {* lifted section *}
+
+lemma fin_finsert_iff[simp]:
+  "x |\<in>| finsert y S = (x = y \<or> x |\<in>| S)"
+by (lifting memb_cons_iff) 
+
+lemma
+  shows finsertI1: "x |\<in>| finsert x S"
+  and   finsertI2: "x |\<in>| S \<Longrightarrow> x |\<in>| finsert y S"
+  by (lifting memb_consI1, lifting memb_consI2)
+
+
+lemma finsert_absorb [simp]: 
+  shows "x |\<in>| S \<Longrightarrow> finsert x S = S"
+  by (lifting memb_absorb)
+
+
+section {* Singletons *}
+
+text {* raw section *}
+
+lemma singleton_list_eq:
+  shows "[x] \<approx> [y] \<longleftrightarrow> x = y"
+  by (simp add: id_simps) auto
+
+text {* lifted section *}
+
+lemma fempty_not_finsert[simp]:
+  shows "{||} \<noteq> finsert x S"
+  by (lifting nil_not_cons)
+
+lemma fsingleton_eq[simp]:
+  shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
+  by (lifting singleton_list_eq)
+
+section {* Union *}
+
+quotient_definition
+  funion  (infixl "|\<union>|" 65)
+where
+  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+is
+  "op @"
+
+section {* Cardinality of finite sets *}
+
+fun
+  fcard_raw :: "'a list \<Rightarrow> nat"
+where
+  fcard_raw_nil:  "fcard_raw [] = 0"
+| fcard_raw_cons: "fcard_raw (x # xs) = (if memb x xs then fcard_raw xs else Suc (fcard_raw xs))"
+
+quotient_definition
+  "fcard :: 'a fset \<Rightarrow> nat" 
+is "fcard_raw"
+
+text {* raw section *}
+
+lemma fcard_raw_ge_0:
+  assumes a: "x \<in> set xs"
+  shows "0 < fcard_raw xs"
+using a
+by (induct xs) (auto simp add: memb_def)
+
+lemma fcard_raw_delete_one:
+  "fcard_raw ([x \<leftarrow> xs. x \<noteq> y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
+by (induct xs) (auto dest: fcard_raw_ge_0 simp add: memb_def)
+
+lemma fcard_raw_rsp_aux:
+  assumes a: "a \<approx> b"
+  shows "fcard_raw a = fcard_raw b"
+using a
+apply(induct a arbitrary: b)
+apply(auto simp add: memb_def)
+apply(metis)
+apply(drule_tac x="[x \<leftarrow> b. x \<noteq> a1]" in meta_spec)
+apply(simp add: fcard_raw_delete_one)
+apply(metis Suc_pred' fcard_raw_ge_0 fcard_raw_delete_one memb_def)
+done
+
+lemma fcard_raw_rsp[quot_respect]:
+  "(op \<approx> ===> op =) fcard_raw fcard_raw"
+  by (simp add: fcard_raw_rsp_aux)
+
+text {* lifted section *}
+
+lemma fcard_fempty [simp]:
+  shows "fcard {||} = 0"
+by (lifting fcard_raw_nil)
+
+lemma fcard_finsert_if [simp]:
+  shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))"
+by (lifting fcard_raw_cons)
+
+
+section {* Induction and Cases rules for finite sets *}
+
+lemma fset_exhaust[case_names fempty finsert, cases type: fset]:
+  shows "\<lbrakk>S = {||} \<Longrightarrow> P; \<And>x S'. S = finsert x S' \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+by (lifting list.exhaust)
+
+lemma fset_induct[case_names fempty finsert]:
+  shows "\<lbrakk>P {||}; \<And>x S. P S \<Longrightarrow> P (finsert x S)\<rbrakk> \<Longrightarrow> P S"
+by (lifting list.induct)
+
+lemma fset_induct2[case_names fempty finsert, induct type: fset]:
+  assumes prem1: "P {||}" 
+  and     prem2: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
+  shows "P S"
+proof(induct S rule: fset_induct)
+  case fempty
+  show "P {||}" by (rule prem1)
+next
+  case (finsert x S)
+  have asm: "P S" by fact
+  show "P (finsert x S)"
+  proof(cases "x |\<in>| S")
+    case True
+    have "x |\<in>| S" by fact
+    then show "P (finsert x S)" using asm by simp
+  next
+    case False
+    have "x |\<notin>| S" by fact
+    then show "P (finsert x S)" using prem2 asm by simp
+  qed
+qed
+
+
+section {* fmap and fset comprehension *}
+
+quotient_definition
+  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
+is
+ "map"
+
+quotient_definition
+  "fconcat :: ('a fset) fset => 'a fset"
+is
+ "concat"
+
+(*lemma fconcat_rsp[quot_respect]:
+  shows "((list_rel op \<approx>) ===> op \<approx>) concat concat"
+apply(auto)
+sorry
+*)
+
+(* PROBLEM: these lemmas needs to be restated, since  *)
+(* concat.simps(1) and concat.simps(2) contain the    *)
+(* type variables ?'a1.0 (which are turned into frees *)
+(* 'a_1                                               *)
+lemma concat1:
+  shows "concat [] \<approx> []"
+by (simp add: id_simps)
+
+lemma concat2:
+  shows "concat (x # xs) \<approx> x @ concat xs"
+by (simp add: id_simps)
+
+lemma concat_rsp[quot_respect]:
+  shows "(list_rel op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
+sorry
+
+lemma nil_rsp2[quot_respect]: "(list_rel op \<approx> OOO op \<approx>) [] []"
+  apply (metis FSet3.nil_rsp list_rel.simps(1) pred_comp.intros)
+  done
+
+lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
+  apply (rule eq_reflection)
+  apply auto
+  done
+
+lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
+  unfolding list_eq.simps
+  apply(simp only: set_map set_in_eq)
+  done
+
+lemma quotient_compose_list_pre:
+  "(list_rel op \<approx> OOO op \<approx>) r s =
+  ((list_rel op \<approx> OOO op \<approx>) r r \<and> (list_rel op \<approx> OOO op \<approx>) s s \<and>
+  abs_fset (map abs_fset r) = abs_fset (map abs_fset s))"
+  apply rule
+  apply rule
+  apply rule
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply rule
+  apply (rule equivp_reflp[OF fset_equivp])
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply(rule)
+  apply rule
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply rule
+  apply (rule equivp_reflp[OF fset_equivp])
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply (subgoal_tac "map abs_fset r \<approx> map abs_fset s")
+  apply (metis Quotient_rel[OF Quotient_fset])
+  apply (auto simp only:)[1]
+  apply (subgoal_tac "map abs_fset r = map abs_fset b")
+  prefer 2
+  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
+  apply (subgoal_tac "map abs_fset s = map abs_fset ba")
+  prefer 2
+  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
+  apply (simp only: map_rel_cong)
+  apply rule
+  apply (rule rep_abs_rsp[of "list_rel op \<approx>" "map abs_fset"])
+  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply rule
+  prefer 2
+  apply (rule rep_abs_rsp_left[of "list_rel op \<approx>" "map abs_fset"])
+  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply (erule conjE)+
+  apply (subgoal_tac "map abs_fset r \<approx> map abs_fset s")
+  prefer 2
+  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp)
+  apply (rule map_rel_cong)
+  apply (assumption)
+  done
+
+lemma quotient_compose_list[quot_thm]:
+  shows  "Quotient ((list_rel op \<approx>) OOO (op \<approx>))
+    (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
+  unfolding Quotient_def comp_def
+  apply (rule)+
+  apply (simp add: abs_o_rep[OF Quotient_fset] id_simps Quotient_abs_rep[OF Quotient_fset])
+  apply (rule)
+  apply (rule)
+  apply (rule)
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply (rule)
+  apply (rule equivp_reflp[OF fset_equivp])
+  apply (rule list_rel_refl)
+  apply (metis equivp_def fset_equivp)
+  apply rule
+  apply rule
+  apply(rule quotient_compose_list_pre)
+  done
+
+lemma fconcat_empty:
+  shows "fconcat {||} = {||}"
+apply(lifting concat1)
+apply(cleaning)
+apply(simp add: comp_def)
+apply(fold fempty_def[simplified id_simps])
+apply(rule refl)
+done
+
+(* Should be true *)
+
+lemma insert_rsp2[quot_respect]:
+  "(op \<approx> ===> list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx>) op # op #"
+apply auto
+apply (simp add: set_in_eq)
+sorry
+
+lemma append_rsp[quot_respect]:
+  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
+  by (auto)
+
+lemma fconcat_insert:
+  shows "fconcat (finsert x S) = x |\<union>| fconcat S"
+apply(lifting concat2)
+apply(cleaning)
+apply (simp add: finsert_def fconcat_def comp_def)
+apply cleaning
+done
+
+text {* raw section *}
+
+lemma map_rsp_aux:
+  assumes a: "a \<approx> b"
+  shows "map f a \<approx> map f b"
+  using a
+apply(induct a arbitrary: b)
+apply(auto)
+apply(metis rev_image_eqI)
+done
+
+lemma map_rsp[quot_respect]:
+  shows "(op = ===> op \<approx> ===> op \<approx>) map map"
+by (auto simp add: map_rsp_aux)
+
+
+text {* lifted section *}
+
+(* TBD *)
+
+text {* syntax for fset comprehensions (adapted from lists) *}
+
+nonterminals fsc_qual fsc_quals
+
+syntax
+"_fsetcompr" :: "'a \<Rightarrow> fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> 'a fset"  ("{|_ . __")
+"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ <- _")
+"_fsc_test" :: "bool \<Rightarrow> fsc_qual" ("_")
+"_fsc_end" :: "fsc_quals" ("|}")
+"_fsc_quals" :: "fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> fsc_quals" (", __")
+"_fsc_abs" :: "'a => 'b fset => 'b fset"
+
+syntax (xsymbols)
+"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
+syntax (HTML output)
+"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
+
+parse_translation (advanced) {*
+let
+  val femptyC = Syntax.const @{const_name fempty};
+  val finsertC = Syntax.const @{const_name finsert};
+  val fmapC = Syntax.const @{const_name fmap};
+  val fconcatC = Syntax.const @{const_name fconcat};
+  val IfC = Syntax.const @{const_name If};
+  fun fsingl x = finsertC $ x $ femptyC;
+
+  fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
+    let
+      val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
+      val e = if opti then fsingl e else e;
+      val case1 = Syntax.const "_case1" $ p $ e;
+      val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
+                                        $ femptyC;
+      val cs = Syntax.const "_case2" $ case1 $ case2
+      val ft = Datatype_Case.case_tr false Datatype.info_of_constr
+                 ctxt [x, cs]
+    in lambda x ft end;
+
+  fun abs_tr ctxt (p as Free(s,T)) e opti =
+        let val thy = ProofContext.theory_of ctxt;
+            val s' = Sign.intern_const thy s
+        in if Sign.declared_const thy s'
+           then (pat_tr ctxt p e opti, false)
+           else (lambda p e, true)
+        end
+    | abs_tr ctxt p e opti = (pat_tr ctxt p e opti, false);
+
+  fun fsc_tr ctxt [e, Const("_fsc_test",_) $ b, qs] =
+        let 
+          val res = case qs of 
+                      Const("_fsc_end",_) => fsingl e
+                    | Const("_fsc_quals",_)$ q $ qs => fsc_tr ctxt [e, q, qs];
+        in 
+          IfC $ b $ res $ femptyC 
+        end
+
+    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_end",_)] =
+         (case abs_tr ctxt p e true of
+            (f,true) => fmapC $ f $ es
+          | (f, false) => fconcatC $ (fmapC $ f $ es))
+       
+    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_quals",_) $ q $ qs] =
+        let
+          val e' = fsc_tr ctxt [e, q, qs];
+        in 
+          fconcatC $ (fmapC $ (fst (abs_tr ctxt p e' false)) $ es) 
+        end
+
+in [("_fsetcompr", fsc_tr)] end
+*}
+
+
+(* NEEDS FIXING *)
+(* examles *)
+(*
+term "{|(x,y,z). b|}"
+term "{|x. x \<leftarrow> xs|}"
+term "{|(x,y,z). x\<leftarrow>xs|}"
+term "{|e x y. x\<leftarrow>xs, y\<leftarrow>ys|}"
+term "{|(x,y,z). x<a, x>b|}"
+term "{|(x,y,z). x\<leftarrow>xs, x>b|}"
+term "{|(x,y,z). x<a, x\<leftarrow>xs|}"
+term "{|(x,y). Cons True x \<leftarrow> xs|}"
+term "{|(x,y,z). Cons x [] \<leftarrow> xs|}"
+term "{|(x,y,z). x<a, x>b, x=d|}"
+term "{|(x,y,z). x<a, x>b, y\<leftarrow>ys|}"
+term "{|(x,y,z). x<a, x\<leftarrow>xs,y>b|}"
+term "{|(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys|}"
+term "{|(x,y,z). x\<leftarrow>xs, x>b, y<a|}"
+term "{|(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys|}"
+term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x|}"
+term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs|}"
+*)
+
+(* BELOW CONSTRUCTION SITE *)
+
+
+lemma no_mem_nil: 
+  "(\<forall>a. a \<notin> set A) = (A = [])"
+by (induct A) (auto)
+
+lemma none_mem_nil: 
+  "(\<forall>a. a \<notin> set A) = (A \<approx> [])"
+by simp
+
+lemma mem_cons: 
+  "a \<in> set A \<Longrightarrow> a # A \<approx> A"
+by auto
+
+lemma cons_left_comm:
+  "x # y # A \<approx> y # x # A"
+by (auto simp add: id_simps)
+
+lemma cons_left_idem:
+  "x # x # A \<approx> x # A"
+by (auto simp add: id_simps)
+
+lemma finite_set_raw_strong_cases:
+  "(X = []) \<or> (\<exists>a Y. ((a \<notin> set Y) \<and> (X \<approx> a # Y)))"
+  apply (induct X)
+  apply (simp)
+  apply (rule disjI2)
+  apply (erule disjE)
+  apply (rule_tac x="a" in exI)
+  apply (rule_tac x="[]" in exI)
+  apply (simp)
+  apply (erule exE)+
+  apply (case_tac "a = aa")
+  apply (rule_tac x="a" in exI)
+  apply (rule_tac x="Y" in exI)
+  apply (simp)
+  apply (rule_tac x="aa" in exI)
+  apply (rule_tac x="a # Y" in exI)
+  apply (auto)
+  done
+
+fun
+  delete_raw :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
+where
+  "delete_raw [] x = []"
+| "delete_raw (a # A) x = (if (a = x) then delete_raw A x else a # (delete_raw A x))"
+
+lemma mem_delete_raw:
+  "x \<in> set (delete_raw A a) = (x \<in> set A \<and> \<not>(x = a))"
+  by (induct A arbitrary: x a) (auto)
+
+lemma mem_delete_raw_ident:
+  "\<not>(a \<in> set (delete_raw A a))"
+by (induct A) (auto)
+
+lemma not_mem_delete_raw_ident:
+  "b \<notin> set A \<Longrightarrow> (delete_raw A b = A)"
+by (induct A) (auto)
+
+lemma delete_raw_RSP:
+  "A \<approx> B \<Longrightarrow> delete_raw A a \<approx> delete_raw B a"
+apply(induct A arbitrary: B a)
+apply(auto)
+sorry
+
+lemma cons_delete_raw:
+  "a # (delete_raw A a) \<approx> (if a \<in> set A then A else (a # A))"
+sorry
+
+lemma mem_cons_delete_raw:
+    "a \<in> set A \<Longrightarrow> a # (delete_raw A a) \<approx> A"
+sorry
+
+lemma finite_set_raw_delete_raw_cases:
+    "X = [] \<or> (\<exists>a. a mem X \<and> X \<approx> a # delete_raw X a)"
+  by (induct X) (auto)
+
+
+
+
+
+lemma list2set_thm:
+  shows "set [] = {}"
+  and "set (h # t) = insert h (set t)"
+  by (auto)
+
+lemma list2set_RSP:
+  "A \<approx> B \<Longrightarrow> set A = set B"
+  by auto
+
+definition
+  rsp_fold
+where
+  "rsp_fold f = (\<forall>u v w. (f u (f v w) = f v (f u w)))"
+
+primrec
+  fold_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b"
+where
+  "fold_raw f z [] = z"
+| "fold_raw f z (a # A) =
+     (if (rsp_fold f) then
+       if a mem A then fold_raw f z A
+       else f a (fold_raw f z A)
+     else z)"
+
+lemma mem_lcommuting_fold_raw:
+  "rsp_fold f \<Longrightarrow> h mem B \<Longrightarrow> fold_raw f z B = f h (fold_raw f z (delete_raw B h))"
+sorry
+
+lemma fold_rsp[quot_respect]:
+  "(op = ===> op = ===> op \<approx> ===> op =) fold_raw fold_raw"
+apply(auto)
+sorry
+
+primrec
+  inter_raw
+where
+  "inter_raw [] B = []"
+| "inter_raw (a # A) B = (if a mem B then a # inter_raw A B else inter_raw A B)"
+
+lemma mem_inter_raw:
+  "x mem (inter_raw A B) = x mem A \<and> x mem B"
+sorry
+
+lemma inter_raw_RSP:
+  "A1 \<approx> A2 \<and> B1 \<approx> B2 \<Longrightarrow> (inter_raw A1 B1) \<approx> (inter_raw A2 B2)"
+sorry
+
+
+(* LIFTING DEFS *)
+
+
+section {* Constants on the Quotient Type *} 
+
+
+quotient_definition
+  "fdelete :: 'a fset \<Rightarrow> 'a \<Rightarrow> 'a fset" 
+  is "delete_raw"
+
+quotient_definition
+  finter ("_ \<inter>f _" [70, 71] 70)
+where
+  "finter :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
+  is "inter_raw"
+
+quotient_definition
+  "ffold :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" 
+  is "fold_raw"
+
+quotient_definition
+  "fset_to_set :: 'a fset \<Rightarrow> 'a set" 
+  is "set"
+
+
+section {* Lifted Theorems *}
+
+thm list.cases (* ??? *)
+
+thm cons_left_comm
+lemma "finsert a (finsert b S) = finsert b (finsert a S)"
+by (lifting cons_left_comm)
+
+thm cons_left_idem
+lemma "finsert a (finsert a S) = finsert a S"
+by (lifting cons_left_idem)
+
+(* thm MEM:
+  MEM x [] = F
+  MEM x (h::t) = (x=h) \/ MEM x t *)
+thm none_mem_nil
+(*lemma "(\<forall>a. a \<notin>f A) = (A = fempty)"*)
+
+thm mem_cons
+thm finite_set_raw_strong_cases
+(*thm card_raw.simps*)
+(*thm not_mem_card_raw*)
+(*thm card_raw_suc*)
+
+lemma "fcard X = Suc n \<Longrightarrow> (\<exists>a S. a \<notin>f S & X = finsert a S)"
+(*by (lifting card_raw_suc)*)
+sorry
+
+(*thm card_raw_cons_gt_0
+thm mem_card_raw_gt_0
+thm not_nil_equiv_cons
+*)
+thm delete_raw.simps
+(*thm mem_delete_raw*)
+(*thm card_raw_delete_raw*)
+thm cons_delete_raw
+thm mem_cons_delete_raw
+thm finite_set_raw_delete_raw_cases
+thm append.simps
+(* MEM_APPEND: MEM e (APPEND l1 l2) = MEM e l1 \/ MEM e l2 *)
+thm inter_raw.simps
+thm mem_inter_raw
+thm fold_raw.simps
+thm list2set_thm
+thm list_eq_def
+thm list.induct
+lemma "\<lbrakk>P fempty; \<And>a x. P x \<Longrightarrow> P (finsert a x)\<rbrakk> \<Longrightarrow> P l"
+by (lifting list.induct)
+
+(* We also have map and some properties of it in FSet *)
+(* and the following which still lifts ok *)
+lemma "funion (funion x xa) xb = funion x (funion xa xb)"
+by (lifting append_assoc)
+
+quotient_definition
+  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
+is
+  "list_case"
+
+(* NOT SURE IF TRUE *)
+lemma list_case_rsp[quot_respect]:
+  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
+  apply (auto)
+  sorry
+
+lemma "fset_case (f1::'t) f2 (finsert a xa) = f2 a xa"
+apply (lifting list.cases(2))
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/IntEx.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,277 @@
+theory IntEx
+imports "../Quotient_Product" "../Quotient_List"
+begin
+
+fun
+  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
+where
+  "intrel (x, y) (u, v) = (x + v = u + y)"
+
+quotient_type 
+  my_int = "nat \<times> nat" / intrel
+  apply(unfold equivp_def)
+  apply(auto simp add: mem_def expand_fun_eq)
+  done
+
+quotient_definition
+  "ZERO :: my_int"
+is
+  "(0::nat, 0::nat)"
+
+quotient_definition
+  "ONE :: my_int"
+is
+  "(1::nat, 0::nat)"
+
+fun
+  my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "my_plus (x, y) (u, v) = (x + u, y + v)"
+
+quotient_definition
+  "PLUS :: my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+is
+  "my_plus"
+
+fun
+  my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "my_neg (x, y) = (y, x)"
+
+quotient_definition
+  "NEG :: my_int \<Rightarrow> my_int"
+is
+  "my_neg"
+
+definition
+  MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+where
+  "MINUS z w = PLUS z (NEG w)"
+
+fun
+  my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
+
+quotient_definition
+  "MULT :: my_int \<Rightarrow> my_int \<Rightarrow> my_int"
+is
+  "my_mult"
+
+
+(* NOT SURE WETHER THIS DEFINITION IS CORRECT *)
+fun
+  my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
+where
+  "my_le (x, y) (u, v) = (x+v \<le> u+y)"
+
+quotient_definition
+   "LE :: my_int \<Rightarrow> my_int \<Rightarrow> bool"
+is
+  "my_le"
+
+term LE
+thm LE_def
+
+
+definition
+  LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
+where
+  "LESS z w = (LE z w \<and> z \<noteq> w)"
+
+term LESS
+thm LESS_def
+
+definition
+  ABS :: "my_int \<Rightarrow> my_int"
+where
+  "ABS i = (if (LESS i ZERO) then (NEG i) else i)"
+
+definition
+  SIGN :: "my_int \<Rightarrow> my_int"
+where
+ "SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))"
+
+print_quotconsts
+
+lemma plus_sym_pre:
+  shows "my_plus a b \<approx> my_plus b a"
+  apply(cases a)
+  apply(cases b)
+  apply(auto)
+  done
+
+lemma plus_rsp[quot_respect]:
+  shows "(intrel ===> intrel ===> intrel) my_plus my_plus"
+by (simp)
+
+lemma neg_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx>) my_neg my_neg"
+by simp
+
+lemma test1: "my_plus a b = my_plus a b"
+apply(rule refl)
+done
+
+lemma "PLUS a b = PLUS a b"
+apply(lifting_setup test1)
+apply(regularize)
+apply(injection)
+apply(cleaning)
+done
+
+thm lambda_prs
+
+lemma test2: "my_plus a = my_plus a"
+apply(rule refl)
+done
+
+lemma "PLUS a = PLUS a"
+apply(lifting_setup test2)
+apply(rule impI)
+apply(rule ballI)
+apply(rule apply_rsp[OF Quotient_my_int plus_rsp])
+apply(simp only: in_respects)
+apply(injection)
+apply(cleaning)
+done
+
+lemma test3: "my_plus = my_plus"
+apply(rule refl)
+done
+
+lemma "PLUS = PLUS"
+apply(lifting_setup test3)
+apply(rule impI)
+apply(rule plus_rsp)
+apply(injection)
+apply(cleaning)
+done
+
+
+lemma "PLUS a b = PLUS b a"
+apply(lifting plus_sym_pre)
+done
+
+lemma plus_assoc_pre:
+  shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"
+  apply (cases i)
+  apply (cases j)
+  apply (cases k)
+  apply (simp)
+  done
+
+lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)"
+apply(lifting plus_assoc_pre)
+done
+
+lemma int_induct_raw:
+  assumes a: "P (0::nat, 0)"
+  and     b: "\<And>i. P i \<Longrightarrow> P (my_plus i (1,0))"
+  and     c: "\<And>i. P i \<Longrightarrow> P (my_plus i (my_neg (1,0)))"
+  shows      "P x"
+  apply(case_tac x) apply(simp)
+  apply(rule_tac x="b" in spec)
+  apply(rule_tac Nat.induct)
+  apply(rule allI)
+  apply(rule_tac Nat.induct)
+  using a b c apply(auto)
+  done
+
+lemma int_induct:
+  assumes a: "P ZERO"
+  and     b: "\<And>i. P i \<Longrightarrow> P (PLUS i ONE)"
+  and     c: "\<And>i. P i \<Longrightarrow> P (PLUS i (NEG ONE))"
+  shows      "P x"
+  using a b c
+  by (lifting int_induct_raw)
+
+lemma ex1tst: "Bex1_rel (op \<approx>) (\<lambda>x :: nat \<times> nat. x \<approx> x)"
+sorry
+
+lemma ex1tst': "\<exists>!(x::my_int). x = x"
+apply(lifting ex1tst)
+done
+
+
+lemma ho_tst: "foldl my_plus x [] = x"
+apply simp
+done
+
+
+term foldl
+lemma "foldl PLUS x [] = x"
+apply(lifting ho_tst)
+done
+
+lemma ho_tst2: "foldl my_plus x (h # t) \<approx> my_plus h (foldl my_plus x t)"
+sorry
+
+lemma "foldl PLUS x (h # t) = PLUS h (foldl PLUS x t)"
+apply(lifting ho_tst2)
+done
+
+lemma ho_tst3: "foldl f (s::nat \<times> nat) ([]::(nat \<times> nat) list) = s"
+by simp
+
+lemma "foldl f (x::my_int) ([]::my_int list) = x"
+apply(lifting ho_tst3)
+done
+
+lemma lam_tst: "(\<lambda>x. (x, x)) y = (y, (y :: nat \<times> nat))"
+by simp
+
+(* Don't know how to keep the goal non-contracted... *)
+lemma "(\<lambda>x. (x, x)) (y::my_int) = (y, y)"
+apply(lifting lam_tst)
+done
+
+lemma lam_tst2: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
+by simp
+
+lemma
+  shows "equivp (op \<approx>)"
+  and "equivp ((op \<approx>) ===> (op \<approx>))"
+(* Nitpick finds a counterexample! *)
+oops
+
+lemma lam_tst3a: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
+by auto
+
+lemma id_rsp:
+  shows "(R ===> R) id id"
+by simp
+
+lemma lam_tst3a_reg: "(op \<approx> ===> op \<approx>) (Babs (Respects op \<approx>) (\<lambda>y. y)) (Babs (Respects op \<approx>) (\<lambda>x. x))"
+apply (rule babs_rsp[OF Quotient_my_int])
+apply (simp add: id_rsp)
+done
+
+lemma "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)"
+apply(lifting lam_tst3a)
+apply(rule impI)
+apply(rule lam_tst3a_reg)
+done
+
+lemma lam_tst3b: "(\<lambda>(y :: nat \<times> nat \<Rightarrow> nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat \<Rightarrow> nat \<times> nat). x)"
+by auto
+
+lemma "(\<lambda>(y :: my_int => my_int). y) = (\<lambda>(x :: my_int => my_int). x)"
+apply(lifting lam_tst3b)
+apply(rule impI)
+apply(rule babs_rsp[OF fun_quotient[OF Quotient_my_int Quotient_my_int]])
+apply(simp add: id_rsp)
+done
+
+lemma lam_tst4: "map (\<lambda>x. my_plus x (0,0)) l = l"
+apply (induct l)
+apply simp
+apply (case_tac a)
+apply simp
+done
+
+lemma "map (\<lambda>x. PLUS x ZERO) l = l"
+apply(lifting lam_tst4)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/IntEx2.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,445 @@
+theory IntEx2
+imports "../Quotient" "../Quotient_Product" Nat
+(*uses
+  ("Tools/numeral.ML")
+  ("Tools/numeral_syntax.ML")
+  ("Tools/int_arith.ML")*)
+begin
+
+fun
+  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
+where
+  "intrel (x, y) (u, v) = (x + v = u + y)"
+
+quotient_type int = "nat \<times> nat" / intrel
+  unfolding equivp_def
+  by (auto simp add: mem_def expand_fun_eq)
+
+instantiation int :: "{zero, one, plus, uminus, minus, times, ord, abs, sgn}"
+begin
+
+ML {* @{term "0 \<Colon> int"} *}
+
+quotient_definition
+  "0 \<Colon> int" is "(0\<Colon>nat, 0\<Colon>nat)"
+
+quotient_definition
+  "1 \<Colon> int" is "(1\<Colon>nat, 0\<Colon>nat)"
+
+fun
+  plus_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "plus_raw (x, y) (u, v) = (x + u, y + v)"
+
+quotient_definition
+  "(op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int)" is "plus_raw"
+
+fun
+  uminus_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "uminus_raw (x, y) = (y, x)"
+
+quotient_definition
+  "(uminus \<Colon> (int \<Rightarrow> int))" is "uminus_raw"
+
+definition
+  minus_int_def [code del]:  "z - w = z + (-w\<Colon>int)"
+
+fun
+  mult_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
+where
+  "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
+
+quotient_definition
+  mult_int_def: "(op *) :: (int \<Rightarrow> int \<Rightarrow> int)" is "mult_raw"
+
+fun
+  le_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
+where
+  "le_raw (x, y) (u, v) = (x+v \<le> u+y)"
+
+quotient_definition
+  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" is "le_raw"
+
+definition
+  less_int_def [code del]: "(z\<Colon>int) < w = (z \<le> w \<and> z \<noteq> w)"
+
+definition
+  zabs_def: "\<bar>i\<Colon>int\<bar> = (if i < 0 then - i else i)"
+
+definition
+  zsgn_def: "sgn (i\<Colon>int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
+
+instance ..
+
+end
+
+lemma plus_raw_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) plus_raw plus_raw"
+by auto
+
+lemma uminus_raw_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx>) uminus_raw uminus_raw"
+  by auto
+
+lemma mult_raw_fst:
+  assumes a: "x \<approx> z"
+  shows "mult_raw x y \<approx> mult_raw z y"
+using a
+apply(cases x, cases y, cases z)
+apply(auto simp add: mult_raw.simps intrel.simps)
+apply(rename_tac u v w x y z)
+apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
+apply(simp add: mult_ac)
+apply(simp add: add_mult_distrib [symmetric])
+done
+
+lemma mult_raw_snd:
+  assumes a: "x \<approx> z"
+  shows "mult_raw y x \<approx> mult_raw y z"
+using a
+apply(cases x, cases y, cases z)
+apply(auto simp add: mult_raw.simps intrel.simps)
+apply(rename_tac u v w x y z)
+apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
+apply(simp add: mult_ac)
+apply(simp add: add_mult_distrib [symmetric])
+done
+
+lemma mult_raw_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) mult_raw mult_raw"
+apply(simp only: fun_rel_def)
+apply(rule allI | rule impI)+
+apply(rule equivp_transp[OF int_equivp])
+apply(rule mult_raw_fst)
+apply(assumption)
+apply(rule mult_raw_snd)
+apply(assumption)
+done
+
+lemma le_raw_rsp[quot_respect]:
+  shows "(op \<approx> ===> op \<approx> ===> op =) le_raw le_raw"
+by auto
+
+lemma plus_assoc_raw:
+  shows "plus_raw (plus_raw i j) k \<approx> plus_raw i (plus_raw j k)"
+by (cases i, cases j, cases k) (simp)
+
+lemma plus_sym_raw:
+  shows "plus_raw i j \<approx> plus_raw j i"
+by (cases i, cases j) (simp)
+
+lemma plus_zero_raw:
+  shows "plus_raw  (0, 0) i \<approx> i"
+by (cases i) (simp)
+
+lemma plus_minus_zero_raw:
+  shows "plus_raw (uminus_raw i) i \<approx> (0, 0)"
+by (cases i) (simp)
+
+lemma times_assoc_raw:
+  shows "mult_raw (mult_raw i j) k \<approx> mult_raw i (mult_raw j k)"
+by (cases i, cases j, cases k) 
+   (simp add: algebra_simps)
+
+lemma times_sym_raw:
+  shows "mult_raw i j \<approx> mult_raw j i"
+by (cases i, cases j) (simp add: algebra_simps)
+
+lemma times_one_raw:
+  shows "mult_raw  (1, 0) i \<approx> i"
+by (cases i) (simp)
+
+lemma times_plus_comm_raw:
+  shows "mult_raw (plus_raw i j) k \<approx> plus_raw (mult_raw i k) (mult_raw j k)"
+by (cases i, cases j, cases k) 
+   (simp add: algebra_simps)
+
+lemma one_zero_distinct:
+  shows "\<not> (0, 0) \<approx> ((1::nat), (0::nat))"
+  by simp
+
+text{* The integers form a @{text comm_ring_1}*}
+
+instance int :: comm_ring_1
+proof
+  fix i j k :: int
+  show "(i + j) + k = i + (j + k)"
+    by (lifting plus_assoc_raw)
+  show "i + j = j + i" 
+    by (lifting plus_sym_raw)
+  show "0 + i = (i::int)"
+    by (lifting plus_zero_raw)
+  show "- i + i = 0"
+    by (lifting plus_minus_zero_raw)
+  show "i - j = i + - j"
+    by (simp add: minus_int_def)
+  show "(i * j) * k = i * (j * k)"
+    by (lifting times_assoc_raw)
+  show "i * j = j * i"
+    by (lifting times_sym_raw)
+  show "1 * i = i"
+    by (lifting times_one_raw)
+  show "(i + j) * k = i * k + j * k"
+    by (lifting times_plus_comm_raw)
+  show "0 \<noteq> (1::int)"
+    by (lifting one_zero_distinct)
+qed
+
+lemma plus_raw_rsp_aux:
+  assumes a: "a \<approx> b" "c \<approx> d"
+  shows "plus_raw a c \<approx> plus_raw b d"
+using a
+by (cases a, cases b, cases c, cases d) 
+   (simp)
+
+lemma add:
+     "(abs_int (x,y)) + (abs_int (u,v)) =
+      (abs_int (x + u, y + v))"
+apply(simp add: plus_int_def id_simps)
+apply(fold plus_raw.simps)
+apply(rule Quotient_rel_abs[OF Quotient_int])
+apply(rule plus_raw_rsp_aux)
+apply(simp_all add: rep_abs_rsp_left[OF Quotient_int])
+done
+
+definition int_of_nat_raw: 
+  "int_of_nat_raw m = (m :: nat, 0 :: nat)"
+
+quotient_definition
+  "int_of_nat :: nat \<Rightarrow> int" is "int_of_nat_raw"
+
+lemma[quot_respect]: 
+  shows "(op = ===> op \<approx>) int_of_nat_raw int_of_nat_raw"
+by (simp add: equivp_reflp[OF int_equivp])
+
+lemma int_of_nat:
+  shows "of_nat m = int_of_nat m"
+apply (induct m)
+apply (simp_all add: zero_int_def one_int_def int_of_nat_def int_of_nat_raw add)
+done
+
+lemma le_antisym_raw:
+  shows "le_raw i j \<Longrightarrow> le_raw j i \<Longrightarrow> i \<approx> j"
+by (cases i, cases j) (simp)
+
+lemma le_refl_raw:
+  shows "le_raw i i"
+by (cases i) (simp)
+
+lemma le_trans_raw:
+  shows "le_raw i j \<Longrightarrow> le_raw j k \<Longrightarrow> le_raw i k"
+by (cases i, cases j, cases k) (simp)
+
+lemma le_cases_raw:
+  shows "le_raw i j \<or> le_raw j i"
+by (cases i, cases j) 
+   (simp add: linorder_linear)
+
+instance int :: linorder
+proof
+  fix i j k :: int
+  show antisym: "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j"
+    by (lifting le_antisym_raw)
+  show "(i < j) = (i \<le> j \<and> \<not> j \<le> i)"
+    by (auto simp add: less_int_def dest: antisym) 
+  show "i \<le> i"
+    by (lifting le_refl_raw)
+  show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k"
+    by (lifting le_trans_raw)
+  show "i \<le> j \<or> j \<le> i"
+    by (lifting le_cases_raw)
+qed
+
+instantiation int :: distrib_lattice
+begin
+
+definition
+  "(inf \<Colon> int \<Rightarrow> int \<Rightarrow> int) = min"
+
+definition
+  "(sup \<Colon> int \<Rightarrow> int \<Rightarrow> int) = max"
+
+instance
+  by intro_classes
+    (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1)
+
+end
+
+lemma le_plus_raw:
+  shows "le_raw i j \<Longrightarrow> le_raw (plus_raw k i) (plus_raw k j)"
+by (cases i, cases j, cases k) (simp)
+
+
+instance int :: ordered_cancel_ab_semigroup_add
+proof
+  fix i j k :: int
+  show "i \<le> j \<Longrightarrow> k + i \<le> k + j"
+    by (lifting le_plus_raw)
+qed
+
+abbreviation
+  "less_raw i j \<equiv> le_raw i j \<and> \<not>(i \<approx> j)"
+
+lemma zmult_zless_mono2_lemma:
+  fixes i j::int
+  and   k::nat
+  shows "i < j \<Longrightarrow> 0 < k \<Longrightarrow> of_nat k * i < of_nat k * j"
+apply(induct "k")
+apply(simp)
+apply(case_tac "k = 0")
+apply(simp_all add: left_distrib add_strict_mono)
+done
+
+lemma zero_le_imp_eq_int_raw:
+  fixes k::"(nat \<times> nat)"
+  shows "less_raw (0, 0) k \<Longrightarrow> (\<exists>n > 0. k \<approx> int_of_nat_raw n)"
+apply(cases k)
+apply(simp add:int_of_nat_raw)
+apply(auto)
+apply(rule_tac i="b" and j="a" in less_Suc_induct)
+apply(auto)
+done
+
+lemma zero_le_imp_eq_int:
+  fixes k::int
+  shows "0 < k \<Longrightarrow> \<exists>n > 0. k = of_nat n"
+  unfolding less_int_def int_of_nat
+  by (lifting zero_le_imp_eq_int_raw)
+
+lemma zmult_zless_mono2: 
+  fixes i j k::int
+  assumes a: "i < j" "0 < k"
+  shows "k * i < k * j"
+using a
+using a
+apply(drule_tac zero_le_imp_eq_int)
+apply(auto simp add: zmult_zless_mono2_lemma)
+done
+
+text{*The integers form an ordered integral domain*}
+instance int :: linordered_idom
+proof
+  fix i j k :: int
+  show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"
+    by (rule zmult_zless_mono2)
+  show "\<bar>i\<bar> = (if i < 0 then -i else i)"
+    by (simp only: zabs_def)
+  show "sgn (i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
+    by (simp only: zsgn_def)
+qed
+
+lemmas int_distrib =
+  left_distrib [of "z1::int" "z2" "w", standard]
+  right_distrib [of "w::int" "z1" "z2", standard]
+  left_diff_distrib [of "z1::int" "z2" "w", standard]
+  right_diff_distrib [of "w::int" "z1" "z2", standard]
+
+
+subsection {* Embedding of the Integers into any @{text ring_1}: @{text of_int}*}
+
+(*
+context ring_1
+begin
+
+ 
+definition 
+  of_int :: "int \<Rightarrow> 'a" 
+where
+  "of_int 
+*)
+
+
+subsection {* Binary representation *}
+
+text {*
+  This formalization defines binary arithmetic in terms of the integers
+  rather than using a datatype. This avoids multiple representations (leading
+  zeroes, etc.)  See @{text "ZF/Tools/twos-compl.ML"}, function @{text
+  int_of_binary}, for the numerical interpretation.
+
+  The representation expects that @{text "(m mod 2)"} is 0 or 1,
+  even if m is negative;
+  For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
+  @{text "-5 = (-3)*2 + 1"}.
+  
+  This two's complement binary representation derives from the paper 
+  "An Efficient Representation of Arithmetic for Term Rewriting" by
+  Dave Cohen and Phil Watson, Rewriting Techniques and Applications,
+  Springer LNCS 488 (240-251), 1991.
+*}
+
+subsubsection {* The constructors @{term Bit0}, @{term Bit1}, @{term Pls} and @{term Min} *}
+
+definition
+  Pls :: int where
+  [code del]: "Pls = 0"
+
+definition
+  Min :: int where
+  [code del]: "Min = - 1"
+
+definition
+  Bit0 :: "int \<Rightarrow> int" where
+  [code del]: "Bit0 k = k + k"
+
+definition
+  Bit1 :: "int \<Rightarrow> int" where
+  [code del]: "Bit1 k = 1 + k + k"
+
+class number = -- {* for numeric types: nat, int, real, \dots *}
+  fixes number_of :: "int \<Rightarrow> 'a"
+
+(*use "~~/src/HOL/Tools/numeral.ML"
+
+syntax
+  "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
+
+use "~~/src/HOL/Tools/numeral_syntax.ML"
+
+setup NumeralSyntax.setup
+
+abbreviation
+  "Numeral0 \<equiv> number_of Pls"
+
+abbreviation
+  "Numeral1 \<equiv> number_of (Bit1 Pls)"
+
+lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)"
+  -- {* Unfold all @{text let}s involving constants *}
+  unfolding Let_def ..
+
+definition
+  succ :: "int \<Rightarrow> int" where
+  [code del]: "succ k = k + 1"
+
+definition
+  pred :: "int \<Rightarrow> int" where
+  [code del]: "pred k = k - 1"
+
+lemmas
+  max_number_of [simp] = max_def
+    [of "number_of u" "number_of v", standard, simp]
+and
+  min_number_of [simp] = min_def 
+    [of "number_of u" "number_of v", standard, simp]
+  -- {* unfolding @{text minx} and @{text max} on numerals *}
+
+lemmas numeral_simps = 
+  succ_def pred_def Pls_def Min_def Bit0_def Bit1_def
+
+text {* Removal of leading zeroes *}
+
+lemma Bit0_Pls [simp, code_post]:
+  "Bit0 Pls = Pls"
+  unfolding numeral_simps by simp
+
+lemma Bit1_Min [simp, code_post]:
+  "Bit1 Min = Min"
+  unfolding numeral_simps by simp
+
+lemmas normalize_bin_simps =
+  Bit0_Pls Bit1_Min
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/LFex.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,297 @@
+theory LFex
+imports Nominal "../Quotient_List"
+begin
+
+atom_decl name ident
+
+nominal_datatype kind = 
+    Type
+  | KPi "ty" "name" "kind"
+and ty =  
+    TConst "ident"
+  | TApp "ty" "trm"
+  | TPi "ty" "name" "ty"
+and trm = 
+    Const "ident"
+  | Var "name"
+  | App "trm" "trm"
+  | Lam "ty" "name" "trm" 
+
+function
+    fv_kind :: "kind \<Rightarrow> name set"
+and fv_ty   :: "ty \<Rightarrow> name set"
+and fv_trm  :: "trm \<Rightarrow> name set"
+where
+  "fv_kind (Type) = {}"
+| "fv_kind (KPi A x K) = (fv_ty A) \<union> ((fv_kind K) - {x})"
+| "fv_ty (TConst i) = {}"
+| "fv_ty (TApp A M) = (fv_ty A) \<union> (fv_trm M)"
+| "fv_ty (TPi A x B) = (fv_ty A) \<union> ((fv_ty B) - {x})"
+| "fv_trm (Const i) = {}"
+| "fv_trm (Var x) = {x}"
+| "fv_trm (App M N) = (fv_trm M) \<union> (fv_trm N)"
+| "fv_trm (Lam A x M) = (fv_ty A) \<union> ((fv_trm M) - {x})"
+sorry
+
+termination fv_kind sorry
+
+inductive
+    akind :: "kind \<Rightarrow> kind \<Rightarrow> bool" ("_ \<approx>ki _" [100, 100] 100)
+and aty   :: "ty \<Rightarrow> ty \<Rightarrow> bool"     ("_ \<approx>ty _" [100, 100] 100)
+and atrm  :: "trm \<Rightarrow> trm \<Rightarrow> bool"   ("_ \<approx>tr _" [100, 100] 100)
+where
+  a1:  "(Type) \<approx>ki (Type)"
+| a21: "\<lbrakk>A \<approx>ty A'; K \<approx>ki K'\<rbrakk> \<Longrightarrow> (KPi A x K) \<approx>ki (KPi A' x K')"
+| a22: "\<lbrakk>A \<approx>ty A'; K \<approx>ki ([(x,x')]\<bullet>K'); x \<notin> (fv_ty A'); x \<notin> ((fv_kind K') - {x'})\<rbrakk> 
+        \<Longrightarrow> (KPi A x K) \<approx>ki (KPi A' x' K')"
+| a3:  "i = j \<Longrightarrow> (TConst i) \<approx>ty (TConst j)"
+| a4:  "\<lbrakk>A \<approx>ty A'; M \<approx>tr M'\<rbrakk> \<Longrightarrow> (TApp A M) \<approx>ty (TApp A' M')"
+| a51: "\<lbrakk>A \<approx>ty A'; B \<approx>ty B'\<rbrakk> \<Longrightarrow> (TPi A x B) \<approx>ty (TPi A' x B')"
+| a52: "\<lbrakk>A \<approx>ty A'; B \<approx>ty ([(x,x')]\<bullet>B'); x \<notin> (fv_ty B'); x \<notin> ((fv_ty B') - {x'})\<rbrakk> 
+        \<Longrightarrow> (TPi A x B) \<approx>ty (TPi A' x' B')"
+| a6:  "i = j \<Longrightarrow> (Const i) \<approx>trm (Const j)"
+| a7:  "x = y \<Longrightarrow> (Var x) \<approx>trm (Var y)"
+| a8:  "\<lbrakk>M \<approx>trm M'; N \<approx>tr N'\<rbrakk> \<Longrightarrow> (App M N) \<approx>tr (App M' N')"
+| a91: "\<lbrakk>A \<approx>ty A'; M \<approx>tr M'\<rbrakk> \<Longrightarrow> (Lam A x M) \<approx>tr (Lam A' x M')"
+| a92: "\<lbrakk>A \<approx>ty A'; M \<approx>tr ([(x,x')]\<bullet>M'); x \<notin> (fv_ty B'); x \<notin> ((fv_trm M') - {x'})\<rbrakk> 
+        \<Longrightarrow> (Lam A x M) \<approx>tr (Lam A' x' M')"
+
+lemma al_refl:
+  fixes K::"kind" 
+  and   A::"ty"
+  and   M::"trm"
+  shows "K \<approx>ki K"
+  and   "A \<approx>ty A"
+  and   "M \<approx>tr M"
+  apply(induct K and A and M rule: kind_ty_trm.inducts)
+  apply(auto intro: akind_aty_atrm.intros)
+  done
+
+lemma alpha_equivps:
+  shows "equivp akind"
+  and   "equivp aty"
+  and   "equivp atrm"
+sorry
+
+quotient_type KIND = kind / akind
+  by (rule alpha_equivps)
+
+quotient_type 
+    TY = ty / aty and   
+    TRM = trm / atrm
+  by (auto intro: alpha_equivps)
+
+quotient_definition
+   "TYP :: KIND"
+is
+  "Type"
+
+quotient_definition
+   "KPI :: TY \<Rightarrow> name \<Rightarrow> KIND \<Rightarrow> KIND"
+is
+  "KPi"
+
+quotient_definition
+   "TCONST :: ident \<Rightarrow> TY"
+is
+  "TConst"
+
+quotient_definition
+   "TAPP :: TY \<Rightarrow> TRM \<Rightarrow> TY"
+is
+  "TApp"
+
+quotient_definition
+   "TPI :: TY \<Rightarrow> name \<Rightarrow> TY \<Rightarrow> TY"
+is
+  "TPi"
+
+(* FIXME: does not work with CONST *)
+quotient_definition
+   "CONS :: ident \<Rightarrow> TRM"
+is
+  "Const"
+
+quotient_definition
+   "VAR :: name \<Rightarrow> TRM"
+is
+  "Var"
+
+quotient_definition
+   "APP :: TRM \<Rightarrow> TRM \<Rightarrow> TRM"
+is
+  "App"
+
+quotient_definition
+   "LAM :: TY \<Rightarrow> name \<Rightarrow> TRM \<Rightarrow> TRM"
+is
+  "Lam"
+
+thm TYP_def
+thm KPI_def
+thm TCONST_def
+thm TAPP_def
+thm TPI_def
+thm VAR_def
+thm CONS_def
+thm APP_def
+thm LAM_def
+
+(* FIXME: print out a warning if the type contains a liftet type, like kind \<Rightarrow> name set *)
+quotient_definition
+   "FV_kind :: KIND \<Rightarrow> name set"
+is
+  "fv_kind"
+
+quotient_definition
+   "FV_ty :: TY \<Rightarrow> name set"
+is
+  "fv_ty"
+
+quotient_definition
+   "FV_trm :: TRM \<Rightarrow> name set"
+is
+  "fv_trm"
+
+thm FV_kind_def
+thm FV_ty_def
+thm FV_trm_def
+
+(* FIXME: does not work yet *)
+overloading
+    perm_kind \<equiv> "perm :: 'x prm \<Rightarrow> KIND \<Rightarrow> KIND"   (unchecked)
+    perm_ty   \<equiv> "perm :: 'x prm \<Rightarrow> TY \<Rightarrow> TY"       (unchecked)
+    perm_trm  \<equiv> "perm :: 'x prm \<Rightarrow> TRM \<Rightarrow> TRM"     (unchecked) 
+begin
+
+quotient_definition
+   "perm_kind :: 'x prm \<Rightarrow> KIND \<Rightarrow> KIND"
+is
+  "(perm::'x prm \<Rightarrow> kind \<Rightarrow> kind)"
+
+quotient_definition
+   "perm_ty :: 'x prm \<Rightarrow> TY \<Rightarrow> TY"
+is
+  "(perm::'x prm \<Rightarrow> ty \<Rightarrow> ty)"
+
+quotient_definition
+   "perm_trm :: 'x prm \<Rightarrow> TRM \<Rightarrow> TRM"
+is
+  "(perm::'x prm \<Rightarrow> trm \<Rightarrow> trm)"
+
+end
+
+(* TODO/FIXME: Think whether these RSP theorems are true. *)
+lemma kpi_rsp[quot_respect]: 
+  "(aty ===> op = ===> akind ===> akind) KPi KPi" sorry
+lemma tconst_rsp[quot_respect]: 
+  "(op = ===> aty) TConst TConst" sorry
+lemma tapp_rsp[quot_respect]: 
+  "(aty ===> atrm ===> aty) TApp TApp" sorry
+lemma tpi_rsp[quot_respect]: 
+  "(aty ===> op = ===> aty ===> aty) TPi TPi" sorry
+lemma var_rsp[quot_respect]: 
+  "(op = ===> atrm) Var Var" sorry
+lemma app_rsp[quot_respect]: 
+  "(atrm ===> atrm ===> atrm) App App" sorry
+lemma const_rsp[quot_respect]: 
+  "(op = ===> atrm) Const Const" sorry
+lemma lam_rsp[quot_respect]: 
+  "(aty ===> op = ===> atrm ===> atrm) Lam Lam" sorry
+
+lemma perm_kind_rsp[quot_respect]: 
+  "(op = ===> akind ===> akind) op \<bullet> op \<bullet>" sorry
+lemma perm_ty_rsp[quot_respect]: 
+  "(op = ===> aty ===> aty) op \<bullet> op \<bullet>" sorry
+lemma perm_trm_rsp[quot_respect]: 
+  "(op = ===> atrm ===> atrm) op \<bullet> op \<bullet>" sorry
+
+lemma fv_ty_rsp[quot_respect]: 
+  "(aty ===> op =) fv_ty fv_ty" sorry
+lemma fv_kind_rsp[quot_respect]: 
+  "(akind ===> op =) fv_kind fv_kind" sorry
+lemma fv_trm_rsp[quot_respect]: 
+  "(atrm ===> op =) fv_trm fv_trm" sorry
+
+
+thm akind_aty_atrm.induct
+thm kind_ty_trm.induct
+
+
+lemma 
+  assumes a0:
+  "P1 TYP TYP"
+  and a1: 
+  "\<And>A A' K K' x. \<lbrakk>(A::TY) = A'; P2 A A'; (K::KIND) = K'; P1 K K'\<rbrakk> 
+  \<Longrightarrow> P1 (KPI A x K) (KPI A' x K')"
+  and a2:    
+  "\<And>A A' K K' x x'. \<lbrakk>(A ::TY) = A'; P2 A A'; (K :: KIND) = ([(x, x')] \<bullet> K'); P1 K ([(x, x')] \<bullet> K'); 
+    x \<notin> FV_ty A'; x \<notin> FV_kind K' - {x'}\<rbrakk> \<Longrightarrow> P1 (KPI A x K) (KPI A' x' K')"
+  and a3: 
+  "\<And>i j. i = j \<Longrightarrow> P2 (TCONST i) (TCONST j)"
+  and a4:
+  "\<And>A A' M M'. \<lbrakk>(A ::TY) = A'; P2 A A'; (M :: TRM) = M'; P3 M M'\<rbrakk> \<Longrightarrow> P2 (TAPP A M) (TAPP A' M')"
+  and a5:
+  "\<And>A A' B B' x. \<lbrakk>(A ::TY) = A'; P2 A A'; (B ::TY) = B'; P2 B B'\<rbrakk> \<Longrightarrow> P2 (TPI A x B) (TPI A' x B')"
+  and a6:
+  "\<And>A A' B x x' B'. \<lbrakk>(A ::TY) = A'; P2 A A'; (B ::TY) = ([(x, x')] \<bullet> B'); P2 B ([(x, x')] \<bullet> B'); 
+  x \<notin> FV_ty B'; x \<notin> FV_ty B' - {x'}\<rbrakk> \<Longrightarrow> P2 (TPI A x B) (TPI A' x' B')"
+  and a7:
+  "\<And>i j m. i = j \<Longrightarrow> P3 (CONS i) (m (CONS j))"
+  and a8:
+  "\<And>x y m. x = y \<Longrightarrow> P3 (VAR x) (m (VAR y))"
+  and a9:
+  "\<And>M m M' N N'. \<lbrakk>(M :: TRM) = m M'; P3 M (m M'); (N :: TRM) = N'; P3 N N'\<rbrakk> \<Longrightarrow> P3 (APP M N) (APP M' N')"
+  and a10: 
+  "\<And>A A' M M' x. \<lbrakk>(A ::TY) = A'; P2 A A'; (M :: TRM) = M'; P3 M M'\<rbrakk> \<Longrightarrow> P3 (LAM A x M) (LAM A' x M')"
+  and a11:
+  "\<And>A A' M x x' M' B'. \<lbrakk>(A ::TY) = A'; P2 A A'; (M :: TRM) = ([(x, x')] \<bullet> M'); P3 M ([(x, x')] \<bullet> M'); 
+  x \<notin> FV_ty B'; x \<notin> FV_trm M' - {x'}\<rbrakk> \<Longrightarrow> P3 (LAM A x M) (LAM A' x' M')"
+  shows "((x1 :: KIND) = x2 \<longrightarrow> P1 x1 x2) \<and>
+         ((x3 ::TY) = x4 \<longrightarrow> P2 x3 x4) \<and> 
+         ((x5 :: TRM) = x6 \<longrightarrow> P3 x5 x6)"
+using a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10 a11
+apply(lifting akind_aty_atrm.induct)
+(*
+Profiling:
+ML_prf {* fun ith i =  (#concl (fst (Subgoal.focus @{context} i (#goal (Isar.goal ()))))) *}
+ML_prf {* profile 2 Seq.list_of ((clean_tac @{context} quot defs 1) (ith 3)) *}
+ML_prf {* profile 2 Seq.list_of ((regularize_tac @{context} @{thms alpha_equivps} 1) (ith 1)) *}
+ML_prf {* PolyML.profiling 1 *}
+ML_prf {* profile 2 Seq.list_of ((all_inj_repabs_tac @{context} quot rel_refl trans2 1) (#goal (Isar.goal ()))) *}
+*)
+ done
+
+(* Does not work:
+lemma
+  assumes a0: "P1 TYP"
+  and     a1: "\<And>ty name kind. \<lbrakk>P2 ty; P1 kind\<rbrakk> \<Longrightarrow> P1 (KPI ty name kind)"
+  and     a2: "\<And>id. P2 (TCONST id)"
+  and     a3: "\<And>ty trm. \<lbrakk>P2 ty; P3 trm\<rbrakk> \<Longrightarrow> P2 (TAPP ty trm)"
+  and     a4: "\<And>ty1 name ty2. \<lbrakk>P2 ty1; P2 ty2\<rbrakk> \<Longrightarrow> P2 (TPI ty1 name ty2)"
+  and     a5: "\<And>id. P3 (CONS id)"
+  and     a6: "\<And>name. P3 (VAR name)"
+  and     a7: "\<And>trm1 trm2. \<lbrakk>P3 trm1; P3 trm2\<rbrakk> \<Longrightarrow> P3 (APP trm1 trm2)"
+  and     a8: "\<And>ty name trm. \<lbrakk>P2 ty; P3 trm\<rbrakk> \<Longrightarrow> P3 (LAM ty name trm)"
+  shows "P1 mkind \<and> P2 mty \<and> P3 mtrm"
+using a0 a1 a2 a3 a4 a5 a6 a7 a8
+*)
+
+
+lemma "\<lbrakk>P TYP;
+  \<And>ty name kind. \<lbrakk>Q ty; P kind\<rbrakk> \<Longrightarrow> P (KPI ty name kind);
+  \<And>id. Q (TCONST id);
+  \<And>ty trm. \<lbrakk>Q ty; R trm\<rbrakk> \<Longrightarrow> Q (TAPP ty trm);
+  \<And>ty1 name ty2. \<lbrakk>Q ty1; Q ty2\<rbrakk> \<Longrightarrow> Q (TPI ty1 name ty2);
+  \<And>id. R (CONS id); \<And>name. R (VAR name);
+  \<And>trm1 trm2. \<lbrakk>R trm1; R trm2\<rbrakk> \<Longrightarrow> R (APP trm1 trm2);
+  \<And>ty name trm. \<lbrakk>Q ty; R trm\<rbrakk> \<Longrightarrow> R (LAM ty name trm)\<rbrakk>
+  \<Longrightarrow> P mkind \<and> Q mty \<and> R mtrm"
+apply(lifting kind_ty_trm.induct)
+done
+
+end
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/LamEx.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,636 @@
+theory LamEx
+imports Nominal "../Quotient" "../Quotient_List"
+begin
+
+atom_decl name
+
+datatype rlam =
+  rVar "name"
+| rApp "rlam" "rlam"
+| rLam "name" "rlam"
+
+fun
+  rfv :: "rlam \<Rightarrow> name set"
+where
+  rfv_var: "rfv (rVar a) = {a}"
+| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
+| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
+
+overloading
+  perm_rlam \<equiv> "perm :: 'x prm \<Rightarrow> rlam \<Rightarrow> rlam"   (unchecked)
+begin
+
+fun
+  perm_rlam
+where
+  "perm_rlam pi (rVar a) = rVar (pi \<bullet> a)"
+| "perm_rlam pi (rApp t1 t2) = rApp (perm_rlam pi t1) (perm_rlam pi t2)"
+| "perm_rlam pi (rLam a t) = rLam (pi \<bullet> a) (perm_rlam pi t)"
+
+end
+
+declare perm_rlam.simps[eqvt]
+
+instance rlam::pt_name
+  apply(default)
+  apply(induct_tac [!] x rule: rlam.induct)
+  apply(simp_all add: pt_name2 pt_name3)
+  done
+
+instance rlam::fs_name
+  apply(default)
+  apply(induct_tac [!] x rule: rlam.induct)
+  apply(simp add: supp_def)
+  apply(fold supp_def)
+  apply(simp add: supp_atm)
+  apply(simp add: supp_def Collect_imp_eq Collect_neg_eq)
+  apply(simp add: supp_def)
+  apply(simp add: supp_def Collect_imp_eq Collect_neg_eq[symmetric])
+  apply(fold supp_def)
+  apply(simp add: supp_atm)
+  done
+
+declare set_diff_eqvt[eqvt]
+
+lemma rfv_eqvt[eqvt]:
+  fixes pi::"name prm"
+  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
+apply(induct t)
+apply(simp_all)
+apply(simp add: perm_set_eq)
+apply(simp add: union_eqvt)
+apply(simp add: set_diff_eqvt)
+apply(simp add: perm_set_eq)
+done
+
+inductive
+    alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
+| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
+| a3: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s \<and> (pi \<bullet> a) = b)
+       \<Longrightarrow> rLam a t \<approx> rLam b s"
+
+
+(* should be automatic with new version of eqvt-machinery *)
+lemma alpha_eqvt:
+  fixes pi::"name prm"
+  shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)"
+apply(induct rule: alpha.induct)
+apply(simp add: a1)
+apply(simp add: a2)
+apply(simp)
+apply(rule a3)
+apply(erule conjE)
+apply(erule exE)
+apply(erule conjE)
+apply(rule_tac x="pi \<bullet> pia" in exI)
+apply(rule conjI)
+apply(rule_tac pi1="rev pi" in perm_bij[THEN iffD1])
+apply(perm_simp add: eqvts)
+apply(rule conjI)
+apply(rule_tac pi1="rev pi" in pt_fresh_star_bij(1)[OF pt_name_inst at_name_inst, THEN iffD1])
+apply(perm_simp add: eqvts)
+apply(rule conjI)
+apply(subst perm_compose[symmetric])
+apply(simp)
+apply(subst perm_compose[symmetric])
+apply(simp)
+done
+
+lemma alpha_refl:
+  shows "t \<approx> t"
+apply(induct t rule: rlam.induct)
+apply(simp add: a1)
+apply(simp add: a2)
+apply(rule a3)
+apply(rule_tac x="[]" in exI)
+apply(simp_all add: fresh_star_def fresh_list_nil)
+done
+
+lemma alpha_sym:
+  shows "t \<approx> s \<Longrightarrow> s \<approx> t"
+apply(induct rule: alpha.induct)
+apply(simp add: a1)
+apply(simp add: a2)
+apply(rule a3)
+apply(erule exE)
+apply(rule_tac x="rev pi" in exI)
+apply(simp)
+apply(simp add: fresh_star_def fresh_list_rev)
+apply(rule conjI)
+apply(erule conjE)+
+apply(rotate_tac 3)
+apply(drule_tac pi="rev pi" in alpha_eqvt)
+apply(perm_simp)
+apply(rule pt_bij2[OF pt_name_inst at_name_inst])
+apply(simp)
+done
+
+lemma alpha_trans:
+  shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
+apply(induct arbitrary: t3 rule: alpha.induct)
+apply(erule alpha.cases)
+apply(simp_all)
+apply(simp add: a1)
+apply(rotate_tac 4)
+apply(erule alpha.cases)
+apply(simp_all)
+apply(simp add: a2)
+apply(rotate_tac 1)
+apply(erule alpha.cases)
+apply(simp_all)
+apply(erule conjE)+
+apply(erule exE)+
+apply(erule conjE)+
+apply(rule a3)
+apply(rule_tac x="pia @ pi" in exI)
+apply(simp add: fresh_star_def fresh_list_append)
+apply(simp add: pt_name2)
+apply(drule_tac x="rev pia \<bullet> sa" in spec)
+apply(drule mp)
+apply(rotate_tac 8)
+apply(drule_tac pi="rev pia" in alpha_eqvt)
+apply(perm_simp)
+apply(rotate_tac 11)
+apply(drule_tac pi="pia" in alpha_eqvt)
+apply(perm_simp)
+done
+
+lemma alpha_equivp:
+  shows "equivp alpha"
+apply(rule equivpI)
+unfolding reflp_def symp_def transp_def
+apply(auto intro: alpha_refl alpha_sym alpha_trans)
+done
+
+lemma alpha_rfv:
+  shows "t \<approx> s \<Longrightarrow> rfv t = rfv s"
+apply(induct rule: alpha.induct)
+apply(simp)
+apply(simp)
+apply(simp)
+done
+
+quotient_type lam = rlam / alpha
+  by (rule alpha_equivp)
+
+
+quotient_definition
+  "Var :: name \<Rightarrow> lam"
+is
+  "rVar"
+
+quotient_definition
+   "App :: lam \<Rightarrow> lam \<Rightarrow> lam"
+is
+  "rApp"
+
+quotient_definition
+  "Lam :: name \<Rightarrow> lam \<Rightarrow> lam"
+is
+  "rLam"
+
+quotient_definition
+  "fv :: lam \<Rightarrow> name set"
+is
+  "rfv"
+
+(* definition of overloaded permutation function *)
+(* for the lifted type lam                       *)
+overloading
+  perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
+begin
+
+quotient_definition
+  "perm_lam :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"
+is
+  "perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam"
+
+end
+
+lemma perm_rsp[quot_respect]:
+  "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
+  apply(auto)
+  (* this is propably true if some type conditions are imposed ;o) *)
+  sorry
+
+lemma fresh_rsp:
+  "(op = ===> alpha ===> op =) fresh fresh"
+  apply(auto)
+  (* this is probably only true if some type conditions are imposed *)
+  sorry
+
+lemma rVar_rsp[quot_respect]:
+  "(op = ===> alpha) rVar rVar"
+  by (auto intro: a1)
+
+lemma rApp_rsp[quot_respect]: "(alpha ===> alpha ===> alpha) rApp rApp"
+  by (auto intro: a2)
+
+lemma rLam_rsp[quot_respect]: "(op = ===> alpha ===> alpha) rLam rLam"
+  apply(auto)
+  apply(rule a3)
+  apply(rule_tac x="[]" in exI)
+  unfolding fresh_star_def
+  apply(simp add: fresh_list_nil)
+  apply(simp add: alpha_rfv)
+  done
+
+lemma rfv_rsp[quot_respect]: 
+  "(alpha ===> op =) rfv rfv"
+apply(simp add: alpha_rfv)
+done
+
+section {* lifted theorems *}
+
+lemma lam_induct:
+  "\<lbrakk>\<And>name. P (Var name);
+    \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
+    \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> 
+    \<Longrightarrow> P lam"
+  by (lifting rlam.induct)
+
+ML {* show_all_types := true *}
+
+lemma perm_lam [simp]:
+  fixes pi::"'a prm"
+  shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
+  and   "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)"
+  and   "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)"
+apply(lifting perm_rlam.simps)
+ML_prf {*
+  List.last (map (symmetric o #def) (Quotient_Info.qconsts_dest @{context}));
+  List.last (map (Thm.varifyT o symmetric o #def) (Quotient_Info.qconsts_dest @{context}))
+*}
+done
+
+instance lam::pt_name
+apply(default)
+apply(induct_tac [!] x rule: lam_induct)
+apply(simp_all add: pt_name2 pt_name3)
+done
+
+lemma fv_lam [simp]: 
+  shows "fv (Var a) = {a}"
+  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
+  and   "fv (Lam a t) = fv t - {a}"
+apply(lifting rfv_var rfv_app rfv_lam)
+done
+
+
+lemma a1: 
+  "a = b \<Longrightarrow> Var a = Var b"
+  by  (lifting a1)
+
+lemma a2: 
+  "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
+  by  (lifting a2)
+
+lemma a3: 
+  "\<lbrakk>\<exists>pi::name prm. (fv t - {a} = fv s - {b} \<and> (fv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) = s \<and> (pi \<bullet> a) = b)\<rbrakk> 
+   \<Longrightarrow> Lam a t = Lam b s"
+  by  (lifting a3)
+
+lemma alpha_cases: 
+  "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
+    \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
+    \<And>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s; 
+         \<exists>pi::name prm. fv t - {a} = fv s - {b} \<and> (fv t - {a}) \<sharp>* pi \<and> (pi \<bullet> t) = s \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> P\<rbrakk>
+    \<Longrightarrow> P"
+  by (lifting alpha.cases)
+
+lemma alpha_induct: 
+  "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b);
+    \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
+     \<And>t a s b.
+        \<lbrakk>\<exists>pi::name prm. fv t - {a} = fv s - {b} \<and>
+         (fv t - {a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s) \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
+    \<Longrightarrow> qxb qx qxa"
+  by (lifting alpha.induct)
+
+lemma lam_inject [simp]: 
+  shows "(Var a = Var b) = (a = b)"
+  and   "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)"
+apply(lifting rlam.inject(1) rlam.inject(2))
+apply(auto)
+apply(drule alpha.cases)
+apply(simp_all)
+apply(simp add: alpha.a1)
+apply(drule alpha.cases)
+apply(simp_all)
+apply(drule alpha.cases)
+apply(simp_all)
+apply(rule alpha.a2)
+apply(simp_all)
+done
+
+lemma rlam_distinct:
+  shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')"
+  and   "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)"
+  and   "\<not>(rVar nam \<approx> rLam nam' rlam')"
+  and   "\<not>(rLam nam' rlam' \<approx> rVar nam)"
+  and   "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')"
+  and   "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)"
+apply auto
+apply(erule alpha.cases)
+apply simp_all
+apply(erule alpha.cases)
+apply simp_all
+apply(erule alpha.cases)
+apply simp_all
+apply(erule alpha.cases)
+apply simp_all
+apply(erule alpha.cases)
+apply simp_all
+apply(erule alpha.cases)
+apply simp_all
+done
+
+lemma lam_distinct[simp]:
+  shows "Var nam \<noteq> App lam1' lam2'"
+  and   "App lam1' lam2' \<noteq> Var nam"
+  and   "Var nam \<noteq> Lam nam' lam'"
+  and   "Lam nam' lam' \<noteq> Var nam"
+  and   "App lam1 lam2 \<noteq> Lam nam' lam'"
+  and   "Lam nam' lam' \<noteq> App lam1 lam2"
+apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6))
+done
+
+lemma var_supp1:
+  shows "(supp (Var a)) = ((supp a)::name set)"
+  by (simp add: supp_def)
+
+lemma var_supp:
+  shows "(supp (Var a)) = {a::name}"
+  using var_supp1 by (simp add: supp_atm)
+
+lemma app_supp:
+  shows "supp (App t1 t2) = (supp t1) \<union> ((supp t2)::name set)"
+apply(simp only: perm_lam supp_def lam_inject)
+apply(simp add: Collect_imp_eq Collect_neg_eq)
+done
+
+lemma lam_supp:
+  shows "supp (Lam x t) = ((supp ([x].t))::name set)"
+apply(simp add: supp_def)
+apply(simp add: abs_perm)
+sorry
+
+
+instance lam::fs_name
+apply(default)
+apply(induct_tac x rule: lam_induct)
+apply(simp add: var_supp)
+apply(simp add: app_supp)
+apply(simp add: lam_supp abs_supp)
+done
+
+lemma fresh_lam:
+  "(a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> a \<sharp> t)"
+apply(simp add: fresh_def)
+apply(simp add: lam_supp abs_supp)
+apply(auto)
+done
+
+lemma lam_induct_strong:
+  fixes a::"'a::fs_name"
+  assumes a1: "\<And>name b. P b (Var name)"
+  and     a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)"
+  and     a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; name \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
+  shows "P a lam"
+proof -
+  have "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" 
+  proof (induct lam rule: lam_induct)
+    case (1 name pi)
+    show "P a (pi \<bullet> Var name)"
+      apply (simp)
+      apply (rule a1)
+      done
+  next
+    case (2 lam1 lam2 pi)
+    have b1: "\<And>(pi::name prm) a. P a (pi \<bullet> lam1)" by fact
+    have b2: "\<And>(pi::name prm) a. P a (pi \<bullet> lam2)" by fact
+    show "P a (pi \<bullet> App lam1 lam2)"
+      apply (simp)
+      apply (rule a2)
+      apply (rule b1)
+      apply (rule b2)
+      done
+  next
+    case (3 name lam pi a)
+    have b: "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" by fact
+    obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
+      apply(rule exists_fresh[of "(a, pi\<bullet>name, pi\<bullet>lam)"])
+      apply(simp_all add: fs_name1)
+      done
+    from b fr have p: "P a (Lam c (([(c, pi\<bullet>name)]@pi)\<bullet>lam))" 
+      apply -
+      apply(rule a3)
+      apply(blast)
+      apply(simp)
+      done
+    have eq: "[(c, pi\<bullet>name)] \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
+      apply(rule perm_fresh_fresh)
+      using fr
+      apply(simp add: fresh_lam)
+      apply(simp add: fresh_lam)
+      done
+    show "P a (pi \<bullet> Lam name lam)" 
+      apply (simp)
+      apply(subst eq[symmetric])
+      using p
+      apply(simp only: perm_lam pt_name2 swap_simps)
+      done
+  qed
+  then have "P a (([]::name prm) \<bullet> lam)" by blast
+  then show "P a lam" by simp 
+qed
+
+
+lemma var_fresh:
+  fixes a::"name"
+  shows "(a \<sharp> (Var b)) = (a \<sharp> b)"
+  apply(simp add: fresh_def)
+  apply(simp add: var_supp1)
+  done
+
+(* lemma hom_reg: *)
+
+lemma rlam_rec_eqvt:
+  fixes pi::"name prm"
+  and   f1::"name \<Rightarrow> ('a::pt_name)"
+  shows "(pi\<bullet>rlam_rec f1 f2 f3 t) = rlam_rec (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3) (pi\<bullet>t)"
+apply(induct t)
+apply(simp_all)
+apply(simp add: perm_fun_def)
+apply(perm_simp)
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+back
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+apply(simp)
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+back
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
+apply(simp)
+done
+ 
+
+lemma rlam_rec_respects:
+  assumes f1: "f_var \<in> Respects (op= ===> op=)"
+  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
+  and     f3: "f_lam \<in> Respects (op= ===> alpha ===> op= ===> op=)"
+  shows "rlam_rec f_var f_app f_lam \<in> Respects (alpha ===> op =)"
+apply(simp add: mem_def)
+apply(simp add: Respects_def)
+apply(rule allI)
+apply(rule allI)
+apply(rule impI)
+apply(erule alpha.induct)
+apply(simp)
+apply(simp)
+using f2
+apply(simp add: mem_def)
+apply(simp add: Respects_def)
+using f3[simplified mem_def Respects_def]
+apply(simp)
+apply(case_tac "a=b")
+apply(clarify)
+apply(simp)
+(* probably true *)
+sorry
+
+function
+  term1_hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
+                (rlam \<Rightarrow> rlam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
+                ((name \<Rightarrow> rlam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> rlam \<Rightarrow> 'a"
+where
+  "term1_hom var app abs' (rVar x) = (var x)"
+| "term1_hom var app abs' (rApp t u) =
+     app t u (term1_hom var app abs' t) (term1_hom var app abs' u)"
+| "term1_hom var app abs' (rLam x u) =
+     abs' (\<lambda>y. [(x, y)] \<bullet> u) (\<lambda>y. term1_hom var app abs' ([(x, y)] \<bullet> u))"
+apply(pat_completeness)
+apply(auto)
+done
+
+lemma pi_size:
+  fixes pi::"name prm"
+  and   t::"rlam"
+  shows "size (pi \<bullet> t) = size t"
+apply(induct t)
+apply(auto)
+done
+
+termination term1_hom
+  apply(relation "measure (\<lambda>(f1, f2, f3, t). size t)")
+apply(auto simp add: pi_size)
+done
+
+lemma lam_exhaust:
+  "\<lbrakk>\<And>name. y = Var name \<Longrightarrow> P; \<And>rlam1 rlam2. y = App rlam1 rlam2 \<Longrightarrow> P; \<And>name rlam. y = Lam name rlam \<Longrightarrow> P\<rbrakk>
+    \<Longrightarrow> P"
+apply(lifting rlam.exhaust)
+done
+
+(* THIS IS NOT TRUE, but it lets prove the existence of the hom function *)
+lemma lam_inject':
+  "(Lam a x = Lam b y) = ((\<lambda>c. [(a, c)] \<bullet> x) = (\<lambda>c. [(b, c)] \<bullet> y))"
+sorry
+
+function
+  hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
+                (lam \<Rightarrow> lam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
+                ((name \<Rightarrow> lam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> lam \<Rightarrow> 'a"
+where
+  "hom f_var f_app f_lam (Var x) = f_var x"
+| "hom f_var f_app f_lam (App l r) = f_app l r (hom f_var f_app f_lam l) (hom f_var f_app f_lam r)"
+| "hom f_var f_app f_lam (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom f_var f_app f_lam ([(a,b)] \<bullet> x))"
+defer
+apply(simp_all add: lam_inject') (* inject, distinct *)
+apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+apply(rule refl)
+apply(rule ext)
+apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
+apply simp_all
+apply(erule conjE)+
+apply(rule_tac x="b" in cong)
+apply simp_all
+apply auto
+apply(rule_tac y="b" in lam_exhaust)
+apply simp_all
+apply auto
+apply meson
+apply(simp_all add: lam_inject')
+apply metis
+done
+
+termination hom
+  apply -
+(*
+ML_prf {* Size.size_thms @{theory} "LamEx.lam" *}
+*)
+sorry
+
+thm hom.simps
+
+lemma term1_hom_rsp:
+  "\<lbrakk>(alpha ===> alpha ===> op =) f_app f_app; ((op = ===> alpha) ===> op =) f_lam f_lam\<rbrakk>
+       \<Longrightarrow> (alpha ===> op =) (term1_hom f_var f_app f_lam) (term1_hom f_var f_app f_lam)"
+apply(simp)
+apply(rule allI)+
+apply(rule impI)
+apply(erule alpha.induct)
+apply(auto)[1]
+apply(auto)[1]
+apply(simp)
+apply(erule conjE)+
+apply(erule exE)+
+apply(erule conjE)+
+apply(clarify)
+sorry
+
+lemma hom: "
+\<forall>f_var. \<forall>f_app \<in> Respects(alpha ===> alpha ===> op =).
+\<forall>f_lam \<in> Respects((op = ===> alpha) ===> op =).
+\<exists>hom\<in>Respects (alpha ===> op =). 
+    ((\<forall>x. hom (rVar x) = f_var x) \<and>
+     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
+     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
+apply(rule allI)
+apply(rule ballI)+
+apply(rule_tac x="term1_hom f_var f_app f_lam" in bexI)
+apply(simp_all)
+apply(simp only: in_respects)
+apply(rule term1_hom_rsp)
+apply(assumption)+
+done
+
+lemma hom':
+"\<exists>hom.
+  ((\<forall>x. hom (Var x) = f_var x) \<and>
+   (\<forall>l r. hom (App l r) = f_app l r (hom l) (hom r)) \<and>
+   (\<forall>x a. hom (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
+apply (lifting hom)
+done
+
+(* test test
+lemma raw_hom_correct: 
+  assumes f1: "f_var \<in> Respects (op= ===> op=)"
+  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
+  and     f3: "f_lam \<in> Respects ((op= ===> alpha) ===> (op= ===> op=) ===> op=)"
+  shows "\<exists>!hom\<in>Respects (alpha ===> op =). 
+    ((\<forall>x. hom (rVar x) = f_var x) \<and>
+     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
+     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
+unfolding Bex1_def
+apply(rule ex1I)
+sorry
+*)
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/LarryDatatype.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,394 @@
+theory LarryDatatype
+imports Main "../Quotient" "../Quotient_Syntax"
+begin
+
+subsection{*Defining the Free Algebra*}
+
+datatype
+  freemsg = NONCE  nat
+        | MPAIR  freemsg freemsg
+        | CRYPT  nat freemsg  
+        | DECRYPT  nat freemsg
+
+inductive 
+  msgrel::"freemsg \<Rightarrow> freemsg \<Rightarrow> bool" (infixl "\<sim>" 50)
+where 
+  CD:    "CRYPT K (DECRYPT K X) \<sim> X"
+| DC:    "DECRYPT K (CRYPT K X) \<sim> X"
+| NONCE: "NONCE N \<sim> NONCE N"
+| MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
+| CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
+| DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
+| SYM:   "X \<sim> Y \<Longrightarrow> Y \<sim> X"
+| TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
+
+lemmas msgrel.intros[intro]
+
+text{*Proving that it is an equivalence relation*}
+
+lemma msgrel_refl: "X \<sim> X"
+by (induct X, (blast intro: msgrel.intros)+)
+
+theorem equiv_msgrel: "equivp msgrel"
+proof (rule equivpI)
+  show "reflp msgrel" by (simp add: reflp_def msgrel_refl)
+  show "symp msgrel" by (simp add: symp_def, blast intro: msgrel.SYM)
+  show "transp msgrel" by (simp add: transp_def, blast intro: msgrel.TRANS)
+qed
+
+subsection{*Some Functions on the Free Algebra*}
+
+subsubsection{*The Set of Nonces*}
+
+fun
+  freenonces :: "freemsg \<Rightarrow> nat set"
+where
+  "freenonces (NONCE N) = {N}"
+| "freenonces (MPAIR X Y) = freenonces X \<union> freenonces Y"
+| "freenonces (CRYPT K X) = freenonces X"
+| "freenonces (DECRYPT K X) = freenonces X"
+
+theorem msgrel_imp_eq_freenonces: 
+  assumes a: "U \<sim> V"
+  shows "freenonces U = freenonces V"
+  using a by (induct) (auto) 
+
+subsubsection{*The Left Projection*}
+
+text{*A function to return the left part of the top pair in a message.  It will
+be lifted to the initial algrebra, to serve as an example of that process.*}
+fun
+  freeleft :: "freemsg \<Rightarrow> freemsg"
+where
+  "freeleft (NONCE N) = NONCE N"
+| "freeleft (MPAIR X Y) = X"
+| "freeleft (CRYPT K X) = freeleft X"
+| "freeleft (DECRYPT K X) = freeleft X"
+
+text{*This theorem lets us prove that the left function respects the
+equivalence relation.  It also helps us prove that MPair
+  (the abstract constructor) is injective*}
+lemma msgrel_imp_eqv_freeleft_aux:
+  shows "freeleft U \<sim> freeleft U"
+  by (induct rule: freeleft.induct) (auto)
+
+theorem msgrel_imp_eqv_freeleft:
+  assumes a: "U \<sim> V" 
+  shows "freeleft U \<sim> freeleft V"
+  using a
+  by (induct) (auto intro: msgrel_imp_eqv_freeleft_aux)
+
+subsubsection{*The Right Projection*}
+
+text{*A function to return the right part of the top pair in a message.*}
+fun
+  freeright :: "freemsg \<Rightarrow> freemsg"
+where
+  "freeright (NONCE N) = NONCE N"
+| "freeright (MPAIR X Y) = Y"
+| "freeright (CRYPT K X) = freeright X"
+| "freeright (DECRYPT K X) = freeright X"
+
+text{*This theorem lets us prove that the right function respects the
+equivalence relation.  It also helps us prove that MPair
+  (the abstract constructor) is injective*}
+lemma msgrel_imp_eqv_freeright_aux:
+  shows "freeright U \<sim> freeright U"
+  by (induct rule: freeright.induct) (auto)
+
+theorem msgrel_imp_eqv_freeright:
+  assumes a: "U \<sim> V" 
+  shows "freeright U \<sim> freeright V"
+  using a
+  by (induct) (auto intro: msgrel_imp_eqv_freeright_aux)
+
+subsubsection{*The Discriminator for Constructors*}
+
+text{*A function to distinguish nonces, mpairs and encryptions*}
+fun 
+  freediscrim :: "freemsg \<Rightarrow> int"
+where
+   "freediscrim (NONCE N) = 0"
+ | "freediscrim (MPAIR X Y) = 1"
+ | "freediscrim (CRYPT K X) = freediscrim X + 2"
+ | "freediscrim (DECRYPT K X) = freediscrim X - 2"
+
+text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
+theorem msgrel_imp_eq_freediscrim:
+  assumes a: "U \<sim> V"
+  shows "freediscrim U = freediscrim V"
+  using a by (induct) (auto)
+
+subsection{*The Initial Algebra: A Quotiented Message Type*}
+
+quotient_type msg = freemsg / msgrel
+  by (rule equiv_msgrel)
+
+text{*The abstract message constructors*}
+
+quotient_definition
+  "Nonce :: nat \<Rightarrow> msg"
+is
+  "NONCE"
+
+quotient_definition
+  "MPair :: msg \<Rightarrow> msg \<Rightarrow> msg"
+is
+  "MPAIR"
+
+quotient_definition
+  "Crypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
+is
+  "CRYPT"
+
+quotient_definition
+  "Decrypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
+is
+  "DECRYPT"
+
+lemma [quot_respect]:
+  shows "(op = ===> op \<sim> ===> op \<sim>) CRYPT CRYPT"
+by (auto intro: CRYPT)
+
+lemma [quot_respect]:
+  shows "(op = ===> op \<sim> ===> op \<sim>) DECRYPT DECRYPT"
+by (auto intro: DECRYPT)
+
+text{*Establishing these two equations is the point of the whole exercise*}
+theorem CD_eq [simp]: 
+  shows "Crypt K (Decrypt K X) = X"
+  by (lifting CD)
+
+theorem DC_eq [simp]: 
+  shows "Decrypt K (Crypt K X) = X"
+  by (lifting DC)
+
+subsection{*The Abstract Function to Return the Set of Nonces*}
+
+quotient_definition
+   "nonces:: msg \<Rightarrow> nat set"
+is
+  "freenonces"
+
+text{*Now prove the four equations for @{term nonces}*}
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op =) freenonces freenonces"
+  by (simp add: msgrel_imp_eq_freenonces)
+
+lemma [quot_respect]:
+  shows "(op = ===> op \<sim>) NONCE NONCE"
+  by (simp add: NONCE)
+
+lemma nonces_Nonce [simp]: 
+  shows "nonces (Nonce N) = {N}"
+  by (lifting freenonces.simps(1))
+ 
+lemma [quot_respect]:
+  shows " (op \<sim> ===> op \<sim> ===> op \<sim>) MPAIR MPAIR"
+  by (simp add: MPAIR)
+
+lemma nonces_MPair [simp]: 
+  shows "nonces (MPair X Y) = nonces X \<union> nonces Y"
+  by (lifting freenonces.simps(2))
+
+lemma nonces_Crypt [simp]: 
+  shows "nonces (Crypt K X) = nonces X"
+  by (lifting freenonces.simps(3))
+
+lemma nonces_Decrypt [simp]: 
+  shows "nonces (Decrypt K X) = nonces X"
+  by (lifting freenonces.simps(4))
+
+subsection{*The Abstract Function to Return the Left Part*}
+
+quotient_definition
+  "left:: msg \<Rightarrow> msg"
+is
+  "freeleft"
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op \<sim>) freeleft freeleft"
+  by (simp add: msgrel_imp_eqv_freeleft)
+
+lemma left_Nonce [simp]: 
+  shows "left (Nonce N) = Nonce N"
+  by (lifting freeleft.simps(1))
+
+lemma left_MPair [simp]: 
+  shows "left (MPair X Y) = X"
+  by (lifting freeleft.simps(2))
+
+lemma left_Crypt [simp]: 
+  shows "left (Crypt K X) = left X"
+  by (lifting freeleft.simps(3))
+
+lemma left_Decrypt [simp]: 
+  shows "left (Decrypt K X) = left X"
+  by (lifting freeleft.simps(4))
+
+subsection{*The Abstract Function to Return the Right Part*}
+
+quotient_definition
+  "right:: msg \<Rightarrow> msg"
+is
+  "freeright"
+
+text{*Now prove the four equations for @{term right}*}
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op \<sim>) freeright freeright"
+  by (simp add: msgrel_imp_eqv_freeright)
+
+lemma right_Nonce [simp]: 
+  shows "right (Nonce N) = Nonce N"
+  by (lifting freeright.simps(1))
+
+lemma right_MPair [simp]: 
+  shows "right (MPair X Y) = Y"
+  by (lifting freeright.simps(2))
+
+lemma right_Crypt [simp]: 
+  shows "right (Crypt K X) = right X"
+  by (lifting freeright.simps(3))
+
+lemma right_Decrypt [simp]: 
+  shows "right (Decrypt K X) = right X"
+  by (lifting freeright.simps(4))
+
+subsection{*Injectivity Properties of Some Constructors*}
+
+lemma NONCE_imp_eq: 
+  shows "NONCE m \<sim> NONCE n \<Longrightarrow> m = n"
+  by (drule msgrel_imp_eq_freenonces, simp)
+
+text{*Can also be proved using the function @{term nonces}*}
+lemma Nonce_Nonce_eq [iff]: 
+  shows "(Nonce m = Nonce n) = (m = n)"
+proof
+  assume "Nonce m = Nonce n"
+  then show "m = n" by (lifting NONCE_imp_eq)
+next
+  assume "m = n" 
+  then show "Nonce m = Nonce n" by simp
+qed
+
+lemma MPAIR_imp_eqv_left: 
+  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> X \<sim> X'"
+  by (drule msgrel_imp_eqv_freeleft) (simp)
+
+lemma MPair_imp_eq_left: 
+  assumes eq: "MPair X Y = MPair X' Y'" 
+  shows "X = X'"
+  using eq by (lifting MPAIR_imp_eqv_left)
+
+lemma MPAIR_imp_eqv_right: 
+  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> Y \<sim> Y'"
+  by (drule msgrel_imp_eqv_freeright) (simp)
+
+lemma MPair_imp_eq_right: 
+  shows "MPair X Y = MPair X' Y' \<Longrightarrow> Y = Y'"
+  by (lifting  MPAIR_imp_eqv_right)
+
+theorem MPair_MPair_eq [iff]: 
+  shows "(MPair X Y = MPair X' Y') = (X=X' & Y=Y')" 
+  by (blast dest: MPair_imp_eq_left MPair_imp_eq_right)
+
+lemma NONCE_neqv_MPAIR: 
+  shows "\<not>(NONCE m \<sim> MPAIR X Y)"
+  by (auto dest: msgrel_imp_eq_freediscrim)
+
+theorem Nonce_neq_MPair [iff]: 
+  shows "Nonce N \<noteq> MPair X Y"
+  by (lifting NONCE_neqv_MPAIR)
+
+text{*Example suggested by a referee*}
+
+lemma CRYPT_NONCE_neq_NONCE:
+  shows "\<not>(CRYPT K (NONCE M) \<sim> NONCE N)"
+  by (auto dest: msgrel_imp_eq_freediscrim)
+
+theorem Crypt_Nonce_neq_Nonce: 
+  shows "Crypt K (Nonce M) \<noteq> Nonce N"
+  by (lifting CRYPT_NONCE_neq_NONCE)
+
+text{*...and many similar results*}
+lemma CRYPT2_NONCE_neq_NONCE: 
+  shows "\<not>(CRYPT K (CRYPT K' (NONCE M)) \<sim> NONCE N)"
+  by (auto dest: msgrel_imp_eq_freediscrim)  
+
+theorem Crypt2_Nonce_neq_Nonce: 
+  shows "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"
+  by (lifting CRYPT2_NONCE_neq_NONCE) 
+
+theorem Crypt_Crypt_eq [iff]: 
+  shows "(Crypt K X = Crypt K X') = (X=X')" 
+proof
+  assume "Crypt K X = Crypt K X'"
+  hence "Decrypt K (Crypt K X) = Decrypt K (Crypt K X')" by simp
+  thus "X = X'" by simp
+next
+  assume "X = X'"
+  thus "Crypt K X = Crypt K X'" by simp
+qed
+
+theorem Decrypt_Decrypt_eq [iff]: 
+  shows "(Decrypt K X = Decrypt K X') = (X=X')" 
+proof
+  assume "Decrypt K X = Decrypt K X'"
+  hence "Crypt K (Decrypt K X) = Crypt K (Decrypt K X')" by simp
+  thus "X = X'" by simp
+next
+  assume "X = X'"
+  thus "Decrypt K X = Decrypt K X'" by simp
+qed
+
+lemma msg_induct_aux:
+  shows "\<lbrakk>\<And>N. P (Nonce N);
+          \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y);
+          \<And>K X. P X \<Longrightarrow> P (Crypt K X);
+          \<And>K X. P X \<Longrightarrow> P (Decrypt K X)\<rbrakk> \<Longrightarrow> P msg"
+  by (lifting freemsg.induct)
+
+lemma msg_induct [case_names Nonce MPair Crypt Decrypt, cases type: msg]:
+  assumes N: "\<And>N. P (Nonce N)"
+      and M: "\<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y)"
+      and C: "\<And>K X. P X \<Longrightarrow> P (Crypt K X)"
+      and D: "\<And>K X. P X \<Longrightarrow> P (Decrypt K X)"
+  shows "P msg"
+  using N M C D by (rule msg_induct_aux)
+
+subsection{*The Abstract Discriminator*}
+
+text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
+need this function in order to prove discrimination theorems.*}
+
+quotient_definition
+  "discrim:: msg \<Rightarrow> int"
+is
+  "freediscrim"
+
+text{*Now prove the four equations for @{term discrim}*}
+
+lemma [quot_respect]:
+  shows "(op \<sim> ===> op =) freediscrim freediscrim"
+  by (auto simp add: msgrel_imp_eq_freediscrim)
+
+lemma discrim_Nonce [simp]: 
+  shows "discrim (Nonce N) = 0"
+  by (lifting freediscrim.simps(1))
+
+lemma discrim_MPair [simp]: 
+  shows "discrim (MPair X Y) = 1"
+  by (lifting freediscrim.simps(2))
+
+lemma discrim_Crypt [simp]: 
+  shows "discrim (Crypt K X) = discrim X + 2"
+  by (lifting freediscrim.simps(3))
+
+lemma discrim_Decrypt [simp]: 
+  shows "discrim (Decrypt K X) = discrim X - 2"
+  by (lifting freediscrim.simps(4))
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/LarryInt.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,395 @@
+
+header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
+
+theory LarryInt
+imports Nat "../Quotient" "../Quotient_Product"
+begin
+
+fun
+  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" 
+where
+  "intrel (x, y) (u, v) = (x + v = u + y)"
+
+quotient_type int = "nat \<times> nat" / intrel
+  by (auto simp add: equivp_def expand_fun_eq)
+
+instantiation int :: "{zero, one, plus, uminus, minus, times, ord}"
+begin
+
+quotient_definition
+  Zero_int_def: "0::int" is "(0::nat, 0::nat)"
+
+quotient_definition
+  One_int_def: "1::int" is "(1::nat, 0::nat)"
+
+definition
+  "add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
+
+quotient_definition
+  "(op +) :: int \<Rightarrow> int \<Rightarrow> int" 
+is
+  "add_raw"
+
+definition
+  "uminus_raw \<equiv> \<lambda>(x::nat, y::nat). (y, x)"
+
+quotient_definition
+  "uminus :: int \<Rightarrow> int" 
+is
+  "uminus_raw"
+
+fun
+  mult_raw::"nat \<times> nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
+where
+  "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
+
+quotient_definition
+  "(op *) :: int \<Rightarrow> int \<Rightarrow> int" 
+is
+  "mult_raw"
+
+definition
+  "le_raw \<equiv> \<lambda>(x, y) (u, v). (x+v \<le> u+(y::nat))"
+
+quotient_definition 
+  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" 
+is
+  "le_raw"
+
+definition
+  less_int_def: "z < (w::int) \<equiv> (z \<le> w & z \<noteq> w)"
+
+definition
+  diff_int_def:  "z - (w::int) \<equiv> z + (-w)"
+
+instance ..
+
+end
+
+subsection{*Construction of the Integers*}
+
+lemma zminus_zminus_raw:
+  "uminus_raw (uminus_raw z) = z"
+  by (cases z) (simp add: uminus_raw_def)
+
+lemma [quot_respect]:
+  shows "(intrel ===> intrel) uminus_raw uminus_raw"
+  by (simp add: uminus_raw_def)
+  
+lemma zminus_zminus:
+  fixes z::"int"
+  shows "- (- z) = z"
+  by(lifting zminus_zminus_raw)
+
+lemma zminus_0_raw:
+  shows "uminus_raw (0, 0) = (0, 0::nat)"
+  by (simp add: uminus_raw_def)
+
+lemma zminus_0: 
+  shows "- 0 = (0::int)"
+  by (lifting zminus_0_raw)
+
+subsection{*Integer Addition*}
+
+lemma zminus_zadd_distrib_raw:
+  shows "uminus_raw (add_raw z w) = add_raw (uminus_raw z) (uminus_raw w)"
+by (cases z, cases w)
+   (auto simp add: add_raw_def uminus_raw_def)
+
+lemma [quot_respect]:
+  shows "(intrel ===> intrel ===> intrel) add_raw add_raw"
+by (simp add: add_raw_def)
+
+lemma zminus_zadd_distrib: 
+  fixes z w::"int"
+  shows "- (z + w) = (- z) + (- w)"
+  by(lifting zminus_zadd_distrib_raw)
+
+lemma zadd_commute_raw:
+  shows "add_raw z w = add_raw w z"
+by (cases z, cases w)
+   (simp add: add_raw_def)
+
+lemma zadd_commute:
+  fixes z w::"int"
+  shows "(z::int) + w = w + z"
+  by (lifting zadd_commute_raw)
+
+lemma zadd_assoc_raw:
+  shows "add_raw (add_raw z1 z2) z3 = add_raw z1 (add_raw z2 z3)"
+  by (cases z1, cases z2, cases z3) (simp add: add_raw_def)
+
+lemma zadd_assoc: 
+  fixes z1 z2 z3::"int"
+  shows "(z1 + z2) + z3 = z1 + (z2 + z3)"
+  by (lifting zadd_assoc_raw)
+
+lemma zadd_0_raw:
+  shows "add_raw (0, 0) z = z"
+  by (simp add: add_raw_def)
+
+
+text {*also for the instance declaration int :: plus_ac0*}
+lemma zadd_0: 
+  fixes z::"int"
+  shows "0 + z = z"
+  by (lifting zadd_0_raw)
+
+lemma zadd_zminus_inverse_raw:
+  shows "intrel (add_raw (uminus_raw z) z) (0, 0)"
+  by (cases z) (simp add: add_raw_def uminus_raw_def)
+
+lemma zadd_zminus_inverse2: 
+  fixes z::"int"
+  shows "(- z) + z = 0"
+  by (lifting zadd_zminus_inverse_raw)
+
+subsection{*Integer Multiplication*}
+
+lemma zmult_zminus_raw:
+  shows "mult_raw (uminus_raw z) w = uminus_raw (mult_raw z w)"
+apply(cases z, cases w)
+apply(simp add: uminus_raw_def)
+done
+
+lemma mult_raw_fst:
+  assumes a: "intrel x z"
+  shows "intrel (mult_raw x y) (mult_raw z y)"
+using a
+apply(cases x, cases y, cases z)
+apply(auto simp add: mult_raw.simps intrel.simps)
+apply(rename_tac u v w x y z)
+apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
+apply(simp add: mult_ac)
+apply(simp add: add_mult_distrib [symmetric])
+done
+
+lemma mult_raw_snd:
+  assumes a: "intrel x z"
+  shows "intrel (mult_raw y x) (mult_raw y z)"
+using a
+apply(cases x, cases y, cases z)
+apply(auto simp add: mult_raw.simps intrel.simps)
+apply(rename_tac u v w x y z)
+apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
+apply(simp add: mult_ac)
+apply(simp add: add_mult_distrib [symmetric])
+done
+
+lemma [quot_respect]:
+  shows "(intrel ===> intrel ===> intrel) mult_raw mult_raw"
+apply(simp only: fun_rel_def)
+apply(rule allI | rule impI)+
+apply(rule equivp_transp[OF int_equivp])
+apply(rule mult_raw_fst)
+apply(assumption)
+apply(rule mult_raw_snd)
+apply(assumption)
+done
+
+lemma zmult_zminus: 
+  fixes z w::"int"
+  shows "(- z) * w = - (z * w)"
+  by (lifting zmult_zminus_raw)
+
+lemma zmult_commute_raw: 
+  shows "mult_raw z w = mult_raw w z"
+apply(cases z, cases w)
+apply(simp add: add_ac mult_ac)
+done
+
+lemma zmult_commute: 
+  fixes z w::"int"
+  shows "z * w = w * z"
+  by (lifting zmult_commute_raw)
+
+lemma zmult_assoc_raw:
+  shows "mult_raw (mult_raw z1 z2) z3 = mult_raw z1 (mult_raw z2 z3)"
+apply(cases z1, cases z2, cases z3)
+apply(simp add: add_mult_distrib2 mult_ac)
+done
+
+lemma zmult_assoc: 
+  fixes z1 z2 z3::"int"
+  shows "(z1 * z2) * z3 = z1 * (z2 * z3)"
+  by (lifting zmult_assoc_raw)
+
+lemma zadd_mult_distrib_raw:
+  shows "mult_raw (add_raw z1 z2) w = add_raw (mult_raw z1 w) (mult_raw z2 w)"
+apply(cases z1, cases z2, cases w)
+apply(simp add: add_mult_distrib2 mult_ac add_raw_def)
+done
+
+lemma zadd_zmult_distrib: 
+  fixes z1 z2 w::"int"
+  shows "(z1 + z2) * w = (z1 * w) + (z2 * w)"
+  by(lifting zadd_mult_distrib_raw)
+
+lemma zadd_zmult_distrib2: 
+  fixes w z1 z2::"int"
+  shows "w * (z1 + z2) = (w * z1) + (w * z2)"
+  by (simp add: zmult_commute [of w] zadd_zmult_distrib)
+
+lemma zdiff_zmult_distrib: 
+  fixes w z1 z2::"int"
+  shows "(z1 - z2) * w = (z1 * w) - (z2 * w)"
+  by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
+
+lemma zdiff_zmult_distrib2: 
+  fixes w z1 z2::"int"
+  shows "w * (z1 - z2) = (w * z1) - (w * z2)"
+  by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
+
+lemmas int_distrib =
+  zadd_zmult_distrib zadd_zmult_distrib2
+  zdiff_zmult_distrib zdiff_zmult_distrib2
+
+lemma zmult_1_raw:
+  shows "mult_raw (1, 0) z = z"
+  by (cases z) (auto)
+
+lemma zmult_1:
+  fixes z::"int"
+  shows "1 * z = z"
+  by (lifting zmult_1_raw)
+
+lemma zmult_1_right: 
+  fixes z::"int"
+  shows "z * 1 = z"
+  by (rule trans [OF zmult_commute zmult_1])
+
+lemma zero_not_one:
+  shows "\<not>(intrel (0, 0) (1::nat, 0::nat))"
+  by auto
+
+text{*The Integers Form A Ring*}
+instance int :: comm_ring_1
+proof
+  fix i j k :: int
+  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
+  show "i + j = j + i" by (simp add: zadd_commute)
+  show "0 + i = i" by (rule zadd_0)
+  show "- i + i = 0" by (rule zadd_zminus_inverse2)
+  show "i - j = i + (-j)" by (simp add: diff_int_def)
+  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
+  show "i * j = j * i" by (rule zmult_commute)
+  show "1 * i = i" by (rule zmult_1) 
+  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
+  show "0 \<noteq> (1::int)" by (lifting zero_not_one)
+qed
+
+
+subsection{*The @{text "\<le>"} Ordering*}
+
+lemma zle_refl_raw:
+  shows "le_raw w w"
+  by (cases w) (simp add: le_raw_def)
+
+lemma [quot_respect]:
+  shows "(intrel ===> intrel ===> op =) le_raw le_raw"
+  by (auto) (simp_all add: le_raw_def)
+
+lemma zle_refl: 
+  fixes w::"int"
+  shows "w \<le> w"
+  by (lifting zle_refl_raw)
+
+
+lemma zle_trans_raw:
+  shows "\<lbrakk>le_raw i j; le_raw j k\<rbrakk> \<Longrightarrow> le_raw i k"
+apply(cases i, cases j, cases k)
+apply(auto simp add: le_raw_def)
+done
+
+lemma zle_trans: 
+  fixes i j k::"int"
+  shows "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> k"
+  by (lifting zle_trans_raw)
+
+lemma zle_anti_sym_raw:
+  shows "\<lbrakk>le_raw z w; le_raw w z\<rbrakk> \<Longrightarrow> intrel z w"
+apply(cases z, cases w)
+apply(auto iff: le_raw_def)
+done
+
+lemma zle_anti_sym: 
+  fixes z w::"int"
+  shows "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = w"
+  by (lifting zle_anti_sym_raw)
+
+
+(* Axiom 'order_less_le' of class 'order': *)
+lemma zless_le: 
+  fixes w z::"int"
+  shows "(w < z) = (w \<le> z & w \<noteq> z)"
+  by (simp add: less_int_def)
+
+instance int :: order
+apply (default)
+apply(auto simp add: zless_le zle_anti_sym)[1]
+apply(rule zle_refl)
+apply(erule zle_trans, assumption)
+apply(erule zle_anti_sym, assumption)
+done
+
+(* Axiom 'linorder_linear' of class 'linorder': *)
+
+lemma zle_linear_raw:
+  shows "le_raw z w \<or> le_raw w z"
+apply(cases w, cases z)
+apply(auto iff: le_raw_def)
+done
+
+lemma zle_linear: 
+  fixes z w::"int"
+  shows "z \<le> w \<or> w \<le> z"
+  by (lifting zle_linear_raw)
+
+instance int :: linorder
+apply(default)
+apply(rule zle_linear)
+done
+
+lemma zadd_left_mono_raw:
+  shows "le_raw i j \<Longrightarrow> le_raw (add_raw k i) (add_raw k j)"
+apply(cases k)
+apply(auto simp add: add_raw_def le_raw_def)
+done
+
+lemma zadd_left_mono: 
+  fixes i j::"int"
+  shows "i \<le> j \<Longrightarrow> k + i \<le> k + j"
+  by (lifting zadd_left_mono_raw)
+
+
+subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
+
+definition
+  "nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
+
+quotient_definition
+  "nat2::int \<Rightarrow> nat"
+is
+  "nat_raw"
+
+abbreviation
+  "less_raw x y \<equiv> (le_raw x y \<and> \<not>(intrel x y))"
+
+lemma nat_le_eq_zle_raw:
+  shows "less_raw (0, 0) w \<or> le_raw (0, 0) z \<Longrightarrow> (nat_raw w \<le> nat_raw z) = (le_raw w z)"
+  apply (cases w)
+  apply (cases z)
+  apply (simp add: nat_raw_def le_raw_def)
+  by auto
+
+lemma [quot_respect]:
+  shows "(intrel ===> op =) nat_raw nat_raw"
+  by (auto iff: nat_raw_def)
+
+lemma nat_le_eq_zle: 
+  fixes w z::"int"
+  shows "0 < w \<or> 0 \<le> z \<Longrightarrow> (nat2 w \<le> nat2 z) = (w\<le>z)"
+  unfolding less_int_def
+  by (lifting nat_le_eq_zle_raw)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/SigmaEx.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,253 @@
+theory SigmaEx
+imports Nominal "../Quotient" "../Quotient_List" "../Quotient_Product"
+begin
+
+atom_decl name
+
+datatype robj =
+  rVar "name"
+| rObj "(string \<times> rmethod) list"
+| rInv "robj" "string"
+| rUpd "robj" "string" "rmethod"
+and rmethod =
+  rSig "name" "robj"
+
+inductive
+    alpha_obj :: "robj \<Rightarrow> robj \<Rightarrow> bool" ("_ \<approx>o _" [100, 100] 100)
+and alpha_method :: "rmethod \<Rightarrow> rmethod \<Rightarrow> bool" ("_ \<approx>m _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (rVar a) \<approx>o (rVar b)"
+| a2: "rObj [] \<approx>o rObj []"
+| a3: "rObj t1 \<approx>o rObj t2 \<Longrightarrow> m1 \<approx>m r2 \<Longrightarrow> rObj ((l1, m1) # t1) \<approx>o rObj ((l2, m2) # t2)"
+| a4: "x \<approx>o y \<Longrightarrow> rInv x l1 \<approx>o rInv y l2"
+| a5: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx>o s \<and> (pi \<bullet> a) = b)
+       \<Longrightarrow> rSig a t \<approx>m rSig b s"
+
+lemma alpha_equivps:
+  shows "equivp alpha_obj"
+  and   "equivp alpha_method"
+sorry
+
+quotient_type
+    obj = robj / alpha_obj
+and method = rmethod / alpha_method
+  by (auto intro: alpha_equivps)
+
+quotient_definition
+  "Var :: name \<Rightarrow> obj"
+is
+  "rVar"
+
+quotient_definition
+  "Obj :: (string \<times> method) list \<Rightarrow> obj"
+is
+  "rObj"
+
+quotient_definition
+  "Inv :: obj \<Rightarrow> string \<Rightarrow> obj"
+is
+  "rInv"
+
+quotient_definition
+  "Upd :: obj \<Rightarrow> string \<Rightarrow> method \<Rightarrow> obj"
+is
+  "rUpd"
+
+quotient_definition
+  "Sig :: name \<Rightarrow> obj \<Rightarrow> method"
+is
+  "rSig"
+
+overloading
+  perm_obj \<equiv> "perm :: 'x prm \<Rightarrow> obj \<Rightarrow> obj" (unchecked)
+  perm_method   \<equiv> "perm :: 'x prm \<Rightarrow> method \<Rightarrow> method" (unchecked)
+begin
+
+quotient_definition
+  "perm_obj :: 'x prm \<Rightarrow> obj \<Rightarrow> obj"
+is
+  "(perm::'x prm \<Rightarrow> robj \<Rightarrow> robj)"
+
+quotient_definition
+  "perm_method :: 'x prm \<Rightarrow> method \<Rightarrow> method"
+is
+  "(perm::'x prm \<Rightarrow> rmethod \<Rightarrow> rmethod)"
+
+end
+
+
+
+lemma tolift:
+"\<forall> fvar.
+ \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
+ \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
+ \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<forall> fnil.
+ \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
+ \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
+
+ Ex1 (\<lambda>x.
+(x \<in> (Respects (prod_rel (alpha_obj ===> op =)
+     (prod_rel (list_rel (prod_rel (op =) alpha_method) ===> op =)
+       (prod_rel ((prod_rel (op =) alpha_method) ===> op =)
+         (alpha_method ===> op =)
+       )
+     )))) \<and>
+(\<lambda> (hom_o\<Colon>robj \<Rightarrow> 'a, hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b, hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c, hom_m\<Colon>rmethod \<Rightarrow> 'd).
+
+((\<forall>x. hom_o (rVar x) = fvar x) \<and>
+ (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
+ (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
+ (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
+ (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
+ (hom_d [] = fnil) \<and>
+ (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
+ (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
+)) x) "
+sorry
+
+lemma test_to: "Ex1 (\<lambda>x. (x \<in> (Respects alpha_obj)) \<and> P x)"
+ML_prf {* prop_of (#goal (Isar.goal ())) *}
+sorry
+lemma test_tod: "Ex1 (P :: obj \<Rightarrow> bool)"
+apply (lifting test_to)
+done
+
+
+
+
+(*syntax
+  "_expttrn"        :: "pttrn => bool => bool"      ("(3\<exists>\<exists> _./ _)" [0, 10] 10)
+
+translations
+  "\<exists>\<exists> x. P"   == "Ex (%x. P)"
+*)
+
+lemma rvar_rsp[quot_respect]: "(op = ===> alpha_obj) rVar rVar"
+  by (simp add: a1)
+
+lemma robj_rsp[quot_respect]: "(list_rel (prod_rel op = alpha_method) ===> alpha_obj) rObj rObj"
+sorry
+lemma rinv_rsp[quot_respect]: "(alpha_obj ===> op = ===> alpha_obj) rInv rInv"
+sorry
+lemma rupd_rsp[quot_respect]: "(alpha_obj ===> op = ===> alpha_method ===> alpha_obj) rUpd rUpd"
+sorry
+lemma rsig_rsp[quot_respect]: "(op = ===> alpha_obj ===> alpha_method) rSig rSig"
+sorry
+lemma operm_rsp[quot_respect]: "(op = ===> alpha_obj ===> alpha_obj) op \<bullet> op \<bullet>"
+sorry
+
+  
+lemma bex1_bex1reg: "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"
+apply (simp add: Ex1_def Bex1_rel_def in_respects)
+apply clarify
+apply auto
+apply (rule bexI)
+apply assumption
+apply (simp add: in_respects)
+apply (simp add: in_respects)
+apply auto
+done
+
+lemma liftd: "
+Ex1 (\<lambda>(hom_o, hom_d, hom_e, hom_m).
+
+ (\<forall>x. hom_o (Var x) = fvar x) \<and>
+ (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
+ (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
+ (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
+ (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
+ (hom_d [] = fnil) \<and>
+ (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
+ (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
+)"
+apply (lifting tolift)
+done
+
+lemma tolift':
+"\<forall> fvar.
+ \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
+ \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
+ \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<forall> fnil.
+ \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
+ \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
+
+ \<exists> hom_o\<Colon>robj \<Rightarrow> 'a \<in> Respects (alpha_obj ===> op =).
+ \<exists> hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b \<in> Respects (list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<exists> hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c \<in> Respects ((prod_rel (op =) alpha_method) ===> op =).
+ \<exists> hom_m\<Colon>rmethod \<Rightarrow> 'd \<in> Respects (alpha_method ===> op =).
+(
+ (\<forall>x. hom_o (rVar x) = fvar x) \<and>
+ (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
+ (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
+ (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
+ (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
+ (hom_d [] = fnil) \<and>
+ (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
+ (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
+)"
+sorry
+
+lemma liftd': "
+\<exists>hom_o. \<exists>hom_d. \<exists>hom_e. \<exists>hom_m.
+(
+ (\<forall>x. hom_o (Var x) = fvar x) \<and>
+ (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
+ (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
+ (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
+ (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
+ (hom_d [] = fnil) \<and>
+ (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
+ (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
+)"
+apply (lifting tolift')
+done
+
+lemma tolift'':
+"\<forall> fvar.
+ \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
+ \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
+ \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
+ \<forall> fnil.
+ \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
+ \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
+
+ Bex1_rel (alpha_obj ===> op =) (\<lambda>hom_o\<Colon>robj \<Rightarrow> 'a .
+ Bex1_rel (list_rel (prod_rel (op =) alpha_method) ===> op =) (\<lambda>hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b.
+ Bex1_rel ((prod_rel (op =) alpha_method) ===> op =) (\<lambda>hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c.
+ Bex1_rel (alpha_method ===> op =) (\<lambda>hom_m\<Colon>rmethod \<Rightarrow> 'd.
+(
+ (\<forall>x. hom_o (rVar x) = fvar x) \<and>
+ (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
+ (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
+ (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
+ (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
+ (hom_d [] = fnil) \<and>
+ (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
+ (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
+)
+))))"
+sorry
+
+lemma liftd'': "
+\<exists>!hom_o. \<exists>!hom_d. \<exists>!hom_e. \<exists>!hom_m.
+(
+ (\<forall>x. hom_o (Var x) = fvar x) \<and>
+ (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
+ (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
+ (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
+ (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
+ (hom_d [] = fnil) \<and>
+ (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
+ (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
+)"
+apply (lifting tolift'')
+done
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Examples/Terms.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,271 @@
+theory Terms
+imports Nominal "../Quotient" "../Quotient_List"
+begin
+
+atom_decl name
+
+text {* primrec seems to be genarally faster than fun *}
+
+section {*** lets with binding patterns ***}
+
+datatype trm1 =
+  Vr1 "name"
+| Ap1 "trm1" "trm1"
+| Lm1 "name" "trm1"        --"name is bound in trm1"
+| Lt1 "bp" "trm1" "trm1"   --"all variables in bp are bound in the 2nd trm1"
+and bp =
+  BUnit
+| BVr "name"
+| BPr "bp" "bp"
+
+(* to be given by the user *)
+primrec 
+  bv1
+where
+  "bv1 (BUnit) = {}"
+| "bv1 (BVr x) = {x}"
+| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)"
+
+(* needs to be calculated by the package *)
+primrec 
+  fv_trm1 and fv_bp
+where
+  "fv_trm1 (Vr1 x) = {x}"
+| "fv_trm1 (Ap1 t1 t2) = (fv_trm1 t1) \<union> (fv_trm1 t2)"
+| "fv_trm1 (Lm1 x t) = (fv_trm1 t) - {x}"
+| "fv_trm1 (Lt1 bp t1 t2) = (fv_trm1 t1) \<union> (fv_trm1 t2 - bv1 bp)"
+| "fv_bp (BUnit) = {}"
+| "fv_bp (BVr x) = {x}"
+| "fv_bp (BPr b1 b2) = (fv_bp b1) \<union> (fv_bp b2)"
+
+(* needs to be stated by the package *)
+overloading
+  perm_trm1 \<equiv> "perm :: 'x prm \<Rightarrow> trm1 \<Rightarrow> trm1"   (unchecked)
+  perm_bp \<equiv> "perm :: 'x prm \<Rightarrow> bp \<Rightarrow> bp" (unchecked)
+begin
+
+primrec
+  perm_trm1 and perm_bp
+where
+  "perm_trm1 pi (Vr1 a) = Vr1 (pi \<bullet> a)"
+| "perm_trm1 pi (Ap1 t1 t2) = Ap1 (perm_trm1 pi t1) (perm_trm1 pi t2)"
+| "perm_trm1 pi (Lm1 a t) = Lm1 (pi \<bullet> a) (perm_trm1 pi t)"
+| "perm_trm1 pi (Lt1 bp t1 t2) = Lt1 (perm_bp pi bp) (perm_trm1 pi t1) (perm_trm1 pi t2)"
+| "perm_bp pi (BUnit) = BUnit"
+| "perm_bp pi (BVr a) = BVr (pi \<bullet> a)"
+| "perm_bp pi (BPr bp1 bp2) = BPr (perm_bp pi bp1) (perm_bp pi bp2)"
+
+end
+
+inductive
+  alpha1 :: "trm1 \<Rightarrow> trm1 \<Rightarrow> bool" ("_ \<approx>1 _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (Vr1 a) \<approx>1 (Vr1 b)"
+| a2: "\<lbrakk>t1 \<approx>1 t2; s1 \<approx>1 s2\<rbrakk> \<Longrightarrow> Ap1 t1 s1 \<approx>1 Ap1 t2 s2"
+| a3: "\<exists>pi::name prm. (fv_trm1 t - {a} = fv_trm1 s - {b} \<and> 
+                      (fv_trm1 t - {a})\<sharp>* pi \<and> 
+                      (pi \<bullet> t) \<approx>1 s \<and> (pi \<bullet> a) = b)
+       \<Longrightarrow> Lm1 a t \<approx>1 Lm1 b s"
+| a4: "\<exists>pi::name prm.(
+         t1 \<approx>1 t2 \<and>
+         (fv_trm1 s1 - fv_bp b1 = fv_trm1 s2 - fv_bp b2) \<and>
+         (fv_trm1 s1 - fv_bp b1) \<sharp>* pi \<and>
+         (pi \<bullet> s1 = s2)                    (* Optional: \<and> (pi \<bullet> b1 = b2) *)
+       ) \<Longrightarrow> Lt1 b1 t1 s1 \<approx>1 Lt1 b2 t2 s2"
+
+lemma alpha1_equivp: "equivp alpha1" sorry
+
+quotient_type qtrm1 = trm1 / alpha1
+  by (rule alpha1_equivp)
+
+
+section {*** lets with single assignments ***}
+
+datatype trm2 =
+  Vr2 "name"
+| Ap2 "trm2" "trm2"
+| Lm2 "name" "trm2"
+| Lt2 "assign" "trm2"
+and assign =
+  As "name" "trm2"
+
+(* to be given by the user *)
+primrec 
+  bv2
+where
+  "bv2 (As x t) = {x}"
+
+(* needs to be calculated by the package *)
+primrec
+  fv_trm2 and fv_assign
+where
+  "fv_trm2 (Vr2 x) = {x}"
+| "fv_trm2 (Ap2 t1 t2) = (fv_trm2 t1) \<union> (fv_trm2 t2)"
+| "fv_trm2 (Lm2 x t) = (fv_trm2 t) - {x}"
+| "fv_trm2 (Lt2 as t) = (fv_trm2 t - bv2 as) \<union> (fv_assign as)"
+| "fv_assign (As x t) = (fv_trm2 t)"
+
+(* needs to be stated by the package *)
+overloading
+  perm_trm2 \<equiv> "perm :: 'x prm \<Rightarrow> trm2 \<Rightarrow> trm2"   (unchecked)
+  perm_assign \<equiv> "perm :: 'x prm \<Rightarrow> assign \<Rightarrow> assign" (unchecked)
+begin
+
+primrec
+  perm_trm2 and perm_assign
+where
+  "perm_trm2 pi (Vr2 a) = Vr2 (pi \<bullet> a)"
+| "perm_trm2 pi (Ap2 t1 t2) = Ap2 (perm_trm2 pi t1) (perm_trm2 pi t2)"
+| "perm_trm2 pi (Lm2 a t) = Lm2 (pi \<bullet> a) (perm_trm2 pi t)"
+| "perm_trm2 pi (Lt2 as t) = Lt2 (perm_assign pi as) (perm_trm2 pi t)"
+| "perm_assign pi (As a t) = As (pi \<bullet> a) (perm_trm2 pi t)"
+
+end
+
+inductive
+  alpha2 :: "trm2 \<Rightarrow> trm2 \<Rightarrow> bool" ("_ \<approx>2 _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (Vr2 a) \<approx>2 (Vr2 b)"
+| a2: "\<lbrakk>t1 \<approx>2 t2; s1 \<approx>2 s2\<rbrakk> \<Longrightarrow> Ap2 t1 s1 \<approx>2 Ap2 t2 s2"
+| a3: "\<exists>pi::name prm. (fv_trm2 t - {a} = fv_trm2 s - {b} \<and> 
+                      (fv_trm2 t - {a})\<sharp>* pi \<and> 
+                      (pi \<bullet> t) \<approx>2 s \<and> 
+                      (pi \<bullet> a) = b)
+       \<Longrightarrow> Lm2 a t \<approx>2 Lm2 b s"
+| a4: "\<exists>pi::name prm. (
+         fv_trm2 t1 - fv_assign b1 = fv_trm2 t2 - fv_assign b2 \<and>
+         (fv_trm2 t1 - fv_assign b1) \<sharp>* pi \<and>
+         pi \<bullet> t1 = t2       (* \<and> (pi \<bullet> b1 = b2) *)
+       ) \<Longrightarrow> Lt2 b1 t1 \<approx>2 Lt2 b2 t2"
+
+lemma alpha2_equivp: "equivp alpha2" sorry
+
+quotient_type qtrm2 = trm2 / alpha2
+  by (rule alpha2_equivp)
+
+section {*** lets with many assignments ***}
+
+datatype trm3 =
+  Vr3 "name"
+| Ap3 "trm3" "trm3"
+| Lm3 "name" "trm3"
+| Lt3 "assigns" "trm3"
+and assigns =
+  ANil
+| ACons "name" "trm3" "assigns"
+
+(* to be given by the user *)
+primrec 
+  bv3
+where
+  "bv3 ANil = {}"
+| "bv3 (ACons x t as) = {x} \<union> (bv3 as)"
+
+primrec
+  fv_trm3 and fv_assigns
+where
+  "fv_trm3 (Vr3 x) = {x}"
+| "fv_trm3 (Ap3 t1 t2) = (fv_trm3 t1) \<union> (fv_trm3 t2)"
+| "fv_trm3 (Lm3 x t) = (fv_trm3 t) - {x}"
+| "fv_trm3 (Lt3 as t) = (fv_trm3 t - bv3 as) \<union> (fv_assigns as)"
+| "fv_assigns (ANil) = {}"
+| "fv_assigns (ACons x t as) = (fv_trm3 t) \<union> (fv_assigns as)"
+
+(* needs to be stated by the package *)
+overloading
+  perm_trm3 \<equiv> "perm :: 'x prm \<Rightarrow> trm3 \<Rightarrow> trm3"   (unchecked)
+  perm_assigns \<equiv> "perm :: 'x prm \<Rightarrow> assigns \<Rightarrow> assigns" (unchecked)
+begin
+
+primrec
+  perm_trm3 and perm_assigns
+where
+  "perm_trm3 pi (Vr3 a) = Vr3 (pi \<bullet> a)"
+| "perm_trm3 pi (Ap3 t1 t2) = Ap3 (perm_trm3 pi t1) (perm_trm3 pi t2)"
+| "perm_trm3 pi (Lm3 a t) = Lm3 (pi \<bullet> a) (perm_trm3 pi t)"
+| "perm_trm3 pi (Lt3 as t) = Lt3 (perm_assigns pi as) (perm_trm3 pi t)"
+| "perm_assigns pi (ANil) = ANil"
+| "perm_assigns pi (ACons a t as) = ACons (pi \<bullet> a) (perm_trm3 pi t) (perm_assigns pi as)"
+
+end
+
+inductive
+  alpha3 :: "trm3 \<Rightarrow> trm3 \<Rightarrow> bool" ("_ \<approx>3 _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (Vr3 a) \<approx>3 (Vr3 b)"
+| a2: "\<lbrakk>t1 \<approx>3 t2; s1 \<approx>3 s2\<rbrakk> \<Longrightarrow> Ap3 t1 s1 \<approx>3 Ap3 t2 s2"
+| a3: "\<exists>pi::name prm. (fv_trm3 t - {a} = fv_trm3 s - {b} \<and> 
+                      (fv_trm3 t - {a})\<sharp>* pi \<and> 
+                      (pi \<bullet> t) \<approx>3 s \<and> 
+                      (pi \<bullet> a) = b)
+       \<Longrightarrow> Lm3 a t \<approx>3 Lm3 b s"
+| a4: "\<exists>pi::name prm. (
+         fv_trm3 t1 - fv_assigns b1 = fv_trm3 t2 - fv_assigns b2 \<and>
+         (fv_trm3 t1 - fv_assigns b1) \<sharp>* pi \<and>
+         pi \<bullet> t1 = t2      (* \<and> (pi \<bullet> b1 = b2)  *)
+       ) \<Longrightarrow> Lt3 b1 t1 \<approx>3 Lt3 b2 t2"
+
+lemma alpha3_equivp: "equivp alpha3" sorry
+
+quotient_type qtrm3 = trm3 / alpha3
+  by (rule alpha3_equivp)
+
+
+section {*** lam with indirect list recursion ***}
+
+datatype trm4 =
+  Vr4 "name"
+| Ap4 "trm4" "trm4 list"
+| Lm4 "name" "trm4"
+
+thm trm4.recs
+
+primrec
+  fv_trm4 and fv_trm4_list
+where
+  "fv_trm4 (Vr4 x) = {x}"
+| "fv_trm4 (Ap4 t ts) = (fv_trm4 t) \<union> (fv_trm4_list ts)"
+| "fv_trm4 (Lm4 x t) = (fv_trm4 t) - {x}"
+| "fv_trm4_list ([]) = {}"
+| "fv_trm4_list (t#ts) = (fv_trm4 t) \<union> (fv_trm4_list ts)"
+
+
+(* needs to be stated by the package *)
+(* there cannot be a clause for lists, as *) 
+(* permutations are  already defined in Nominal (also functions, options, and so on) *)
+overloading
+  perm_trm4 \<equiv> "perm :: 'x prm \<Rightarrow> trm4 \<Rightarrow> trm4"   (unchecked)
+begin
+
+primrec
+  perm_trm4 
+where
+  "perm_trm4 pi (Vr4 a) = Vr4 (pi \<bullet> a)"
+| "perm_trm4 pi (Ap4 t ts) = Ap4 (perm_trm4 pi t) (pi \<bullet> ts)"
+| "perm_trm4 pi (Lm4 a t) = Lm4 (pi \<bullet> a) (perm_trm4 pi t)"
+
+end
+
+inductive
+    alpha4 :: "trm4 \<Rightarrow> trm4 \<Rightarrow> bool" ("_ \<approx>4 _" [100, 100] 100)
+and alpha4list :: "trm4 list \<Rightarrow> trm4 list \<Rightarrow> bool" ("_ \<approx>4list _" [100, 100] 100) 
+where
+  a1: "a = b \<Longrightarrow> (Vr4 a) \<approx>4 (Vr4 b)"
+| a2: "\<lbrakk>t1 \<approx>4 t2; s1 \<approx>4list s2\<rbrakk> \<Longrightarrow> Ap4 t1 s1 \<approx>4 Ap4 t2 s2"
+| a4: "\<exists>pi::name prm. (fv_trm4 t - {a} = fv_trm4 s - {b} \<and> 
+                      (fv_trm4 t - {a})\<sharp>* pi \<and> 
+                      (pi \<bullet> t) \<approx>4 s \<and> 
+                      (pi \<bullet> a) = b)
+       \<Longrightarrow> Lm4 a t \<approx>4 Lm4 b s"
+| a5: "[] \<approx>4list []"
+| a6: "\<lbrakk>t \<approx>4 s; ts \<approx>4list ss\<rbrakk> \<Longrightarrow> (t#ts) \<approx>4list (s#ss)"
+
+lemma alpha4_equivp: "equivp alpha4" sorry
+lemma alpha4list_equivp: "equivp alpha4list" sorry
+
+quotient_type 
+  qtrm4 = trm4 / alpha4 and
+  qtrm4list = "trm4 list" / alpha4list
+  by (simp_all add: alpha4_equivp alpha4list_equivp)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Quotient.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,797 @@
+(*  Title:      Quotient.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+*)
+
+theory Quotient
+imports Plain ATP_Linkup
+uses
+  ("quotient_info.ML")
+  ("quotient_typ.ML")
+  ("quotient_def.ML")
+  ("quotient_term.ML")
+  ("quotient_tacs.ML")
+begin
+
+
+text {*
+  Basic definition for equivalence relations
+  that are represented by predicates.
+*}
+
+definition
+  "equivp E \<equiv> \<forall>x y. E x y = (E x = E y)"
+
+definition
+  "reflp E \<equiv> \<forall>x. E x x"
+
+definition
+  "symp E \<equiv> \<forall>x y. E x y \<longrightarrow> E y x"
+
+definition
+  "transp E \<equiv> \<forall>x y z. E x y \<and> E y z \<longrightarrow> E x z"
+
+lemma equivp_reflp_symp_transp:
+  shows "equivp E = (reflp E \<and> symp E \<and> transp E)"
+  unfolding equivp_def reflp_def symp_def transp_def expand_fun_eq
+  by blast
+
+lemma equivp_reflp:
+  shows "equivp E \<Longrightarrow> E x x"
+  by (simp only: equivp_reflp_symp_transp reflp_def)
+
+lemma equivp_symp:
+  shows "equivp E \<Longrightarrow> E x y \<Longrightarrow> E y x"
+  by (metis equivp_reflp_symp_transp symp_def)
+
+lemma equivp_transp:
+  shows "equivp E \<Longrightarrow> E x y \<Longrightarrow> E y z \<Longrightarrow> E x z"
+  by (metis equivp_reflp_symp_transp transp_def)
+
+lemma equivpI:
+  assumes "reflp R" "symp R" "transp R"
+  shows "equivp R"
+  using assms by (simp add: equivp_reflp_symp_transp)
+
+lemma identity_equivp:
+  shows "equivp (op =)"
+  unfolding equivp_def
+  by auto
+
+text {* Partial equivalences: not yet used anywhere *}
+
+definition
+  "part_equivp E \<equiv> (\<exists>x. E x x) \<and> (\<forall>x y. E x y = (E x x \<and> E y y \<and> (E x = E y)))"
+
+lemma equivp_implies_part_equivp:
+  assumes a: "equivp E"
+  shows "part_equivp E"
+  using a
+  unfolding equivp_def part_equivp_def
+  by auto
+
+text {* Composition of Relations *}
+
+abbreviation
+  rel_conj (infixr "OOO" 75)
+where
+  "r1 OOO r2 \<equiv> r1 OO r2 OO r1"
+
+lemma eq_comp_r:
+  shows "((op =) OOO R) = R"
+  by (auto simp add: expand_fun_eq)
+
+section {* Respects predicate *}
+
+definition
+  Respects
+where
+  "Respects R x \<equiv> R x x"
+
+lemma in_respects:
+  shows "(x \<in> Respects R) = R x x"
+  unfolding mem_def Respects_def
+  by simp
+
+section {* Function map and function relation *}
+
+definition
+  fun_map (infixr "--->" 55)
+where
+[simp]: "fun_map f g h x = g (h (f x))"
+
+definition
+  fun_rel (infixr "===>" 55)
+where
+[simp]: "fun_rel E1 E2 f g = (\<forall>x y. E1 x y \<longrightarrow> E2 (f x) (g y))"
+
+
+lemma fun_map_id:
+  shows "(id ---> id) = id"
+  by (simp add: expand_fun_eq id_def)
+
+lemma fun_rel_eq:
+  shows "((op =) ===> (op =)) = (op =)"
+  by (simp add: expand_fun_eq)
+
+lemma fun_rel_id:
+  assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)"
+  shows "(R1 ===> R2) f g"
+  using a by simp
+
+lemma fun_rel_id_asm:
+  assumes a: "\<And>x y. R1 x y \<Longrightarrow> (A \<longrightarrow> R2 (f x) (g y))"
+  shows "A \<longrightarrow> (R1 ===> R2) f g"
+  using a by auto
+
+
+section {* Quotient Predicate *}
+
+definition
+  "Quotient E Abs Rep \<equiv>
+     (\<forall>a. Abs (Rep a) = a) \<and> (\<forall>a. E (Rep a) (Rep a)) \<and>
+     (\<forall>r s. E r s = (E r r \<and> E s s \<and> (Abs r = Abs s)))"
+
+lemma Quotient_abs_rep:
+  assumes a: "Quotient E Abs Rep"
+  shows "Abs (Rep a) = a"
+  using a
+  unfolding Quotient_def
+  by simp
+
+lemma Quotient_rep_reflp:
+  assumes a: "Quotient E Abs Rep"
+  shows "E (Rep a) (Rep a)"
+  using a
+  unfolding Quotient_def
+  by blast
+
+lemma Quotient_rel:
+  assumes a: "Quotient E Abs Rep"
+  shows " E r s = (E r r \<and> E s s \<and> (Abs r = Abs s))"
+  using a
+  unfolding Quotient_def
+  by blast
+
+lemma Quotient_rel_rep:
+  assumes a: "Quotient R Abs Rep"
+  shows "R (Rep a) (Rep b) = (a = b)"
+  using a
+  unfolding Quotient_def
+  by metis
+
+lemma Quotient_rep_abs:
+  assumes a: "Quotient R Abs Rep"
+  shows "R r r \<Longrightarrow> R (Rep (Abs r)) r"
+  using a unfolding Quotient_def
+  by blast
+
+lemma Quotient_rel_abs:
+  assumes a: "Quotient E Abs Rep"
+  shows "E r s \<Longrightarrow> Abs r = Abs s"
+  using a unfolding Quotient_def
+  by blast
+
+lemma Quotient_symp:
+  assumes a: "Quotient E Abs Rep"
+  shows "symp E"
+  using a unfolding Quotient_def symp_def
+  by metis
+
+lemma Quotient_transp:
+  assumes a: "Quotient E Abs Rep"
+  shows "transp E"
+  using a unfolding Quotient_def transp_def
+  by metis
+
+lemma identity_quotient:
+  shows "Quotient (op =) id id"
+  unfolding Quotient_def id_def
+  by blast
+
+lemma fun_quotient:
+  assumes q1: "Quotient R1 abs1 rep1"
+  and     q2: "Quotient R2 abs2 rep2"
+  shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)"
+proof -
+  have "\<forall>a. (rep1 ---> abs2) ((abs1 ---> rep2) a) = a"
+    using q1 q2
+    unfolding Quotient_def
+    unfolding expand_fun_eq
+    by simp
+  moreover
+  have "\<forall>a. (R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)"
+    using q1 q2
+    unfolding Quotient_def
+    by (simp (no_asm)) (metis)
+  moreover
+  have "\<forall>r s. (R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and>
+        (rep1 ---> abs2) r  = (rep1 ---> abs2) s)"
+    unfolding expand_fun_eq
+    apply(auto)
+    using q1 q2 unfolding Quotient_def
+    apply(metis)
+    using q1 q2 unfolding Quotient_def
+    apply(metis)
+    using q1 q2 unfolding Quotient_def
+    apply(metis)
+    using q1 q2 unfolding Quotient_def
+    apply(metis)
+    done
+  ultimately
+  show "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)"
+    unfolding Quotient_def by blast
+qed
+
+lemma abs_o_rep:
+  assumes a: "Quotient R Abs Rep"
+  shows "Abs o Rep = id"
+  unfolding expand_fun_eq
+  by (simp add: Quotient_abs_rep[OF a])
+
+lemma equals_rsp:
+  assumes q: "Quotient R Abs Rep"
+  and     a: "R xa xb" "R ya yb"
+  shows "R xa ya = R xb yb"
+  using a Quotient_symp[OF q] Quotient_transp[OF q]
+  unfolding symp_def transp_def
+  by blast
+
+lemma lambda_prs:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)"
+  unfolding expand_fun_eq
+  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
+  by simp
+
+lemma lambda_prs1:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "(Rep1 ---> Abs2) (\<lambda>x. (Abs1 ---> Rep2) f x) = (\<lambda>x. f x)"
+  unfolding expand_fun_eq
+  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
+  by simp
+
+lemma rep_abs_rsp:
+  assumes q: "Quotient R Abs Rep"
+  and     a: "R x1 x2"
+  shows "R x1 (Rep (Abs x2))"
+  using a Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]
+  by metis
+
+lemma rep_abs_rsp_left:
+  assumes q: "Quotient R Abs Rep"
+  and     a: "R x1 x2"
+  shows "R (Rep (Abs x1)) x2"
+  using a Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]
+  by metis
+
+text{*
+  In the following theorem R1 can be instantiated with anything,
+  but we know some of the types of the Rep and Abs functions;
+  so by solving Quotient assumptions we can get a unique R1 that
+  will be provable; which is why we need to use apply_rsp and
+  not the primed version *}
+
+lemma apply_rsp:
+  fixes f g::"'a \<Rightarrow> 'c"
+  assumes q: "Quotient R1 Abs1 Rep1"
+  and     a: "(R1 ===> R2) f g" "R1 x y"
+  shows "R2 (f x) (g y)"
+  using a by simp
+
+lemma apply_rsp':
+  assumes a: "(R1 ===> R2) f g" "R1 x y"
+  shows "R2 (f x) (g y)"
+  using a by simp
+
+section {* lemmas for regularisation of ball and bex *}
+
+lemma ball_reg_eqv:
+  fixes P :: "'a \<Rightarrow> bool"
+  assumes a: "equivp R"
+  shows "Ball (Respects R) P = (All P)"
+  using a
+  unfolding equivp_def
+  by (auto simp add: in_respects)
+
+lemma bex_reg_eqv:
+  fixes P :: "'a \<Rightarrow> bool"
+  assumes a: "equivp R"
+  shows "Bex (Respects R) P = (Ex P)"
+  using a
+  unfolding equivp_def
+  by (auto simp add: in_respects)
+
+lemma ball_reg_right:
+  assumes a: "\<And>x. R x \<Longrightarrow> P x \<longrightarrow> Q x"
+  shows "All P \<longrightarrow> Ball R Q"
+  using a by (metis COMBC_def Collect_def Collect_mem_eq)
+
+lemma bex_reg_left:
+  assumes a: "\<And>x. R x \<Longrightarrow> Q x \<longrightarrow> P x"
+  shows "Bex R Q \<longrightarrow> Ex P"
+  using a by (metis COMBC_def Collect_def Collect_mem_eq)
+
+lemma ball_reg_left:
+  assumes a: "equivp R"
+  shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P"
+  using a by (metis equivp_reflp in_respects)
+
+lemma bex_reg_right:
+  assumes a: "equivp R"
+  shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P"
+  using a by (metis equivp_reflp in_respects)
+
+lemma ball_reg_eqv_range:
+  fixes P::"'a \<Rightarrow> bool"
+  and x::"'a"
+  assumes a: "equivp R2"
+  shows   "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))"
+  apply(rule iffI)
+  apply(rule allI)
+  apply(drule_tac x="\<lambda>y. f x" in bspec)
+  apply(simp add: in_respects)
+  apply(rule impI)
+  using a equivp_reflp_symp_transp[of "R2"]
+  apply(simp add: reflp_def)
+  apply(simp)
+  apply(simp)
+  done
+
+lemma bex_reg_eqv_range:
+  assumes a: "equivp R2"
+  shows   "(Bex (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = Ex (\<lambda>f. P (f x)))"
+  apply(auto)
+  apply(rule_tac x="\<lambda>y. f x" in bexI)
+  apply(simp)
+  apply(simp add: Respects_def in_respects)
+  apply(rule impI)
+  using a equivp_reflp_symp_transp[of "R2"]
+  apply(simp add: reflp_def)
+  done
+
+(* Next four lemmas are unused *)
+lemma all_reg:
+  assumes a: "!x :: 'a. (P x --> Q x)"
+  and     b: "All P"
+  shows "All Q"
+  using a b by (metis)
+
+lemma ex_reg:
+  assumes a: "!x :: 'a. (P x --> Q x)"
+  and     b: "Ex P"
+  shows "Ex Q"
+  using a b by metis
+
+lemma ball_reg:
+  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
+  and     b: "Ball R P"
+  shows "Ball R Q"
+  using a b by (metis COMBC_def Collect_def Collect_mem_eq)
+
+lemma bex_reg:
+  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
+  and     b: "Bex R P"
+  shows "Bex R Q"
+  using a b by (metis COMBC_def Collect_def Collect_mem_eq)
+
+
+lemma ball_all_comm:
+  assumes "\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)"
+  shows "(\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y)"
+  using assms by auto
+
+lemma bex_ex_comm:
+  assumes "(\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)"
+  shows "(\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y)"
+  using assms by auto
+
+section {* Bounded abstraction *}
+
+definition
+  Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+where
+  "x \<in> p \<Longrightarrow> Babs p m x = m x"
+
+lemma babs_rsp:
+  assumes q: "Quotient R1 Abs1 Rep1"
+  and     a: "(R1 ===> R2) f g"
+  shows      "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)"
+  apply (auto simp add: Babs_def in_respects)
+  apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
+  using a apply (simp add: Babs_def)
+  apply (simp add: in_respects)
+  using Quotient_rel[OF q]
+  by metis
+
+lemma babs_prs:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f"
+  apply (rule ext)
+  apply (simp)
+  apply (subgoal_tac "Rep1 x \<in> Respects R1")
+  apply (simp add: Babs_def Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
+  apply (simp add: in_respects Quotient_rel_rep[OF q1])
+  done
+
+lemma babs_simp:
+  assumes q: "Quotient R1 Abs Rep"
+  shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)"
+  apply(rule iffI)
+  apply(simp_all only: babs_rsp[OF q])
+  apply(auto simp add: Babs_def)
+  apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
+  apply(metis Babs_def)
+  apply (simp add: in_respects)
+  using Quotient_rel[OF q]
+  by metis
+
+(* If a user proves that a particular functional relation
+   is an equivalence this may be useful in regularising *)
+lemma babs_reg_eqv:
+  shows "equivp R \<Longrightarrow> Babs (Respects R) P = P"
+  by (simp add: expand_fun_eq Babs_def in_respects equivp_reflp)
+
+
+(* 3 lemmas needed for proving repabs_inj *)
+lemma ball_rsp:
+  assumes a: "(R ===> (op =)) f g"
+  shows "Ball (Respects R) f = Ball (Respects R) g"
+  using a by (simp add: Ball_def in_respects)
+
+lemma bex_rsp:
+  assumes a: "(R ===> (op =)) f g"
+  shows "(Bex (Respects R) f = Bex (Respects R) g)"
+  using a by (simp add: Bex_def in_respects)
+
+lemma bex1_rsp:
+  assumes a: "(R ===> (op =)) f g"
+  shows "Ex1 (\<lambda>x. x \<in> Respects R \<and> f x) = Ex1 (\<lambda>x. x \<in> Respects R \<and> g x)"
+  using a
+  by (simp add: Ex1_def in_respects) auto
+
+(* 2 lemmas needed for cleaning of quantifiers *)
+lemma all_prs:
+  assumes a: "Quotient R absf repf"
+  shows "Ball (Respects R) ((absf ---> id) f) = All f"
+  using a unfolding Quotient_def Ball_def in_respects fun_map_def id_apply
+  by metis
+
+lemma ex_prs:
+  assumes a: "Quotient R absf repf"
+  shows "Bex (Respects R) ((absf ---> id) f) = Ex f"
+  using a unfolding Quotient_def Bex_def in_respects fun_map_def id_apply
+  by metis
+
+section {* Bex1_rel quantifier *}
+
+definition
+  Bex1_rel :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+where
+  "Bex1_rel R P \<longleftrightarrow> (\<exists>x \<in> Respects R. P x) \<and> (\<forall>x \<in> Respects R. \<forall>y \<in> Respects R. ((P x \<and> P y) \<longrightarrow> (R x y)))"
+
+lemma bex1_rel_aux:
+  "\<lbrakk>\<forall>xa ya. R xa ya \<longrightarrow> x xa = y ya; Bex1_rel R x\<rbrakk> \<Longrightarrow> Bex1_rel R y"
+  unfolding Bex1_rel_def
+  apply (erule conjE)+
+  apply (erule bexE)
+  apply rule
+  apply (rule_tac x="xa" in bexI)
+  apply metis
+  apply metis
+  apply rule+
+  apply (erule_tac x="xaa" in ballE)
+  prefer 2
+  apply (metis)
+  apply (erule_tac x="ya" in ballE)
+  prefer 2
+  apply (metis)
+  apply (metis in_respects)
+  done
+
+lemma bex1_rel_aux2:
+  "\<lbrakk>\<forall>xa ya. R xa ya \<longrightarrow> x xa = y ya; Bex1_rel R y\<rbrakk> \<Longrightarrow> Bex1_rel R x"
+  unfolding Bex1_rel_def
+  apply (erule conjE)+
+  apply (erule bexE)
+  apply rule
+  apply (rule_tac x="xa" in bexI)
+  apply metis
+  apply metis
+  apply rule+
+  apply (erule_tac x="xaa" in ballE)
+  prefer 2
+  apply (metis)
+  apply (erule_tac x="ya" in ballE)
+  prefer 2
+  apply (metis)
+  apply (metis in_respects)
+  done
+
+lemma bex1_rel_rsp:
+  assumes a: "Quotient R absf repf"
+  shows "((R ===> op =) ===> op =) (Bex1_rel R) (Bex1_rel R)"
+  apply simp
+  apply clarify
+  apply rule
+  apply (simp_all add: bex1_rel_aux bex1_rel_aux2)
+  apply (erule bex1_rel_aux2)
+  apply assumption
+  done
+
+
+lemma ex1_prs:
+  assumes a: "Quotient R absf repf"
+  shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f"
+apply simp
+apply (subst Bex1_rel_def)
+apply (subst Bex_def)
+apply (subst Ex1_def)
+apply simp
+apply rule
+ apply (erule conjE)+
+ apply (erule_tac exE)
+ apply (erule conjE)
+ apply (subgoal_tac "\<forall>y. R y y \<longrightarrow> f (absf y) \<longrightarrow> R x y")
+  apply (rule_tac x="absf x" in exI)
+  apply (simp)
+  apply rule+
+  using a unfolding Quotient_def
+  apply metis
+ apply rule+
+ apply (erule_tac x="x" in ballE)
+  apply (erule_tac x="y" in ballE)
+   apply simp
+  apply (simp add: in_respects)
+ apply (simp add: in_respects)
+apply (erule_tac exE)
+ apply rule
+ apply (rule_tac x="repf x" in exI)
+ apply (simp only: in_respects)
+  apply rule
+ apply (metis Quotient_rel_rep[OF a])
+using a unfolding Quotient_def apply (simp)
+apply rule+
+using a unfolding Quotient_def in_respects
+apply metis
+done
+
+lemma bex1_bexeq_reg: "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"
+  apply (simp add: Ex1_def Bex1_rel_def in_respects)
+  apply clarify
+  apply auto
+  apply (rule bexI)
+  apply assumption
+  apply (simp add: in_respects)
+  apply (simp add: in_respects)
+  apply auto
+  done
+
+section {* Various respects and preserve lemmas *}
+
+lemma quot_rel_rsp:
+  assumes a: "Quotient R Abs Rep"
+  shows "(R ===> R ===> op =) R R"
+  apply(rule fun_rel_id)+
+  apply(rule equals_rsp[OF a])
+  apply(assumption)+
+  done
+
+lemma o_prs:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  and     q3: "Quotient R3 Abs3 Rep3"
+  shows "(Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g)) = f o g"
+  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3]
+  unfolding o_def expand_fun_eq by simp
+
+lemma o_rsp:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  and     q3: "Quotient R3 Abs3 Rep3"
+  and     a1: "(R2 ===> R3) f1 f2"
+  and     a2: "(R1 ===> R2) g1 g2"
+  shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
+  using a1 a2 unfolding o_def expand_fun_eq
+  by (auto)
+
+lemma cond_prs:
+  assumes a: "Quotient R absf repf"
+  shows "absf (if a then repf b else repf c) = (if a then b else c)"
+  using a unfolding Quotient_def by auto
+
+lemma if_prs:
+  assumes q: "Quotient R Abs Rep"
+  shows "Abs (If a (Rep b) (Rep c)) = If a b c"
+  using Quotient_abs_rep[OF q] by auto
+
+(* q not used *)
+lemma if_rsp:
+  assumes q: "Quotient R Abs Rep"
+  and     a: "a1 = a2" "R b1 b2" "R c1 c2"
+  shows "R (If a1 b1 c1) (If a2 b2 c2)"
+  using a by auto
+
+lemma let_prs:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f)) = Let x f"
+  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
+
+lemma let_rsp:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     a1: "(R1 ===> R2) f g"
+  and     a2: "R1 x y"
+  shows "R2 ((Let x f)::'c) ((Let y g)::'c)"
+  using apply_rsp[OF q1 a1] a2 by auto
+
+locale quot_type =
+  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
+  and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
+  assumes equivp: "equivp R"
+  and     rep_prop: "\<And>y. \<exists>x. Rep y = R x"
+  and     rep_inverse: "\<And>x. Abs (Rep x) = x"
+  and     abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
+  and     rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
+begin
+
+definition
+  abs::"'a \<Rightarrow> 'b"
+where
+  "abs x \<equiv> Abs (R x)"
+
+definition
+  rep::"'b \<Rightarrow> 'a"
+where
+  "rep a = Eps (Rep a)"
+
+lemma homeier_lem9:
+  shows "R (Eps (R x)) = R x"
+proof -
+  have a: "R x x" using equivp by (simp add: equivp_reflp_symp_transp reflp_def)
+  then have "R x (Eps (R x))" by (rule someI)
+  then show "R (Eps (R x)) = R x"
+    using equivp unfolding equivp_def by simp
+qed
+
+theorem homeier_thm10:
+  shows "abs (rep a) = a"
+  unfolding abs_def rep_def
+proof -
+  from rep_prop
+  obtain x where eq: "Rep a = R x" by auto
+  have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
+  also have "\<dots> = Abs (R x)" using homeier_lem9 by simp
+  also have "\<dots> = Abs (Rep a)" using eq by simp
+  also have "\<dots> = a" using rep_inverse by simp
+  finally
+  show "Abs (R (Eps (Rep a))) = a" by simp
+qed
+
+lemma homeier_lem7:
+  shows "(R x = R y) = (Abs (R x) = Abs (R y))" (is "?LHS = ?RHS")
+proof -
+  have "?RHS = (Rep (Abs (R x)) = Rep (Abs (R y)))" by (simp add: rep_inject)
+  also have "\<dots> = ?LHS" by (simp add: abs_inverse)
+  finally show "?LHS = ?RHS" by simp
+qed
+
+theorem homeier_thm11:
+  shows "R r r' = (abs r = abs r')"
+  unfolding abs_def
+  by (simp only: equivp[simplified equivp_def] homeier_lem7)
+
+lemma rep_refl:
+  shows "R (rep a) (rep a)"
+  unfolding rep_def
+  by (simp add: equivp[simplified equivp_def])
+
+
+lemma rep_abs_rsp:
+  shows "R f (rep (abs g)) = R f g"
+  and   "R (rep (abs g)) f = R g f"
+  by (simp_all add: homeier_thm10 homeier_thm11)
+
+lemma Quotient:
+  shows "Quotient R abs rep"
+  unfolding Quotient_def
+  apply(simp add: homeier_thm10)
+  apply(simp add: rep_refl)
+  apply(subst homeier_thm11[symmetric])
+  apply(simp add: equivp[simplified equivp_def])
+  done
+
+end
+
+section {* ML setup *}
+
+text {* Auxiliary data for the quotient package *}
+
+use "quotient_info.ML"
+
+declare [[map "fun" = (fun_map, fun_rel)]]
+
+lemmas [quot_thm] = fun_quotient
+lemmas [quot_respect] = quot_rel_rsp
+lemmas [quot_equiv] = identity_equivp
+
+
+text {* Lemmas about simplifying id's. *}
+lemmas [id_simps] =
+  id_def[symmetric]
+  fun_map_id
+  id_apply
+  id_o
+  o_id
+  eq_comp_r
+
+text {* Translation functions for the lifting process. *}
+use "quotient_term.ML"
+
+
+text {* Definitions of the quotient types. *}
+use "quotient_typ.ML"
+
+
+text {* Definitions for quotient constants. *}
+use "quotient_def.ML"
+
+
+text {*
+  An auxiliary constant for recording some information
+  about the lifted theorem in a tactic.
+*}
+definition
+  "Quot_True x \<equiv> True"
+
+lemma
+  shows QT_all: "Quot_True (All P) \<Longrightarrow> Quot_True P"
+  and   QT_ex:  "Quot_True (Ex P) \<Longrightarrow> Quot_True P"
+  and   QT_ex1: "Quot_True (Ex1 P) \<Longrightarrow> Quot_True P"
+  and   QT_lam: "Quot_True (\<lambda>x. P x) \<Longrightarrow> (\<And>x. Quot_True (P x))"
+  and   QT_ext: "(\<And>x. Quot_True (a x) \<Longrightarrow> f x = g x) \<Longrightarrow> (Quot_True a \<Longrightarrow> f = g)"
+  by (simp_all add: Quot_True_def ext)
+
+lemma QT_imp: "Quot_True a \<equiv> Quot_True b"
+  by (simp add: Quot_True_def)
+
+
+text {* Tactics for proving the lifted theorems *}
+use "quotient_tacs.ML"
+
+section {* Methods / Interface *}
+
+method_setup lifting =
+  {* Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt thms))) *}
+  {* lifts theorems to quotient types *}
+
+method_setup lifting_setup =
+  {* Attrib.thm >> (fn thms => fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.procedure_tac ctxt thms))) *}
+  {* sets up the three goals for the quotient lifting procedure *}
+
+method_setup regularize =
+  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.regularize_tac ctxt))) *}
+  {* proves the regularization goals from the quotient lifting procedure *}
+
+method_setup injection =
+  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.all_injection_tac ctxt))) *}
+  {* proves the rep/abs injection goals from the quotient lifting procedure *}
+
+method_setup cleaning =
+  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.clean_tac ctxt))) *}
+  {* proves the cleaning goals from the quotient lifting procedure *}
+
+attribute_setup quot_lifted =
+  {* Scan.succeed Quotient_Tacs.lifted_attrib *}
+  {* lifts theorems to quotient types *}
+
+no_notation
+  rel_conj (infixr "OOO" 75) and
+  fun_map (infixr "--->" 55) and
+  fun_rel (infixr "===>" 55)
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Quotient_List.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,232 @@
+(*  Title:      Quotient_List.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+*)
+theory Quotient_List
+imports Quotient Quotient_Syntax List
+begin
+
+section {* Quotient infrastructure for the list type. *}
+
+fun
+  list_rel
+where
+  "list_rel R [] [] = True"
+| "list_rel R (x#xs) [] = False"
+| "list_rel R [] (x#xs) = False"
+| "list_rel R (x#xs) (y#ys) = (R x y \<and> list_rel R xs ys)"
+
+declare [[map list = (map, list_rel)]]
+
+lemma split_list_all:
+  shows "(\<forall>x. P x) \<longleftrightarrow> P [] \<and> (\<forall>x xs. P (x#xs))"
+  apply(auto)
+  apply(case_tac x)
+  apply(simp_all)
+  done
+
+lemma map_id[id_simps]:
+  shows "map id = id"
+  apply(simp add: expand_fun_eq)
+  apply(rule allI)
+  apply(induct_tac x)
+  apply(simp_all)
+  done
+
+
+lemma list_rel_reflp:
+  shows "equivp R \<Longrightarrow> list_rel R xs xs"
+  apply(induct xs)
+  apply(simp_all add: equivp_reflp)
+  done
+
+lemma list_rel_symp:
+  assumes a: "equivp R"
+  shows "list_rel R xs ys \<Longrightarrow> list_rel R ys xs"
+  apply(induct xs ys rule: list_induct2')
+  apply(simp_all)
+  apply(rule equivp_symp[OF a])
+  apply(simp)
+  done
+
+lemma list_rel_transp:
+  assumes a: "equivp R"
+  shows "list_rel R xs1 xs2 \<Longrightarrow> list_rel R xs2 xs3 \<Longrightarrow> list_rel R xs1 xs3"
+  apply(induct xs1 xs2 arbitrary: xs3 rule: list_induct2')
+  apply(simp_all)
+  apply(case_tac xs3)
+  apply(simp_all)
+  apply(rule equivp_transp[OF a])
+  apply(auto)
+  done
+
+lemma list_equivp[quot_equiv]:
+  assumes a: "equivp R"
+  shows "equivp (list_rel R)"
+  apply(rule equivpI)
+  unfolding reflp_def symp_def transp_def
+  apply(subst split_list_all)
+  apply(simp add: equivp_reflp[OF a] list_rel_reflp[OF a])
+  apply(blast intro: list_rel_symp[OF a])
+  apply(blast intro: list_rel_transp[OF a])
+  done
+
+lemma list_rel_rel:
+  assumes q: "Quotient R Abs Rep"
+  shows "list_rel R r s = (list_rel R r r \<and> list_rel R s s \<and> (map Abs r = map Abs s))"
+  apply(induct r s rule: list_induct2')
+  apply(simp_all)
+  using Quotient_rel[OF q]
+  apply(metis)
+  done
+
+lemma list_quotient[quot_thm]:
+  assumes q: "Quotient R Abs Rep"
+  shows "Quotient (list_rel R) (map Abs) (map Rep)"
+  unfolding Quotient_def
+  apply(subst split_list_all)
+  apply(simp add: Quotient_abs_rep[OF q] abs_o_rep[OF q] map_id)
+  apply(rule conjI)
+  apply(rule allI)
+  apply(induct_tac a)
+  apply(simp)
+  apply(simp)
+  apply(simp add: Quotient_rep_reflp[OF q])
+  apply(rule allI)+
+  apply(rule list_rel_rel[OF q])
+  done
+
+
+lemma cons_prs_aux:
+  assumes q: "Quotient R Abs Rep"
+  shows "(map Abs) ((Rep h) # (map Rep t)) = h # t"
+  by (induct t) (simp_all add: Quotient_abs_rep[OF q])
+
+lemma cons_prs[quot_preserve]:
+  assumes q: "Quotient R Abs Rep"
+  shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)"
+  by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q])
+     (simp)
+
+lemma cons_rsp[quot_respect]:
+  assumes q: "Quotient R Abs Rep"
+  shows "(R ===> list_rel R ===> list_rel R) (op #) (op #)"
+  by (auto)
+
+lemma nil_prs[quot_preserve]:
+  assumes q: "Quotient R Abs Rep"
+  shows "map Abs [] = []"
+  by simp
+
+lemma nil_rsp[quot_respect]:
+  assumes q: "Quotient R Abs Rep"
+  shows "list_rel R [] []"
+  by simp
+
+lemma map_prs_aux:
+  assumes a: "Quotient R1 abs1 rep1"
+  and     b: "Quotient R2 abs2 rep2"
+  shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l"
+  by (induct l)
+     (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
+
+
+lemma map_prs[quot_preserve]:
+  assumes a: "Quotient R1 abs1 rep1"
+  and     b: "Quotient R2 abs2 rep2"
+  shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map"
+  by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b])
+     (simp)
+
+
+lemma map_rsp[quot_respect]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "((R1 ===> R2) ===> (list_rel R1) ===> list_rel R2) map map"
+  apply(simp)
+  apply(rule allI)+
+  apply(rule impI)
+  apply(rule allI)+
+  apply (induct_tac xa ya rule: list_induct2')
+  apply simp_all
+  done
+
+lemma foldr_prs_aux:
+  assumes a: "Quotient R1 abs1 rep1"
+  and     b: "Quotient R2 abs2 rep2"
+  shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e"
+  by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
+
+lemma foldr_prs[quot_preserve]:
+  assumes a: "Quotient R1 abs1 rep1"
+  and     b: "Quotient R2 abs2 rep2"
+  shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr"
+  by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b])
+     (simp)
+
+lemma foldl_prs_aux:
+  assumes a: "Quotient R1 abs1 rep1"
+  and     b: "Quotient R2 abs2 rep2"
+  shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l"
+  by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
+
+
+lemma foldl_prs[quot_preserve]:
+  assumes a: "Quotient R1 abs1 rep1"
+  and     b: "Quotient R2 abs2 rep2"
+  shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl"
+  by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b])
+     (simp)
+
+lemma list_rel_empty:
+  shows "list_rel R [] b \<Longrightarrow> length b = 0"
+  by (induct b) (simp_all)
+
+lemma list_rel_len:
+  shows "list_rel R a b \<Longrightarrow> length a = length b"
+  apply (induct a arbitrary: b)
+  apply (simp add: list_rel_empty)
+  apply (case_tac b)
+  apply simp_all
+  done
+
+(* induct_tac doesn't accept 'arbitrary', so we manually 'spec' *)
+lemma foldl_rsp[quot_respect]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_rel R2 ===> R1) foldl foldl"
+  apply(auto)
+  apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)")
+  apply simp
+  apply (rule_tac x="xa" in spec)
+  apply (rule_tac x="ya" in spec)
+  apply (rule_tac xs="xb" and ys="yb" in list_induct2)
+  apply (rule list_rel_len)
+  apply (simp_all)
+  done
+
+lemma foldr_rsp[quot_respect]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "((R1 ===> R2 ===> R2) ===> list_rel R1 ===> R2 ===> R2) foldr foldr"
+  apply auto
+  apply(subgoal_tac "R2 xb yb \<longrightarrow> list_rel R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)")
+  apply simp
+  apply (rule_tac xs="xa" and ys="ya" in list_induct2)
+  apply (rule list_rel_len)
+  apply (simp_all)
+  done
+
+lemma list_rel_eq[id_simps]:
+  shows "(list_rel (op =)) = (op =)"
+  unfolding expand_fun_eq
+  apply(rule allI)+
+  apply(induct_tac x xa rule: list_induct2')
+  apply(simp_all)
+  done
+
+lemma list_rel_refl:
+  assumes a: "\<And>x y. R x y = (R x = R y)"
+  shows "list_rel R x x"
+  by (induct x) (auto simp add: a)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Quotient_Option.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,80 @@
+(*  Title:      Quotient_Option.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+*)
+theory Quotient_Option
+imports Quotient Quotient_Syntax
+begin
+
+section {* Quotient infrastructure for the option type. *}
+
+fun
+  option_rel
+where
+  "option_rel R None None = True"
+| "option_rel R (Some x) None = False"
+| "option_rel R None (Some x) = False"
+| "option_rel R (Some x) (Some y) = R x y"
+
+declare [[map option = (Option.map, option_rel)]]
+
+text {* should probably be in Option.thy *}
+lemma split_option_all:
+  shows "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>a. P (Some a))"
+  apply(auto)
+  apply(case_tac x)
+  apply(simp_all)
+  done
+
+lemma option_quotient[quot_thm]:
+  assumes q: "Quotient R Abs Rep"
+  shows "Quotient (option_rel R) (Option.map Abs) (Option.map Rep)"
+  unfolding Quotient_def
+  apply(simp add: split_option_all)
+  apply(simp add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q])
+  using q
+  unfolding Quotient_def
+  apply(blast)
+  done
+
+lemma option_equivp[quot_equiv]:
+  assumes a: "equivp R"
+  shows "equivp (option_rel R)"
+  apply(rule equivpI)
+  unfolding reflp_def symp_def transp_def
+  apply(simp_all add: split_option_all)
+  apply(blast intro: equivp_reflp[OF a])
+  apply(blast intro: equivp_symp[OF a])
+  apply(blast intro: equivp_transp[OF a])
+  done
+
+lemma option_None_rsp[quot_respect]:
+  assumes q: "Quotient R Abs Rep"
+  shows "option_rel R None None"
+  by simp
+
+lemma option_Some_rsp[quot_respect]:
+  assumes q: "Quotient R Abs Rep"
+  shows "(R ===> option_rel R) Some Some"
+  by simp
+
+lemma option_None_prs[quot_preserve]:
+  assumes q: "Quotient R Abs Rep"
+  shows "Option.map Abs None = None"
+  by simp
+
+lemma option_Some_prs[quot_preserve]:
+  assumes q: "Quotient R Abs Rep"
+  shows "(Rep ---> Option.map Abs) Some = Some"
+  apply(simp add: expand_fun_eq)
+  apply(simp add: Quotient_abs_rep[OF q])
+  done
+
+lemma option_map_id[id_simps]:
+  shows "Option.map id = id"
+  by (simp add: expand_fun_eq split_option_all)
+
+lemma option_rel_eq[id_simps]:
+  shows "option_rel (op =) = (op =)"
+  by (simp add: expand_fun_eq split_option_all)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Quotient_Product.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,104 @@
+(*  Title:      Quotient_Product.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+*)
+theory Quotient_Product
+imports Quotient Quotient_Syntax
+begin
+
+section {* Quotient infrastructure for the product type. *}
+
+fun
+  prod_rel
+where
+  "prod_rel R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)"
+
+declare [[map * = (prod_fun, prod_rel)]]
+
+
+lemma prod_equivp[quot_equiv]:
+  assumes a: "equivp R1"
+  assumes b: "equivp R2"
+  shows "equivp (prod_rel R1 R2)"
+  apply(rule equivpI)
+  unfolding reflp_def symp_def transp_def
+  apply(simp_all add: split_paired_all)
+  apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
+  apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
+  apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
+  done
+
+lemma prod_quotient[quot_thm]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
+  unfolding Quotient_def
+  apply(simp add: split_paired_all)
+  apply(simp add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
+  apply(simp add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
+  using q1 q2
+  unfolding Quotient_def
+  apply(blast)
+  done
+
+lemma Pair_rsp[quot_respect]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
+  by simp
+
+lemma Pair_prs[quot_preserve]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
+  apply(simp add: expand_fun_eq)
+  apply(simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
+  done
+
+lemma fst_rsp[quot_respect]:
+  assumes "Quotient R1 Abs1 Rep1"
+  assumes "Quotient R2 Abs2 Rep2"
+  shows "(prod_rel R1 R2 ===> R1) fst fst"
+  by simp
+
+lemma fst_prs[quot_preserve]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst"
+  apply(simp add: expand_fun_eq)
+  apply(simp add: Quotient_abs_rep[OF q1])
+  done
+
+lemma snd_rsp[quot_respect]:
+  assumes "Quotient R1 Abs1 Rep1"
+  assumes "Quotient R2 Abs2 Rep2"
+  shows "(prod_rel R1 R2 ===> R2) snd snd"
+  by simp
+
+lemma snd_prs[quot_preserve]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd"
+  apply(simp add: expand_fun_eq)
+  apply(simp add: Quotient_abs_rep[OF q2])
+  done
+
+lemma split_rsp[quot_respect]:
+  shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split"
+  by auto
+
+lemma split_prs[quot_preserve]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  and     q2: "Quotient R2 Abs2 Rep2"
+  shows "(((Abs1 ---> Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> id) split) = split"
+  by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
+
+lemma prod_fun_id[id_simps]:
+  shows "prod_fun id id = id"
+  by (simp add: prod_fun_def)
+
+lemma prod_rel_eq[id_simps]:
+  shows "prod_rel (op =) (op =) = (op =)"
+  by (simp add: expand_fun_eq)
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Quotient_Sum.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,96 @@
+(*  Title:      Quotient_Sum.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+*)
+theory Quotient_Sum
+imports Quotient Quotient_Syntax
+begin
+
+section {* Quotient infrastructure for the sum type. *}
+
+fun
+  sum_rel
+where
+  "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1"
+| "sum_rel R1 R2 (Inl a1) (Inr b2) = False"
+| "sum_rel R1 R2 (Inr a2) (Inl b1) = False"
+| "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2"
+
+fun
+  sum_map
+where
+  "sum_map f1 f2 (Inl a) = Inl (f1 a)"
+| "sum_map f1 f2 (Inr a) = Inr (f2 a)"
+
+declare [[map "+" = (sum_map, sum_rel)]]
+
+
+text {* should probably be in Sum_Type.thy *}
+lemma split_sum_all:
+  shows "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (Inl x)) \<and> (\<forall>x. P (Inr x))"
+  apply(auto)
+  apply(case_tac x)
+  apply(simp_all)
+  done
+
+lemma sum_equivp[quot_equiv]:
+  assumes a: "equivp R1"
+  assumes b: "equivp R2"
+  shows "equivp (sum_rel R1 R2)"
+  apply(rule equivpI)
+  unfolding reflp_def symp_def transp_def
+  apply(simp_all add: split_sum_all)
+  apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
+  apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
+  apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
+  done
+
+lemma sum_quotient[quot_thm]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "Quotient (sum_rel R1 R2) (sum_map Abs1 Abs2) (sum_map Rep1 Rep2)"
+  unfolding Quotient_def
+  apply(simp add: split_sum_all)
+  apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
+  apply(simp_all add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
+  using q1 q2
+  unfolding Quotient_def
+  apply(blast)+
+  done
+
+lemma sum_Inl_rsp[quot_respect]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(R1 ===> sum_rel R1 R2) Inl Inl"
+  by simp
+
+lemma sum_Inr_rsp[quot_respect]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(R2 ===> sum_rel R1 R2) Inr Inr"
+  by simp
+
+lemma sum_Inl_prs[quot_preserve]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(Rep1 ---> sum_map Abs1 Abs2) Inl = Inl"
+  apply(simp add: expand_fun_eq)
+  apply(simp add: Quotient_abs_rep[OF q1])
+  done
+
+lemma sum_Inr_prs[quot_preserve]:
+  assumes q1: "Quotient R1 Abs1 Rep1"
+  assumes q2: "Quotient R2 Abs2 Rep2"
+  shows "(Rep2 ---> sum_map Abs1 Abs2) Inr = Inr"
+  apply(simp add: expand_fun_eq)
+  apply(simp add: Quotient_abs_rep[OF q2])
+  done
+
+lemma sum_map_id[id_simps]:
+  shows "sum_map id id = id"
+  by (simp add: expand_fun_eq split_sum_all)
+
+lemma sum_rel_eq[id_simps]:
+  shows "sum_rel (op =) (op =) = (op =)"
+  by (simp add: expand_fun_eq split_sum_all)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/Quotient_Syntax.thy	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,18 @@
+(*  Title:      Quotient_Syntax.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+*)
+
+header {* Pretty syntax for Quotient operations *}
+
+(*<*)
+theory Quotient_Syntax
+imports Quotient
+begin
+
+notation
+  rel_conj (infixr "OOO" 75) and
+  fun_map (infixr "--->" 55) and
+  fun_rel (infixr "===>" 55)
+
+end
+(*>*)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/ROOT.ML	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,15 @@
+quick_and_dirty := true;
+
+no_document use_thys
+   ["Quotient",
+    "Examples/AbsRepTest",
+    "Examples/FSet",
+    "Examples/FSet2",
+    "Examples/FSet3",
+    "Examples/IntEx",
+    "Examples/IntEx2",
+    "Examples/LFex",
+    "Examples/LamEx",
+    "Examples/LarryDatatype",
+    "Examples/LarryInt",
+    "Examples/Terms"];
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/quotient_def.ML	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,110 @@
+(*  Title:      quotient_def.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+
+    Definitions for constants on quotient types.
+
+*)
+
+signature QUOTIENT_DEF =
+sig
+  val quotient_def: (binding option * mixfix) * (Attrib.binding * (term * term)) ->
+    local_theory -> (term * thm) * local_theory
+
+  val quotdef_cmd: (binding option * mixfix) * (Attrib.binding * (string * string)) ->
+    local_theory -> (term * thm) * local_theory
+
+  val quotient_lift_const: string * term -> local_theory -> (term * thm) * local_theory
+end;
+
+structure Quotient_Def: QUOTIENT_DEF =
+struct
+
+open Quotient_Info;
+open Quotient_Term;
+
+(** Interface and Syntax Setup **)
+
+(* The ML-interface for a quotient definition takes
+   as argument:
+
+    - an optional binding and mixfix annotation
+    - attributes
+    - the new constant as term
+    - the rhs of the definition as term
+
+   It returns the defined constant and its definition
+   theorem; stores the data in the qconsts data slot.
+
+   Restriction: At the moment the right-hand side of the
+   definition must be a constant. Similarly the left-hand 
+   side must be a constant.
+*)
+fun error_msg bind str = 
+let 
+  val name = Binding.name_of bind
+  val pos = Position.str_of (Binding.pos_of bind)
+in
+  error ("Head of quotient_definition " ^ 
+    (quote str) ^ " differs from declaration " ^ name ^ pos)
+end
+
+fun quotient_def ((optbind, mx), (attr, (lhs, rhs))) lthy =
+let
+  val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined."
+  val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction"
+  
+  fun sanity_test NONE _ = true
+    | sanity_test (SOME bind) str =
+        if Name.of_binding bind = str then true
+        else error_msg bind str
+
+  val _ = sanity_test optbind lhs_str
+
+  val qconst_bname = Binding.name lhs_str
+  val absrep_trm = absrep_fun AbsF lthy (fastype_of rhs, lhs_ty) $ rhs
+  val prop = Logic.mk_equals (lhs, Syntax.check_term lthy absrep_trm)
+  val (_, prop') = LocalDefs.cert_def lthy prop
+  val (_, newrhs) = Primitive_Defs.abs_def prop'
+
+  val ((trm, (_ , thm)), lthy') = Local_Theory.define ((qconst_bname, mx), (attr, newrhs)) lthy
+
+  (* data storage *)
+  fun qcinfo phi = transform_qconsts phi {qconst = trm, rconst = rhs, def = thm}
+  fun trans_name phi = (fst o dest_Const o #qconst) (qcinfo phi)
+  val lthy'' = Local_Theory.declaration true
+                 (fn phi => qconsts_update_gen (trans_name phi) (qcinfo phi)) lthy'
+in
+  ((trm, thm), lthy'')
+end
+
+fun quotdef_cmd (decl, (attr, (lhs_str, rhs_str))) lthy =
+let
+  val lhs = Syntax.read_term lthy lhs_str
+  val rhs = Syntax.read_term lthy rhs_str
+  val lthy' = Variable.declare_term lhs lthy
+  val lthy'' = Variable.declare_term rhs lthy'
+in
+  quotient_def (decl, (attr, (lhs, rhs))) lthy''
+end
+
+fun quotient_lift_const (b, t) ctxt =
+  quotient_def ((NONE, NoSyn), (Attrib.empty_binding,
+    (Quotient_Term.quotient_lift_const (b, t) ctxt, t))) ctxt
+
+local
+  structure P = OuterParse;
+in
+
+val quotdef_decl = (P.binding >> SOME) -- P.opt_mixfix' --| P.$$$ "where"
+
+val quotdef_parser =
+  Scan.optional quotdef_decl (NONE, NoSyn) -- 
+    P.!!! (SpecParse.opt_thm_name ":" -- (P.term --| P.$$$ "is" -- P.term))
+end
+
+val _ =
+  OuterSyntax.local_theory "quotient_definition"
+    "definition for constants over the quotient type"
+      OuterKeyword.thy_decl (quotdef_parser >> (snd oo quotdef_cmd))
+
+end; (* structure *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/quotient_info.ML	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,293 @@
+(*  Title:      quotient_info.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+
+    Data slots for the quotient package.
+
+*)
+
+
+signature QUOTIENT_INFO =
+sig
+  exception NotFound
+
+  type maps_info = {mapfun: string, relmap: string}
+  val maps_defined: theory -> string -> bool
+  val maps_lookup: theory -> string -> maps_info     (* raises NotFound *)
+  val maps_update_thy: string -> maps_info -> theory -> theory
+  val maps_update: string -> maps_info -> Proof.context -> Proof.context
+  val print_mapsinfo: Proof.context -> unit
+
+  type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
+  val transform_quotdata: morphism -> quotdata_info -> quotdata_info
+  val quotdata_lookup_raw: theory -> string -> quotdata_info option
+  val quotdata_lookup: theory -> string -> quotdata_info     (* raises NotFound *)
+  val quotdata_update_thy: string -> quotdata_info -> theory -> theory
+  val quotdata_update_gen: string -> quotdata_info -> Context.generic -> Context.generic
+  val quotdata_dest: Proof.context -> quotdata_info list
+  val print_quotinfo: Proof.context -> unit
+
+  type qconsts_info = {qconst: term, rconst: term, def: thm}
+  val transform_qconsts: morphism -> qconsts_info -> qconsts_info
+  val qconsts_lookup: theory -> term -> qconsts_info     (* raises NotFound *)
+  val qconsts_update_thy: string -> qconsts_info -> theory -> theory
+  val qconsts_update_gen: string -> qconsts_info -> Context.generic -> Context.generic
+  val qconsts_dest: Proof.context -> qconsts_info list
+  val print_qconstinfo: Proof.context -> unit
+
+  val equiv_rules_get: Proof.context -> thm list
+  val equiv_rules_add: attribute
+  val rsp_rules_get: Proof.context -> thm list
+  val rsp_rules_add: attribute
+  val prs_rules_get: Proof.context -> thm list
+  val prs_rules_add: attribute
+  val id_simps_get: Proof.context -> thm list
+  val quotient_rules_get: Proof.context -> thm list
+  val quotient_rules_add: attribute
+end;
+
+
+structure Quotient_Info: QUOTIENT_INFO =
+struct
+
+exception NotFound
+
+
+(** data containers **)
+
+(* info about map- and rel-functions for a type *)
+type maps_info = {mapfun: string, relmap: string}
+
+structure MapsData = Theory_Data
+  (type T = maps_info Symtab.table
+   val empty = Symtab.empty
+   val extend = I
+   val merge = Symtab.merge (K true))
+
+fun maps_defined thy s =
+  Symtab.defined (MapsData.get thy) s
+
+fun maps_lookup thy s =
+  case (Symtab.lookup (MapsData.get thy) s) of
+    SOME map_fun => map_fun
+  | NONE => raise NotFound
+
+fun maps_update_thy k minfo = MapsData.map (Symtab.update (k, minfo))
+fun maps_update k minfo = ProofContext.theory (maps_update_thy k minfo)
+
+fun maps_attribute_aux s minfo = Thm.declaration_attribute
+  (fn _ => Context.mapping (maps_update_thy s minfo) (maps_update s minfo))
+
+(* attribute to be used in declare statements *)
+fun maps_attribute (ctxt, (tystr, (mapstr, relstr))) =
+let
+  val thy = ProofContext.theory_of ctxt
+  val tyname = Sign.intern_type thy tystr
+  val mapname = Sign.intern_const thy mapstr
+  val relname = Sign.intern_const thy relstr
+
+  fun sanity_check s = (Const (s, dummyT) |> Syntax.check_term ctxt; ())
+  val _ = List.app sanity_check [mapname, relname]
+in
+  maps_attribute_aux tyname {mapfun = mapname, relmap = relname}
+end
+
+val maps_attr_parser =
+  Args.context -- Scan.lift
+    ((Args.name --| OuterParse.$$$ "=") --
+      (OuterParse.$$$ "(" |-- Args.name --| OuterParse.$$$ "," --
+        Args.name --| OuterParse.$$$ ")"))
+
+val _ = Context.>> (Context.map_theory
+  (Attrib.setup @{binding "map"} (maps_attr_parser >> maps_attribute)
+    "declaration of map information"))
+
+fun print_mapsinfo ctxt =
+let
+  fun prt_map (ty_name, {mapfun, relmap}) =
+    Pretty.block (Library.separate (Pretty.brk 2)
+      (map Pretty.str
+        ["type:", ty_name,
+        "map:", mapfun,
+        "relation map:", relmap]))
+in
+  MapsData.get (ProofContext.theory_of ctxt)
+  |> Symtab.dest
+  |> map (prt_map)
+  |> Pretty.big_list "maps for type constructors:"
+  |> Pretty.writeln
+end
+
+
+(* info about quotient types *)
+type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
+
+structure QuotData = Theory_Data
+  (type T = quotdata_info Symtab.table
+   val empty = Symtab.empty
+   val extend = I
+   val merge = Symtab.merge (K true))
+
+fun transform_quotdata phi {qtyp, rtyp, equiv_rel, equiv_thm} =
+  {qtyp = Morphism.typ phi qtyp,
+   rtyp = Morphism.typ phi rtyp,
+   equiv_rel = Morphism.term phi equiv_rel,
+   equiv_thm = Morphism.thm phi equiv_thm}
+
+fun quotdata_lookup_raw thy str = Symtab.lookup (QuotData.get thy) str
+
+fun quotdata_lookup thy str =
+  case Symtab.lookup (QuotData.get thy) str of
+    SOME qinfo => qinfo
+  | NONE => raise NotFound
+
+fun quotdata_update_thy str qinfo = QuotData.map (Symtab.update (str, qinfo))
+fun quotdata_update_gen str qinfo = Context.mapping (quotdata_update_thy str qinfo) I
+
+fun quotdata_dest lthy =
+  map snd (Symtab.dest (QuotData.get (ProofContext.theory_of lthy)))
+
+fun print_quotinfo ctxt =
+let
+  fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm} =
+    Pretty.block (Library.separate (Pretty.brk 2)
+     [Pretty.str "quotient type:",
+      Syntax.pretty_typ ctxt qtyp,
+      Pretty.str "raw type:",
+      Syntax.pretty_typ ctxt rtyp,
+      Pretty.str "relation:",
+      Syntax.pretty_term ctxt equiv_rel,
+      Pretty.str "equiv. thm:",
+      Syntax.pretty_term ctxt (prop_of equiv_thm)])
+in
+  QuotData.get (ProofContext.theory_of ctxt)
+  |> Symtab.dest
+  |> map (prt_quot o snd)
+  |> Pretty.big_list "quotients:"
+  |> Pretty.writeln
+end
+
+
+(* info about quotient constants *)
+type qconsts_info = {qconst: term, rconst: term, def: thm}
+
+fun qconsts_info_eq (x : qconsts_info, y : qconsts_info) = #qconst x = #qconst y
+
+(* We need to be able to lookup instances of lifted constants,
+   for example given "nat fset" we need to find "'a fset";
+   but overloaded constants share the same name *)
+structure QConstsData = Theory_Data
+  (type T = (qconsts_info list) Symtab.table
+   val empty = Symtab.empty
+   val extend = I
+   val merge = Symtab.merge_list qconsts_info_eq)
+
+fun transform_qconsts phi {qconst, rconst, def} =
+  {qconst = Morphism.term phi qconst,
+   rconst = Morphism.term phi rconst,
+   def = Morphism.thm phi def}
+
+fun qconsts_update_thy name qcinfo = QConstsData.map (Symtab.cons_list (name, qcinfo))
+fun qconsts_update_gen name qcinfo = Context.mapping (qconsts_update_thy name qcinfo) I
+
+fun qconsts_dest lthy =
+  flat (map snd (Symtab.dest (QConstsData.get (ProofContext.theory_of lthy))))
+
+fun qconsts_lookup thy t =
+  let
+    val (name, qty) = dest_Const t
+    fun matches x =
+      let
+        val (name', qty') = dest_Const (#qconst x);
+      in
+        name = name' andalso Sign.typ_instance thy (qty, qty')
+      end
+  in
+    case Symtab.lookup (QConstsData.get thy) name of
+      NONE => raise NotFound
+    | SOME l =>
+      (case (find_first matches l) of
+        SOME x => x
+      | NONE => raise NotFound)
+  end
+
+fun print_qconstinfo ctxt =
+let
+  fun prt_qconst {qconst, rconst, def} =
+    Pretty.block (separate (Pretty.brk 1)
+     [Syntax.pretty_term ctxt qconst,
+      Pretty.str ":=",
+      Syntax.pretty_term ctxt rconst,
+      Pretty.str "as",
+      Syntax.pretty_term ctxt (prop_of def)])
+in
+  QConstsData.get (ProofContext.theory_of ctxt)
+  |> Symtab.dest
+  |> map snd
+  |> flat
+  |> map prt_qconst
+  |> Pretty.big_list "quotient constants:"
+  |> Pretty.writeln
+end
+
+(* equivalence relation theorems *)
+structure EquivRules = Named_Thms
+  (val name = "quot_equiv"
+   val description = "Equivalence relation theorems.")
+
+val equiv_rules_get = EquivRules.get
+val equiv_rules_add = EquivRules.add
+
+(* respectfulness theorems *)
+structure RspRules = Named_Thms
+  (val name = "quot_respect"
+   val description = "Respectfulness theorems.")
+
+val rsp_rules_get = RspRules.get
+val rsp_rules_add = RspRules.add
+
+(* preservation theorems *)
+structure PrsRules = Named_Thms
+  (val name = "quot_preserve"
+   val description = "Preservation theorems.")
+
+val prs_rules_get = PrsRules.get
+val prs_rules_add = PrsRules.add
+
+(* id simplification theorems *)
+structure IdSimps = Named_Thms
+  (val name = "id_simps"
+   val description = "Identity simp rules for maps.")
+
+val id_simps_get = IdSimps.get
+
+(* quotient theorems *)
+structure QuotientRules = Named_Thms
+  (val name = "quot_thm"
+   val description = "Quotient theorems.")
+
+val quotient_rules_get = QuotientRules.get
+val quotient_rules_add = QuotientRules.add
+
+(* setup of the theorem lists *)
+
+val _ = Context.>> (Context.map_theory
+  (EquivRules.setup #>
+   RspRules.setup #>
+   PrsRules.setup #>
+   IdSimps.setup #>
+   QuotientRules.setup))
+
+(* setup of the printing commands *)
+
+fun improper_command (pp_fn, cmd_name, descr_str) =
+  OuterSyntax.improper_command cmd_name descr_str
+    OuterKeyword.diag (Scan.succeed (Toplevel.keep (pp_fn o Toplevel.context_of)))
+
+val _ = map improper_command
+  [(print_mapsinfo, "print_maps", "prints out all map functions"),
+   (print_quotinfo, "print_quotients", "prints out all quotients"),
+   (print_qconstinfo, "print_quotconsts", "prints out all quotient constants")]
+
+
+end; (* structure *)
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/quotient_tacs.ML	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,665 @@
+(*  Title:      quotient_tacs.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+
+    Tactics for solving goal arising from lifting
+    theorems to quotient types.
+*)
+
+signature QUOTIENT_TACS =
+sig
+  val regularize_tac: Proof.context -> int -> tactic
+  val injection_tac: Proof.context -> int -> tactic
+  val all_injection_tac: Proof.context -> int -> tactic
+  val clean_tac: Proof.context -> int -> tactic
+  val procedure_tac: Proof.context -> thm -> int -> tactic
+  val lift_tac: Proof.context -> thm list -> int -> tactic
+  val quotient_tac: Proof.context -> int -> tactic
+  val quot_true_tac: Proof.context -> (term -> term) -> int -> tactic
+  val lifted_attrib: attribute
+end;
+
+structure Quotient_Tacs: QUOTIENT_TACS =
+struct
+
+open Quotient_Info;
+open Quotient_Term;
+
+
+(** various helper fuctions **)
+
+(* Since HOL_basic_ss is too "big" for us, we *)
+(* need to set up our own minimal simpset.    *)
+fun mk_minimal_ss ctxt =
+  Simplifier.context ctxt empty_ss
+    setsubgoaler asm_simp_tac
+    setmksimps (mksimps [])
+
+(* composition of two theorems, used in maps *)
+fun OF1 thm1 thm2 = thm2 RS thm1
+
+(* prints a warning, if the subgoal is not solved *)
+fun WARN (tac, msg) i st =
+ case Seq.pull (SOLVED' tac i st) of
+     NONE    => (warning msg; Seq.single st)
+   | seqcell => Seq.make (fn () => seqcell)
+
+fun RANGE_WARN tacs = RANGE (map WARN tacs)
+
+fun atomize_thm thm =
+let
+  val thm' = Thm.freezeT (forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? *)
+  val thm'' = ObjectLogic.atomize (cprop_of thm')
+in
+  @{thm equal_elim_rule1} OF [thm'', thm']
+end
+
+
+
+(*** Regularize Tactic ***)
+
+(** solvers for equivp and quotient assumptions **)
+
+fun equiv_tac ctxt =
+  REPEAT_ALL_NEW (resolve_tac (equiv_rules_get ctxt))
+
+fun equiv_solver_tac ss = equiv_tac (Simplifier.the_context ss)
+val equiv_solver = Simplifier.mk_solver' "Equivalence goal solver" equiv_solver_tac
+
+fun quotient_tac ctxt =
+  (REPEAT_ALL_NEW (FIRST'
+    [rtac @{thm identity_quotient},
+     resolve_tac (quotient_rules_get ctxt)]))
+
+fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss)
+val quotient_solver =
+  Simplifier.mk_solver' "Quotient goal solver" quotient_solver_tac
+
+fun solve_quotient_assm ctxt thm =
+  case Seq.pull (quotient_tac ctxt 1 thm) of
+    SOME (t, _) => t
+  | _ => error "Solve_quotient_assm failed. Possibly a quotient theorem is missing."
+
+
+fun prep_trm thy (x, (T, t)) =
+  (cterm_of thy (Var (x, T)), cterm_of thy t)
+
+fun prep_ty thy (x, (S, ty)) =
+  (ctyp_of thy (TVar (x, S)), ctyp_of thy ty)
+
+fun get_match_inst thy pat trm =
+let
+  val univ = Unify.matchers thy [(pat, trm)]
+  val SOME (env, _) = Seq.pull univ             (* raises BIND, if no unifier *)
+  val tenv = Vartab.dest (Envir.term_env env)
+  val tyenv = Vartab.dest (Envir.type_env env)
+in
+  (map (prep_ty thy) tyenv, map (prep_trm thy) tenv)
+end
+
+(* Calculates the instantiations for the lemmas:
+
+      ball_reg_eqv_range and bex_reg_eqv_range
+
+   Since the left-hand-side contains a non-pattern '?P (f ?x)'
+   we rely on unification/instantiation to check whether the
+   theorem applies and return NONE if it doesn't.
+*)
+fun calculate_inst ctxt ball_bex_thm redex R1 R2 =
+let
+  val thy = ProofContext.theory_of ctxt
+  fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm))
+  val ty_inst = map (SOME o ctyp_of thy) [domain_type (fastype_of R2)]
+  val trm_inst = map (SOME o cterm_of thy) [R2, R1]
+in
+  case try (Drule.instantiate' ty_inst trm_inst) ball_bex_thm of
+    NONE => NONE
+  | SOME thm' =>
+      (case try (get_match_inst thy (get_lhs thm')) redex of
+        NONE => NONE
+      | SOME inst2 => try (Drule.instantiate inst2) thm')
+end
+
+fun ball_bex_range_simproc ss redex =
+let
+  val ctxt = Simplifier.the_context ss
+in
+  case redex of
+    (Const (@{const_name "Ball"}, _) $ (Const (@{const_name "Respects"}, _) $
+      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
+        calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2
+
+  | (Const (@{const_name "Bex"}, _) $ (Const (@{const_name "Respects"}, _) $
+      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
+        calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2
+
+  | _ => NONE
+end
+
+(* Regularize works as follows:
+
+  0. preliminary simplification step according to
+     ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range
+
+  1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left)
+
+  2. monos
+
+  3. commutation rules for ball and bex (ball_all_comm bex_ex_comm)
+
+  4. then rel-equalities, which need to be instantiated with 'eq_imp_rel'
+     to avoid loops
+
+  5. then simplification like 0
+
+  finally jump back to 1
+*)
+
+fun regularize_tac ctxt =
+let
+  val thy = ProofContext.theory_of ctxt
+  val ball_pat = @{term "Ball (Respects (R1 ===> R2)) P"}
+  val bex_pat  = @{term "Bex (Respects (R1 ===> R2)) P"}
+  val simproc = Simplifier.simproc_i thy "" [ball_pat, bex_pat] (K (ball_bex_range_simproc))
+  val simpset = (mk_minimal_ss ctxt)
+                       addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp}
+                       addsimprocs [simproc]
+                       addSolver equiv_solver addSolver quotient_solver
+  val eq_imp_rel = @{lemma "equivp R ==> a = b --> R a b" by (simp add: equivp_reflp)}
+  val eq_eqvs = map (OF1 eq_imp_rel) (equiv_rules_get ctxt)
+in
+  simp_tac simpset THEN'
+  REPEAT_ALL_NEW (CHANGED o FIRST'
+    [resolve_tac @{thms ball_reg_right bex_reg_left bex1_bexeq_reg},
+     resolve_tac (Inductive.get_monos ctxt),
+     resolve_tac @{thms ball_all_comm bex_ex_comm},
+     resolve_tac eq_eqvs,
+     simp_tac simpset])
+end
+
+
+
+(*** Injection Tactic ***)
+
+(* Looks for Quot_True assumptions, and in case its parameter
+   is an application, it returns the function and the argument.
+*)
+fun find_qt_asm asms =
+let
+  fun find_fun trm =
+    case trm of
+      (Const(@{const_name Trueprop}, _) $ (Const (@{const_name Quot_True}, _) $ _)) => true
+    | _ => false
+in
+ case find_first find_fun asms of
+   SOME (_ $ (_ $ (f $ a))) => SOME (f, a)
+ | _ => NONE
+end
+
+fun quot_true_simple_conv ctxt fnctn ctrm =
+  case (term_of ctrm) of
+    (Const (@{const_name Quot_True}, _) $ x) =>
+    let
+      val fx = fnctn x;
+      val thy = ProofContext.theory_of ctxt;
+      val cx = cterm_of thy x;
+      val cfx = cterm_of thy fx;
+      val cxt = ctyp_of thy (fastype_of x);
+      val cfxt = ctyp_of thy (fastype_of fx);
+      val thm = Drule.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp}
+    in
+      Conv.rewr_conv thm ctrm
+    end
+
+fun quot_true_conv ctxt fnctn ctrm =
+  case (term_of ctrm) of
+    (Const (@{const_name Quot_True}, _) $ _) =>
+      quot_true_simple_conv ctxt fnctn ctrm
+  | _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm
+  | Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm
+  | _ => Conv.all_conv ctrm
+
+fun quot_true_tac ctxt fnctn =
+   CONVERSION
+    ((Conv.params_conv ~1 (fn ctxt =>
+       (Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt)
+
+fun dest_comb (f $ a) = (f, a)
+fun dest_bcomb ((_ $ l) $ r) = (l, r)
+
+fun unlam t =
+  case t of
+    (Abs a) => snd (Term.dest_abs a)
+  | _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0)))
+
+fun dest_fun_type (Type("fun", [T, S])) = (T, S)
+  | dest_fun_type _ = error "dest_fun_type"
+
+val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl
+
+(* We apply apply_rsp only in case if the type needs lifting.
+   This is the case if the type of the data in the Quot_True
+   assumption is different from the corresponding type in the goal.
+*)
+val apply_rsp_tac =
+  Subgoal.FOCUS (fn {concl, asms, context,...} =>
+  let
+    val bare_concl = HOLogic.dest_Trueprop (term_of concl)
+    val qt_asm = find_qt_asm (map term_of asms)
+  in
+    case (bare_concl, qt_asm) of
+      (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) =>
+         if fastype_of qt_fun = fastype_of f
+         then no_tac
+         else
+           let
+             val ty_x = fastype_of x
+             val ty_b = fastype_of qt_arg
+             val ty_f = range_type (fastype_of f)
+             val thy = ProofContext.theory_of context
+             val ty_inst = map (SOME o (ctyp_of thy)) [ty_x, ty_b, ty_f]
+             val t_inst = map (SOME o (cterm_of thy)) [R2, f, g, x, y];
+             val inst_thm = Drule.instantiate' ty_inst
+               ([NONE, NONE, NONE] @ t_inst) @{thm apply_rsp}
+           in
+             (rtac inst_thm THEN' quotient_tac context) 1
+           end
+    | _ => no_tac
+  end)
+
+(* Instantiates and applies 'equals_rsp'. Since the theorem is
+   complex we rely on instantiation to tell us if it applies
+*)
+fun equals_rsp_tac R ctxt =
+let
+  val thy = ProofContext.theory_of ctxt
+in
+  case try (cterm_of thy) R of (* There can be loose bounds in R *)
+    SOME ctm =>
+      let
+        val ty = domain_type (fastype_of R)
+      in
+        case try (Drule.instantiate' [SOME (ctyp_of thy ty)]
+          [SOME (cterm_of thy R)]) @{thm equals_rsp} of
+          SOME thm => rtac thm THEN' quotient_tac ctxt
+        | NONE => K no_tac
+      end
+  | _ => K no_tac
+end
+
+fun rep_abs_rsp_tac ctxt =
+  SUBGOAL (fn (goal, i) =>
+    case (try bare_concl goal) of
+      SOME (rel $ _ $ (rep $ (abs $ _))) =>
+        let
+          val thy = ProofContext.theory_of ctxt;
+          val (ty_a, ty_b) = dest_fun_type (fastype_of abs);
+          val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b];
+        in
+          case try (map (SOME o (cterm_of thy))) [rel, abs, rep] of
+            SOME t_inst =>
+              (case try (Drule.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
+                SOME inst_thm => (rtac inst_thm THEN' quotient_tac ctxt) i
+              | NONE => no_tac)
+          | NONE => no_tac
+        end
+    | _ => no_tac)
+
+
+
+(* Injection means to prove that the regularised theorem implies
+   the abs/rep injected one.
+
+   The deterministic part:
+    - remove lambdas from both sides
+    - prove Ball/Bex/Babs equalities using ball_rsp, bex_rsp, babs_rsp
+    - prove Ball/Bex relations unfolding fun_rel_id
+    - reflexivity of equality
+    - prove equality of relations using equals_rsp
+    - use user-supplied RSP theorems
+    - solve 'relation of relations' goals using quot_rel_rsp
+    - remove rep_abs from the right side
+      (Lambdas under respects may have left us some assumptions)
+
+   Then in order:
+    - split applications of lifted type (apply_rsp)
+    - split applications of non-lifted type (cong_tac)
+    - apply extentionality
+    - assumption
+    - reflexivity of the relation
+*)
+fun injection_match_tac ctxt = SUBGOAL (fn (goal, i) =>
+(case (bare_concl goal) of
+    (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *)
+  (Const (@{const_name fun_rel}, _) $ _ $ _) $ (Abs _) $ (Abs _)
+      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
+
+    (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
+| (Const (@{const_name "op ="},_) $
+    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
+      => rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
+
+    (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *)
+| (Const (@{const_name fun_rel}, _) $ _ $ _) $
+    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
+      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
+
+    (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
+| Const (@{const_name "op ="},_) $
+    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
+      => rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
+
+    (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *)
+| (Const (@{const_name fun_rel}, _) $ _ $ _) $
+    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
+      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
+
+| (Const (@{const_name fun_rel}, _) $ _ $ _) $
+    (Const(@{const_name Bex1_rel},_) $ _) $ (Const(@{const_name Bex1_rel},_) $ _)
+      => rtac @{thm bex1_rel_rsp} THEN' quotient_tac ctxt
+
+| (_ $
+    (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
+    (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
+      => rtac @{thm babs_rsp} THEN' RANGE [quotient_tac ctxt]
+
+| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
+   (rtac @{thm refl} ORELSE'
+    (equals_rsp_tac R ctxt THEN' RANGE [
+       quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
+
+    (* reflexivity of operators arising from Cong_tac *)
+| Const (@{const_name "op ="},_) $ _ $ _ => rtac @{thm refl}
+
+   (* respectfulness of constants; in particular of a simple relation *)
+| _ $ (Const _) $ (Const _)  (* fun_rel, list_rel, etc but not equality *)
+    => resolve_tac (rsp_rules_get ctxt) THEN_ALL_NEW quotient_tac ctxt
+
+    (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
+    (* observe fun_map *)
+| _ $ _ $ _
+    => (rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
+       ORELSE' rep_abs_rsp_tac ctxt
+
+| _ => K no_tac
+) i)
+
+fun injection_step_tac ctxt rel_refl =
+ FIRST' [
+    injection_match_tac ctxt,
+
+    (* R (t $ ...) (t' $ ...) ----> apply_rsp   provided type of t needs lifting *)
+    apply_rsp_tac ctxt THEN'
+                 RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)],
+
+    (* (op =) (t $ ...) (t' $ ...) ----> Cong   provided type of t does not need lifting *)
+    (* merge with previous tactic *)
+    Cong_Tac.cong_tac @{thm cong} THEN'
+                 RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)],
+
+    (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *)
+    rtac @{thm ext} THEN' quot_true_tac ctxt unlam,
+
+    (* resolving with R x y assumptions *)
+    atac,
+
+    (* reflexivity of the basic relations *)
+    (* R ... ... *)
+    resolve_tac rel_refl]
+
+fun injection_tac ctxt =
+let
+  val rel_refl = map (OF1 @{thm equivp_reflp}) (equiv_rules_get ctxt)
+in
+  injection_step_tac ctxt rel_refl
+end
+
+fun all_injection_tac ctxt =
+  REPEAT_ALL_NEW (injection_tac ctxt)
+
+
+
+(*** Cleaning of the Theorem ***)
+
+(* expands all fun_maps, except in front of the (bound) variables listed in xs *)
+fun fun_map_simple_conv xs ctrm =
+  case (term_of ctrm) of
+    ((Const (@{const_name "fun_map"}, _) $ _ $ _) $ h $ _) =>
+        if member (op=) xs h
+        then Conv.all_conv ctrm
+        else Conv.rewr_conv @{thm fun_map_def[THEN eq_reflection]} ctrm
+  | _ => Conv.all_conv ctrm
+
+fun fun_map_conv xs ctxt ctrm =
+  case (term_of ctrm) of
+      _ $ _ => (Conv.comb_conv (fun_map_conv xs ctxt) then_conv
+                fun_map_simple_conv xs) ctrm
+    | Abs _ => Conv.abs_conv (fn (x, ctxt) => fun_map_conv ((term_of x)::xs) ctxt) ctxt ctrm
+    | _ => Conv.all_conv ctrm
+
+fun fun_map_tac ctxt = CONVERSION (fun_map_conv [] ctxt)
+
+(* custom matching functions *)
+fun mk_abs u i t =
+  if incr_boundvars i u aconv t then Bound i else
+  case t of
+    t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2
+  | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t')
+  | Bound j => if i = j then error "make_inst" else t
+  | _ => t
+
+fun make_inst lhs t =
+let
+  val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
+  val _ $ (Abs (_, _, (_ $ g))) = t;
+in
+  (f, Abs ("x", T, mk_abs u 0 g))
+end
+
+fun make_inst_id lhs t =
+let
+  val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
+  val _ $ (Abs (_, _, g)) = t;
+in
+  (f, Abs ("x", T, mk_abs u 0 g))
+end
+
+(* Simplifies a redex using the 'lambda_prs' theorem.
+   First instantiates the types and known subterms.
+   Then solves the quotient assumptions to get Rep2 and Abs1
+   Finally instantiates the function f using make_inst
+   If Rep2 is an identity then the pattern is simpler and
+   make_inst_id is used
+*)
+fun lambda_prs_simple_conv ctxt ctrm =
+  case (term_of ctrm) of
+    (Const (@{const_name fun_map}, _) $ r1 $ a2) $ (Abs _) =>
+      let
+        val thy = ProofContext.theory_of ctxt
+        val (ty_b, ty_a) = dest_fun_type (fastype_of r1)
+        val (ty_c, ty_d) = dest_fun_type (fastype_of a2)
+        val tyinst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c, ty_d]
+        val tinst = [NONE, NONE, SOME (cterm_of thy r1), NONE, SOME (cterm_of thy a2)]
+        val thm1 = Drule.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]}
+        val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1)
+        val thm3 = MetaSimplifier.rewrite_rule @{thms id_apply[THEN eq_reflection]} thm2
+        val (insp, inst) =
+          if ty_c = ty_d
+          then make_inst_id (term_of (Thm.lhs_of thm3)) (term_of ctrm)
+          else make_inst (term_of (Thm.lhs_of thm3)) (term_of ctrm)
+        val thm4 = Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) thm3
+      in
+        Conv.rewr_conv thm4 ctrm
+      end
+  | _ => Conv.all_conv ctrm
+
+fun lambda_prs_conv ctxt = More_Conv.top_conv lambda_prs_simple_conv ctxt
+fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt)
+
+
+(* Cleaning consists of:
+
+  1. unfolding of ---> in front of everything, except
+     bound variables (this prevents lambda_prs from
+     becoming stuck)
+
+  2. simplification with lambda_prs
+
+  3. simplification with:
+
+      - Quotient_abs_rep Quotient_rel_rep
+        babs_prs all_prs ex_prs ex1_prs
+
+      - id_simps and preservation lemmas and
+
+      - symmetric versions of the definitions
+        (that is definitions of quotient constants
+         are folded)
+
+  4. test for refl
+*)
+fun clean_tac lthy =
+let
+  val defs = map (symmetric o #def) (qconsts_dest lthy)
+  val prs = prs_rules_get lthy
+  val ids = id_simps_get lthy
+  val thms = @{thms Quotient_abs_rep Quotient_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs
+
+  val ss = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver
+in
+  EVERY' [fun_map_tac lthy,
+          lambda_prs_tac lthy,
+          simp_tac ss,
+          TRY o rtac refl]
+end
+
+
+
+(** Tactic for Generalising Free Variables in a Goal **)
+
+fun inst_spec ctrm =
+   Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}
+
+fun inst_spec_tac ctrms =
+  EVERY' (map (dtac o inst_spec) ctrms)
+
+fun all_list xs trm =
+  fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm
+
+fun apply_under_Trueprop f =
+  HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop
+
+fun gen_frees_tac ctxt =
+  SUBGOAL (fn (concl, i) =>
+    let
+      val thy = ProofContext.theory_of ctxt
+      val vrs = Term.add_frees concl []
+      val cvrs = map (cterm_of thy o Free) vrs
+      val concl' = apply_under_Trueprop (all_list vrs) concl
+      val goal = Logic.mk_implies (concl', concl)
+      val rule = Goal.prove ctxt [] [] goal
+        (K (EVERY1 [inst_spec_tac (rev cvrs), atac]))
+    in
+      rtac rule i
+    end)
+
+
+(** The General Shape of the Lifting Procedure **)
+
+(* - A is the original raw theorem
+   - B is the regularized theorem
+   - C is the rep/abs injected version of B
+   - D is the lifted theorem
+
+   - 1st prem is the regularization step
+   - 2nd prem is the rep/abs injection step
+   - 3rd prem is the cleaning part
+
+   the Quot_True premise in 2nd records the lifted theorem
+*)
+val lifting_procedure_thm =
+  @{lemma  "[|A;
+              A --> B;
+              Quot_True D ==> B = C;
+              C = D|] ==> D"
+      by (simp add: Quot_True_def)}
+
+fun lift_match_error ctxt msg rtrm qtrm =
+let
+  val rtrm_str = Syntax.string_of_term ctxt rtrm
+  val qtrm_str = Syntax.string_of_term ctxt qtrm
+  val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str,
+    "", "does not match with original theorem", rtrm_str]
+in
+  error msg
+end
+
+fun procedure_inst ctxt rtrm qtrm =
+let
+  val thy = ProofContext.theory_of ctxt
+  val rtrm' = HOLogic.dest_Trueprop rtrm
+  val qtrm' = HOLogic.dest_Trueprop qtrm
+  val reg_goal = regularize_trm_chk ctxt (rtrm', qtrm')
+    handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
+  val inj_goal = inj_repabs_trm_chk ctxt (reg_goal, qtrm')
+    handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
+in
+  Drule.instantiate' []
+    [SOME (cterm_of thy rtrm'),
+     SOME (cterm_of thy reg_goal),
+     NONE,
+     SOME (cterm_of thy inj_goal)] lifting_procedure_thm
+end
+
+(* the tactic leaves three subgoals to be proved *)
+fun procedure_tac ctxt rthm =
+  ObjectLogic.full_atomize_tac
+  THEN' gen_frees_tac ctxt
+  THEN' SUBGOAL (fn (goal, i) =>
+    let
+      val rthm' = atomize_thm rthm
+      val rule = procedure_inst ctxt (prop_of rthm') goal
+    in
+      (rtac rule THEN' rtac rthm') i
+    end)
+
+
+(* Automatic Proofs *)
+
+val msg1 = "The regularize proof failed."
+val msg2 = cat_lines ["The injection proof failed.",
+                      "This is probably due to missing respects lemmas.",
+                      "Try invoking the injection method manually to see",
+                      "which lemmas are missing."]
+val msg3 = "The cleaning proof failed."
+
+fun lift_tac ctxt rthms =
+let
+  fun mk_tac rthm =
+    procedure_tac ctxt rthm
+    THEN' RANGE_WARN
+      [(regularize_tac ctxt, msg1),
+       (all_injection_tac ctxt, msg2),
+       (clean_tac ctxt, msg3)]
+in
+  simp_tac (mk_minimal_ss ctxt) (* unfolding multiple &&& *)
+  THEN' RANGE (map mk_tac rthms)
+end
+
+(* An Attribute which automatically constructs the qthm *)
+fun lifted_attrib_aux context thm =
+let
+  val ctxt = Context.proof_of context
+  val ((_, [thm']), ctxt') = Variable.import false [thm] ctxt
+  val goal = (quotient_lift_all ctxt' o prop_of) thm'
+in
+  Goal.prove ctxt' [] [] goal (K (lift_tac ctxt' [thm] 1))
+  |> singleton (ProofContext.export ctxt' ctxt)
+end;
+
+val lifted_attrib = Thm.rule_attribute lifted_attrib_aux
+
+end; (* structure *)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/quotient_term.ML	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,786 @@
+(*  Title:      quotient_term.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+
+    Constructs terms corresponding to goals from
+    lifting theorems to quotient types.
+*)
+
+signature QUOTIENT_TERM =
+sig
+  exception LIFT_MATCH of string
+
+  datatype flag = AbsF | RepF
+
+  val absrep_fun: flag -> Proof.context -> typ * typ -> term
+  val absrep_fun_chk: flag -> Proof.context -> typ * typ -> term
+
+  (* Allows Nitpick to represent quotient types as single elements from raw type *)
+  val absrep_const_chk: flag -> Proof.context -> string -> term
+
+  val equiv_relation: Proof.context -> typ * typ -> term
+  val equiv_relation_chk: Proof.context -> typ * typ -> term
+
+  val regularize_trm: Proof.context -> term * term -> term
+  val regularize_trm_chk: Proof.context -> term * term -> term
+
+  val inj_repabs_trm: Proof.context -> term * term -> term
+  val inj_repabs_trm_chk: Proof.context -> term * term -> term
+
+  val quotient_lift_const: string * term -> local_theory -> term
+  val quotient_lift_all: Proof.context -> term -> term
+end;
+
+structure Quotient_Term: QUOTIENT_TERM =
+struct
+
+open Quotient_Info;
+
+exception LIFT_MATCH of string
+
+
+
+(*** Aggregate Rep/Abs Function ***)
+
+
+(* The flag RepF is for types in negative position; AbsF is for types
+   in positive position. Because of this, function types need to be
+   treated specially, since there the polarity changes.
+*)
+
+datatype flag = AbsF | RepF
+
+fun negF AbsF = RepF
+  | negF RepF = AbsF
+
+fun is_identity (Const (@{const_name "id"}, _)) = true
+  | is_identity _ = false
+
+fun mk_identity ty = Const (@{const_name "id"}, ty --> ty)
+
+fun mk_fun_compose flag (trm1, trm2) =
+  case flag of
+    AbsF => Const (@{const_name "comp"}, dummyT) $ trm1 $ trm2
+  | RepF => Const (@{const_name "comp"}, dummyT) $ trm2 $ trm1
+
+fun get_mapfun ctxt s =
+let
+  val thy = ProofContext.theory_of ctxt
+  val exn = LIFT_MATCH ("No map function for type " ^ quote s ^ " found.")
+  val mapfun = #mapfun (maps_lookup thy s) handle Quotient_Info.NotFound => raise exn
+in
+  Const (mapfun, dummyT)
+end
+
+(* makes a Free out of a TVar *)
+fun mk_Free (TVar ((x, i), _)) = Free (unprefix "'" x ^ string_of_int i, dummyT)
+
+(* produces an aggregate map function for the
+   rty-part of a quotient definition; abstracts
+   over all variables listed in vs (these variables
+   correspond to the type variables in rty)
+
+   for example for: (?'a list * ?'b)
+   it produces:     %a b. prod_map (map a) b
+*)
+fun mk_mapfun ctxt vs rty =
+let
+  val vs' = map (mk_Free) vs
+
+  fun mk_mapfun_aux rty =
+    case rty of
+      TVar _ => mk_Free rty
+    | Type (_, []) => mk_identity rty
+    | Type (s, tys) => list_comb (get_mapfun ctxt s, map mk_mapfun_aux tys)
+    | _ => raise LIFT_MATCH "mk_mapfun (default)"
+in
+  fold_rev Term.lambda vs' (mk_mapfun_aux rty)
+end
+
+(* looks up the (varified) rty and qty for
+   a quotient definition
+*)
+fun get_rty_qty ctxt s =
+let
+  val thy = ProofContext.theory_of ctxt
+  val exn = LIFT_MATCH ("No quotient type " ^ quote s ^ " found.")
+  val qdata = (quotdata_lookup thy s) handle Quotient_Info.NotFound => raise exn
+in
+  (#rtyp qdata, #qtyp qdata)
+end
+
+(* takes two type-environments and looks
+   up in both of them the variable v, which
+   must be listed in the environment
+*)
+fun double_lookup rtyenv qtyenv v =
+let
+  val v' = fst (dest_TVar v)
+in
+  (snd (the (Vartab.lookup rtyenv v')), snd (the (Vartab.lookup qtyenv v')))
+end
+
+(* matches a type pattern with a type *)
+fun match ctxt err ty_pat ty =
+let
+  val thy = ProofContext.theory_of ctxt
+in
+  Sign.typ_match thy (ty_pat, ty) Vartab.empty
+  handle MATCH_TYPE => err ctxt ty_pat ty
+end
+
+(* produces the rep or abs constant for a qty *)
+fun absrep_const flag ctxt qty_str =
+let
+  val thy = ProofContext.theory_of ctxt
+  val qty_name = Long_Name.base_name qty_str
+in
+  case flag of
+    AbsF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT)
+  | RepF => Const (Sign.full_bname thy ("rep_" ^ qty_name), dummyT)
+end
+
+(* Lets Nitpick represent elements of quotient types as elements of the raw type *)
+fun absrep_const_chk flag ctxt qty_str =
+  Syntax.check_term ctxt (absrep_const flag ctxt qty_str)
+
+fun absrep_match_err ctxt ty_pat ty =
+let
+  val ty_pat_str = Syntax.string_of_typ ctxt ty_pat
+  val ty_str = Syntax.string_of_typ ctxt ty
+in
+  raise LIFT_MATCH (space_implode " "
+    ["absrep_fun (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"])
+end
+
+
+(** generation of an aggregate absrep function **)
+
+(* - In case of equal types we just return the identity.
+
+   - In case of TFrees we also return the identity.
+
+   - In case of function types we recurse taking
+     the polarity change into account.
+
+   - If the type constructors are equal, we recurse for the
+     arguments and build the appropriate map function.
+
+   - If the type constructors are unequal, there must be an
+     instance of quotient types:
+
+       - we first look up the corresponding rty_pat and qty_pat
+         from the quotient definition; the arguments of qty_pat
+         must be some distinct TVars
+       - we then match the rty_pat with rty and qty_pat with qty;
+         if matching fails the types do not correspond -> error
+       - the matching produces two environments; we look up the
+         assignments for the qty_pat variables and recurse on the
+         assignments
+       - we prefix the aggregate map function for the rty_pat,
+         which is an abstraction over all type variables
+       - finally we compose the result with the appropriate
+         absrep function in case at least one argument produced
+         a non-identity function /
+         otherwise we just return the appropriate absrep
+         function
+
+     The composition is necessary for types like
+
+        ('a list) list / ('a foo) foo
+
+     The matching is necessary for types like
+
+        ('a * 'a) list / 'a bar
+
+     The test is necessary in order to eliminate superfluous
+     identity maps.
+*)
+
+fun absrep_fun flag ctxt (rty, qty) =
+  if rty = qty
+  then mk_identity rty
+  else
+    case (rty, qty) of
+      (Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) =>
+        let
+          val arg1 = absrep_fun (negF flag) ctxt (ty1, ty1')
+          val arg2 = absrep_fun flag ctxt (ty2, ty2')
+        in
+          list_comb (get_mapfun ctxt "fun", [arg1, arg2])
+        end
+    | (Type (s, tys), Type (s', tys')) =>
+        if s = s'
+        then
+           let
+             val args = map (absrep_fun flag ctxt) (tys ~~ tys')
+           in
+             list_comb (get_mapfun ctxt s, args)
+           end
+        else
+           let
+             val (rty_pat, qty_pat as Type (_, vs)) = get_rty_qty ctxt s'
+             val rtyenv = match ctxt absrep_match_err rty_pat rty
+             val qtyenv = match ctxt absrep_match_err qty_pat qty
+             val args_aux = map (double_lookup rtyenv qtyenv) vs
+             val args = map (absrep_fun flag ctxt) args_aux
+             val map_fun = mk_mapfun ctxt vs rty_pat
+             val result = list_comb (map_fun, args)
+           in
+             if forall is_identity args
+             then absrep_const flag ctxt s'
+             else mk_fun_compose flag (absrep_const flag ctxt s', result)
+           end
+    | (TFree x, TFree x') =>
+        if x = x'
+        then mk_identity rty
+        else raise (LIFT_MATCH "absrep_fun (frees)")
+    | (TVar _, TVar _) => raise (LIFT_MATCH "absrep_fun (vars)")
+    | _ => raise (LIFT_MATCH "absrep_fun (default)")
+
+fun absrep_fun_chk flag ctxt (rty, qty) =
+  absrep_fun flag ctxt (rty, qty)
+  |> Syntax.check_term ctxt
+
+
+
+
+(*** Aggregate Equivalence Relation ***)
+
+
+(* works very similar to the absrep generation,
+   except there is no need for polarities
+*)
+
+(* instantiates TVars so that the term is of type ty *)
+fun force_typ ctxt trm ty =
+let
+  val thy = ProofContext.theory_of ctxt
+  val trm_ty = fastype_of trm
+  val ty_inst = Sign.typ_match thy (trm_ty, ty) Vartab.empty
+in
+  map_types (Envir.subst_type ty_inst) trm
+end
+
+fun is_eq (Const (@{const_name "op ="}, _)) = true
+  | is_eq _ = false
+
+fun mk_rel_compose (trm1, trm2) =
+  Const (@{const_name "rel_conj"}, dummyT) $ trm1 $ trm2
+
+fun get_relmap ctxt s =
+let
+  val thy = ProofContext.theory_of ctxt
+  val exn = LIFT_MATCH ("get_relmap (no relation map function found for type " ^ s ^ ")")
+  val relmap = #relmap (maps_lookup thy s) handle Quotient_Info.NotFound => raise exn
+in
+  Const (relmap, dummyT)
+end
+
+fun mk_relmap ctxt vs rty =
+let
+  val vs' = map (mk_Free) vs
+
+  fun mk_relmap_aux rty =
+    case rty of
+      TVar _ => mk_Free rty
+    | Type (_, []) => HOLogic.eq_const rty
+    | Type (s, tys) => list_comb (get_relmap ctxt s, map mk_relmap_aux tys)
+    | _ => raise LIFT_MATCH ("mk_relmap (default)")
+in
+  fold_rev Term.lambda vs' (mk_relmap_aux rty)
+end
+
+fun get_equiv_rel ctxt s =
+let
+  val thy = ProofContext.theory_of ctxt
+  val exn = LIFT_MATCH ("get_quotdata (no quotient found for type " ^ s ^ ")")
+in
+  #equiv_rel (quotdata_lookup thy s) handle Quotient_Info.NotFound => raise exn
+end
+
+fun equiv_match_err ctxt ty_pat ty =
+let
+  val ty_pat_str = Syntax.string_of_typ ctxt ty_pat
+  val ty_str = Syntax.string_of_typ ctxt ty
+in
+  raise LIFT_MATCH (space_implode " "
+    ["equiv_relation (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"])
+end
+
+(* builds the aggregate equivalence relation
+   that will be the argument of Respects
+*)
+fun equiv_relation ctxt (rty, qty) =
+  if rty = qty
+  then HOLogic.eq_const rty
+  else
+    case (rty, qty) of
+      (Type (s, tys), Type (s', tys')) =>
+       if s = s'
+       then
+         let
+           val args = map (equiv_relation ctxt) (tys ~~ tys')
+         in
+           list_comb (get_relmap ctxt s, args)
+         end
+       else
+         let
+           val (rty_pat, qty_pat as Type (_, vs)) = get_rty_qty ctxt s'
+           val rtyenv = match ctxt equiv_match_err rty_pat rty
+           val qtyenv = match ctxt equiv_match_err qty_pat qty
+           val args_aux = map (double_lookup rtyenv qtyenv) vs
+           val args = map (equiv_relation ctxt) args_aux
+           val rel_map = mk_relmap ctxt vs rty_pat
+           val result = list_comb (rel_map, args)
+           val eqv_rel = get_equiv_rel ctxt s'
+           val eqv_rel' = force_typ ctxt eqv_rel ([rty, rty] ---> @{typ bool})
+         in
+           if forall is_eq args
+           then eqv_rel'
+           else mk_rel_compose (result, eqv_rel')
+         end
+      | _ => HOLogic.eq_const rty
+
+fun equiv_relation_chk ctxt (rty, qty) =
+  equiv_relation ctxt (rty, qty)
+  |> Syntax.check_term ctxt
+
+
+
+(*** Regularization ***)
+
+(* Regularizing an rtrm means:
+
+ - Quantifiers over types that need lifting are replaced
+   by bounded quantifiers, for example:
+
+      All P  ----> All (Respects R) P
+
+   where the aggregate relation R is given by the rty and qty;
+
+ - Abstractions over types that need lifting are replaced
+   by bounded abstractions, for example:
+
+      %x. P  ----> Ball (Respects R) %x. P
+
+ - Equalities over types that need lifting are replaced by
+   corresponding equivalence relations, for example:
+
+      A = B  ----> R A B
+
+   or
+
+      A = B  ----> (R ===> R) A B
+
+   for more complicated types of A and B
+
+
+ The regularize_trm accepts raw theorems in which equalities
+ and quantifiers match exactly the ones in the lifted theorem
+ but also accepts partially regularized terms.
+
+ This means that the raw theorems can have:
+   Ball (Respects R),  Bex (Respects R), Bex1_rel (Respects R), Babs, R
+ in the places where:
+   All, Ex, Ex1, %, (op =)
+ is required the lifted theorem.
+
+*)
+
+val mk_babs = Const (@{const_name Babs}, dummyT)
+val mk_ball = Const (@{const_name Ball}, dummyT)
+val mk_bex  = Const (@{const_name Bex}, dummyT)
+val mk_bex1_rel = Const (@{const_name Bex1_rel}, dummyT)
+val mk_resp = Const (@{const_name Respects}, dummyT)
+
+(* - applies f to the subterm of an abstraction,
+     otherwise to the given term,
+   - used by regularize, therefore abstracted
+     variables do not have to be treated specially
+*)
+fun apply_subt f (trm1, trm2) =
+  case (trm1, trm2) of
+    (Abs (x, T, t), Abs (_ , _, t')) => Abs (x, T, f (t, t'))
+  | _ => f (trm1, trm2)
+
+fun term_mismatch str ctxt t1 t2 =
+let
+  val t1_str = Syntax.string_of_term ctxt t1
+  val t2_str = Syntax.string_of_term ctxt t2
+  val t1_ty_str = Syntax.string_of_typ ctxt (fastype_of t1)
+  val t2_ty_str = Syntax.string_of_typ ctxt (fastype_of t2)
+in
+  raise LIFT_MATCH (cat_lines [str, t1_str ^ "::" ^ t1_ty_str, t2_str ^ "::" ^ t2_ty_str])
+end
+
+(* the major type of All and Ex quantifiers *)
+fun qnt_typ ty = domain_type (domain_type ty)
+
+(* Checks that two types match, for example:
+     rty -> rty   matches   qty -> qty *)
+fun matches_typ thy rT qT =
+  if rT = qT then true else
+  case (rT, qT) of
+    (Type (rs, rtys), Type (qs, qtys)) =>
+      if rs = qs then
+        if length rtys <> length qtys then false else
+        forall (fn x => x = true) (map2 (matches_typ thy) rtys qtys)
+      else
+        (case Quotient_Info.quotdata_lookup_raw thy qs of
+          SOME quotinfo => Sign.typ_instance thy (rT, #rtyp quotinfo)
+        | NONE => false)
+  | _ => false
+
+
+(* produces a regularized version of rtrm
+
+   - the result might contain dummyTs
+
+   - for regularisation we do not need any
+     special treatment of bound variables
+*)
+fun regularize_trm ctxt (rtrm, qtrm) =
+  case (rtrm, qtrm) of
+    (Abs (x, ty, t), Abs (_, ty', t')) =>
+       let
+         val subtrm = Abs(x, ty, regularize_trm ctxt (t, t'))
+       in
+         if ty = ty' then subtrm
+         else mk_babs $ (mk_resp $ equiv_relation ctxt (ty, ty')) $ subtrm
+       end
+  | (Const (@{const_name "Babs"}, T) $ resrel $ (t as (Abs (_, ty, _))), t' as (Abs (_, ty', _))) =>
+       let
+         val subtrm = regularize_trm ctxt (t, t')
+         val needres = mk_resp $ equiv_relation_chk ctxt (ty, ty')
+       in
+         if resrel <> needres
+         then term_mismatch "regularize (Babs)" ctxt resrel needres
+         else mk_babs $ resrel $ subtrm
+       end
+
+  | (Const (@{const_name "All"}, ty) $ t, Const (@{const_name "All"}, ty') $ t') =>
+       let
+         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
+       in
+         if ty = ty' then Const (@{const_name "All"}, ty) $ subtrm
+         else mk_ball $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
+       end
+
+  | (Const (@{const_name "Ex"}, ty) $ t, Const (@{const_name "Ex"}, ty') $ t') =>
+       let
+         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
+       in
+         if ty = ty' then Const (@{const_name "Ex"}, ty) $ subtrm
+         else mk_bex $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
+       end
+
+  | (Const (@{const_name "Ex1"}, ty) $ (Abs (_, _,
+      (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op :"}, _) $ _ $
+        (Const (@{const_name "Respects"}, _) $ resrel)) $ (t $ _)))),
+     Const (@{const_name "Ex1"}, ty') $ t') =>
+       let
+         val t_ = incr_boundvars (~1) t
+         val subtrm = apply_subt (regularize_trm ctxt) (t_, t')
+         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
+       in
+         if resrel <> needrel
+         then term_mismatch "regularize (Bex1)" ctxt resrel needrel
+         else mk_bex1_rel $ resrel $ subtrm
+       end
+
+  | (Const (@{const_name "Ex1"}, ty) $ t, Const (@{const_name "Ex1"}, ty') $ t') =>
+       let
+         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
+       in
+         if ty = ty' then Const (@{const_name "Ex1"}, ty) $ subtrm
+         else mk_bex1_rel $ (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
+       end
+
+  | (Const (@{const_name "Ball"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t,
+     Const (@{const_name "All"}, ty') $ t') =>
+       let
+         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
+         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
+       in
+         if resrel <> needrel
+         then term_mismatch "regularize (Ball)" ctxt resrel needrel
+         else mk_ball $ (mk_resp $ resrel) $ subtrm
+       end
+
+  | (Const (@{const_name "Bex"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t,
+     Const (@{const_name "Ex"}, ty') $ t') =>
+       let
+         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
+         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
+       in
+         if resrel <> needrel
+         then term_mismatch "regularize (Bex)" ctxt resrel needrel
+         else mk_bex $ (mk_resp $ resrel) $ subtrm
+       end
+
+  | (Const (@{const_name "Bex1_rel"}, ty) $ resrel $ t, Const (@{const_name "Ex1"}, ty') $ t') =>
+       let
+         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
+         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
+       in
+         if resrel <> needrel
+         then term_mismatch "regularize (Bex1_res)" ctxt resrel needrel
+         else mk_bex1_rel $ resrel $ subtrm
+       end
+
+  | (* equalities need to be replaced by appropriate equivalence relations *)
+    (Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) =>
+         if ty = ty' then rtrm
+         else equiv_relation ctxt (domain_type ty, domain_type ty')
+
+  | (* in this case we just check whether the given equivalence relation is correct *)
+    (rel, Const (@{const_name "op ="}, ty')) =>
+       let
+         val rel_ty = fastype_of rel
+         val rel' = equiv_relation_chk ctxt (domain_type rel_ty, domain_type ty')
+       in
+         if rel' aconv rel then rtrm
+         else term_mismatch "regularise (relation mismatch)" ctxt rel rel'
+       end
+
+  | (_, Const _) =>
+       let
+         val thy = ProofContext.theory_of ctxt
+         fun same_const (Const (s, T)) (Const (s', T')) = (s = s') andalso matches_typ thy T T'
+           | same_const _ _ = false
+       in
+         if same_const rtrm qtrm then rtrm
+         else
+           let
+             val rtrm' = #rconst (qconsts_lookup thy qtrm)
+               handle Quotient_Info.NotFound => term_mismatch "regularize(constant notfound)" ctxt rtrm qtrm
+           in
+             if Pattern.matches thy (rtrm', rtrm)
+             then rtrm else term_mismatch "regularize(constant mismatch)" ctxt rtrm qtrm
+           end
+       end
+
+  | (((t1 as Const (@{const_name "split"}, _)) $ Abs (v1, ty, Abs(v1', ty', s1))),
+     ((t2 as Const (@{const_name "split"}, _)) $ Abs (v2, _ , Abs(v2', _  , s2)))) =>
+       regularize_trm ctxt (t1, t2) $ Abs (v1, ty, Abs (v1', ty', regularize_trm ctxt (s1, s2)))
+
+  | (((t1 as Const (@{const_name "split"}, _)) $ Abs (v1, ty, s1)),
+     ((t2 as Const (@{const_name "split"}, _)) $ Abs (v2, _ , s2))) =>
+       regularize_trm ctxt (t1, t2) $ Abs (v1, ty, regularize_trm ctxt (s1, s2))
+
+  | (t1 $ t2, t1' $ t2') =>
+       regularize_trm ctxt (t1, t1') $ regularize_trm ctxt (t2, t2')
+
+  | (Bound i, Bound i') =>
+       if i = i' then rtrm
+       else raise (LIFT_MATCH "regularize (bounds mismatch)")
+
+  | _ =>
+       let
+         val rtrm_str = Syntax.string_of_term ctxt rtrm
+         val qtrm_str = Syntax.string_of_term ctxt qtrm
+       in
+         raise (LIFT_MATCH ("regularize failed (default: " ^ rtrm_str ^ "," ^ qtrm_str ^ ")"))
+       end
+
+fun regularize_trm_chk ctxt (rtrm, qtrm) =
+  regularize_trm ctxt (rtrm, qtrm)
+  |> Syntax.check_term ctxt
+
+
+
+(*** Rep/Abs Injection ***)
+
+(*
+Injection of Rep/Abs means:
+
+  For abstractions:
+
+  * If the type of the abstraction needs lifting, then we add Rep/Abs
+    around the abstraction; otherwise we leave it unchanged.
+
+  For applications:
+
+  * If the application involves a bounded quantifier, we recurse on
+    the second argument. If the application is a bounded abstraction,
+    we always put an Rep/Abs around it (since bounded abstractions
+    are assumed to always need lifting). Otherwise we recurse on both
+    arguments.
+
+  For constants:
+
+  * If the constant is (op =), we leave it always unchanged.
+    Otherwise the type of the constant needs lifting, we put
+    and Rep/Abs around it.
+
+  For free variables:
+
+  * We put a Rep/Abs around it if the type needs lifting.
+
+  Vars case cannot occur.
+*)
+
+fun mk_repabs ctxt (T, T') trm =
+  absrep_fun RepF ctxt (T, T') $ (absrep_fun AbsF ctxt (T, T') $ trm)
+
+fun inj_repabs_err ctxt msg rtrm qtrm =
+let
+  val rtrm_str = Syntax.string_of_term ctxt rtrm
+  val qtrm_str = Syntax.string_of_term ctxt qtrm
+in
+  raise LIFT_MATCH (space_implode " " [msg, quote rtrm_str, "and", quote qtrm_str])
+end
+
+
+(* bound variables need to be treated properly,
+   as the type of subterms needs to be calculated   *)
+fun inj_repabs_trm ctxt (rtrm, qtrm) =
+ case (rtrm, qtrm) of
+    (Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
+       Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
+
+  | (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
+       Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
+
+  | (Const (@{const_name "Babs"}, T) $ r $ t, t' as (Abs _)) =>
+      let
+        val rty = fastype_of rtrm
+        val qty = fastype_of qtrm
+      in
+        mk_repabs ctxt (rty, qty) (Const (@{const_name "Babs"}, T) $ r $ (inj_repabs_trm ctxt (t, t')))
+      end
+
+  | (Abs (x, T, t), Abs (x', T', t')) =>
+      let
+        val rty = fastype_of rtrm
+        val qty = fastype_of qtrm
+        val (y, s) = Term.dest_abs (x, T, t)
+        val (_, s') = Term.dest_abs (x', T', t')
+        val yvar = Free (y, T)
+        val result = Term.lambda_name (y, yvar) (inj_repabs_trm ctxt (s, s'))
+      in
+        if rty = qty then result
+        else mk_repabs ctxt (rty, qty) result
+      end
+
+  | (t $ s, t' $ s') =>
+       (inj_repabs_trm ctxt (t, t')) $ (inj_repabs_trm ctxt (s, s'))
+
+  | (Free (_, T), Free (_, T')) =>
+        if T = T' then rtrm
+        else mk_repabs ctxt (T, T') rtrm
+
+  | (_, Const (@{const_name "op ="}, _)) => rtrm
+
+  | (_, Const (_, T')) =>
+      let
+        val rty = fastype_of rtrm
+      in
+        if rty = T' then rtrm
+        else mk_repabs ctxt (rty, T') rtrm
+      end
+
+  | _ => inj_repabs_err ctxt "injection (default):" rtrm qtrm
+
+fun inj_repabs_trm_chk ctxt (rtrm, qtrm) =
+  inj_repabs_trm ctxt (rtrm, qtrm)
+  |> Syntax.check_term ctxt
+
+
+
+(*** Wrapper for automatically transforming an rthm into a qthm ***)
+
+(* subst_tys takes a list of (rty, qty) substitution pairs
+   and replaces all occurences of rty in the given type
+   by appropriate qty, with substitution *)
+fun subst_ty thy ty (rty, qty) r =
+  if r <> NONE then r else
+  case try (Sign.typ_match thy (rty, ty)) Vartab.empty of
+    SOME inst => SOME (Envir.subst_type inst qty)
+  | NONE => NONE
+fun subst_tys thy substs ty =
+  case fold (subst_ty thy ty) substs NONE of
+    SOME ty => ty
+  | NONE =>
+      (case ty of
+        Type (s, tys) => Type (s, map (subst_tys thy substs) tys)
+      | x => x)
+
+(* subst_trms takes a list of (rtrm, qtrm) substitution pairs
+   and if the given term matches any of the raw terms it
+   returns the appropriate qtrm instantiated. If none of
+   them matched it returns NONE. *)
+fun subst_trm thy t (rtrm, qtrm) s =
+  if s <> NONE then s else
+    case try (Pattern.match thy (rtrm, t)) (Vartab.empty, Vartab.empty) of
+      SOME inst => SOME (Envir.subst_term inst qtrm)
+    | NONE => NONE;
+fun subst_trms thy substs t = fold (subst_trm thy t) substs NONE
+
+(* prepares type and term substitution pairs to be used by above
+   functions that let replace all raw constructs by appropriate
+   lifted counterparts. *)
+fun get_ty_trm_substs ctxt =
+let
+  val thy = ProofContext.theory_of ctxt
+  val quot_infos  = Quotient_Info.quotdata_dest ctxt
+  val const_infos = Quotient_Info.qconsts_dest ctxt
+  val ty_substs = map (fn ri => (#rtyp ri, #qtyp ri)) quot_infos
+  val const_substs = map (fn ci => (#rconst ci, #qconst ci)) const_infos
+  fun rel_eq rel = HOLogic.eq_const (subst_tys thy ty_substs (domain_type (fastype_of rel)))
+  val rel_substs = map (fn ri => (#equiv_rel ri, rel_eq (#equiv_rel ri))) quot_infos
+in
+  (ty_substs, const_substs @ rel_substs)
+end
+
+fun quotient_lift_const (b, t) ctxt =
+let
+  val thy = ProofContext.theory_of ctxt
+  val (ty_substs, _) = get_ty_trm_substs ctxt;
+  val (_, ty) = dest_Const t;
+  val nty = subst_tys thy ty_substs ty;
+in
+  Free(b, nty)
+end
+
+(*
+Takes a term and
+
+* replaces raw constants by the quotient constants
+
+* replaces equivalence relations by equalities
+
+* replaces raw types by the quotient types
+
+*)
+
+fun quotient_lift_all ctxt t =
+let
+  val thy = ProofContext.theory_of ctxt
+  val (ty_substs, substs) = get_ty_trm_substs ctxt
+  fun lift_aux t =
+    case subst_trms thy substs t of
+      SOME x => x
+    | NONE =>
+      (case t of
+        a $ b => lift_aux a $ lift_aux b
+      | Abs(a, ty, s) =>
+          let
+            val (y, s') = Term.dest_abs (a, ty, s)
+            val nty = subst_tys thy ty_substs ty
+          in
+            Abs(y, nty, abstract_over (Free (y, nty), lift_aux s'))
+          end
+      | Free(n, ty) => Free(n, subst_tys thy ty_substs ty)
+      | Var(n, ty) => Var(n, subst_tys thy ty_substs ty)
+      | Bound i => Bound i
+      | Const(s, ty) => Const(s, subst_tys thy ty_substs ty))
+in
+  lift_aux t
+end
+
+
+end; (* structure *)
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Attic/Quot/quotient_typ.ML	Thu Feb 25 07:57:17 2010 +0100
@@ -0,0 +1,309 @@
+(*  Title:      quotient_typ.thy
+    Author:     Cezary Kaliszyk and Christian Urban
+
+    Definition of a quotient type.
+
+*)
+
+signature QUOTIENT_TYPE =
+sig
+  val quotient_type: ((string list * binding * mixfix) * (typ * term)) list
+    -> Proof.context -> Proof.state
+
+  val quotient_type_cmd: ((((string list * binding) * mixfix) * string) * string) list
+    -> Proof.context -> Proof.state
+end;
+
+structure Quotient_Type: QUOTIENT_TYPE =
+struct
+
+open Quotient_Info;
+
+(* wrappers for define, note, Attrib.internal and theorem_i *)
+fun define (name, mx, rhs) lthy =
+let
+  val ((rhs, (_ , thm)), lthy') =
+     Local_Theory.define ((name, mx), (Attrib.empty_binding, rhs)) lthy
+in
+  ((rhs, thm), lthy')
+end
+
+fun note (name, thm, attrs) lthy =
+let
+  val ((_,[thm']), lthy') = Local_Theory.note ((name, attrs), [thm]) lthy
+in
+  (thm', lthy')
+end
+
+fun intern_attr at = Attrib.internal (K at)
+
+fun theorem after_qed goals ctxt =
+let
+  val goals' = map (rpair []) goals
+  fun after_qed' thms = after_qed (the_single thms)
+in
+  Proof.theorem_i NONE after_qed' [goals'] ctxt
+end
+
+
+
+(*** definition of quotient types ***)
+
+val mem_def1 = @{lemma "y : S ==> S y" by (simp add: mem_def)}
+val mem_def2 = @{lemma "S y ==> y : S" by (simp add: mem_def)}
+
+(* constructs the term lambda (c::rty => bool). EX (x::rty). c = rel x *)
+fun typedef_term rel rty lthy =
+let
+  val [x, c] =
+    [("x", rty), ("c", HOLogic.mk_setT rty)]
+    |> Variable.variant_frees lthy [rel]
+    |> map Free
+in
+  lambda c (HOLogic.exists_const rty $
+     lambda x (HOLogic.mk_eq (c, (rel $ x))))
+end
+
+
+(* makes the new type definitions and proves non-emptyness *)
+fun typedef_make (vs, qty_name, mx, rel, rty) lthy =
+let
+  val typedef_tac =
+     EVERY1 (map rtac [@{thm exI}, mem_def2, @{thm exI}, @{thm refl}])
+in
+  Local_Theory.theory_result
+    (Typedef.add_typedef false NONE
+       (qty_name, vs, mx)
+          (typedef_term rel rty lthy)
+             NONE typedef_tac) lthy
+end
+
+
+(* tactic to prove the quot_type theorem for the new type *)
+fun typedef_quot_type_tac equiv_thm (typedef_info: Typedef.info) =
+let
+  val rep_thm = #Rep typedef_info RS mem_def1
+  val rep_inv = #Rep_inverse typedef_info
+  val abs_inv = mem_def2 RS #Abs_inverse typedef_info
+  val rep_inj = #Rep_inject typedef_info
+in
+  (rtac @{thm quot_type.intro} THEN' RANGE [
+    rtac equiv_thm,
+    rtac rep_thm,
+    rtac rep_inv,
+    EVERY' (map rtac [abs_inv, @{thm exI}, @{thm refl}]),
+    rtac rep_inj]) 1
+end
+
+
+(* proves the quot_type theorem for the new type *)
+fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy =
+let
+  val quot_type_const = Const (@{const_name "quot_type"}, dummyT)
+  val goal =
+    HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep)
+    |> Syntax.check_term lthy
+in
+  Goal.prove lthy [] [] goal
+    (K (typedef_quot_type_tac equiv_thm typedef_info))
+end
+
+(* proves the quotient theorem for the new type *)
+fun typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_type_thm) lthy =
+let
+  val quotient_const = Const (@{const_name "Quotient"}, dummyT)
+  val goal =
+    HOLogic.mk_Trueprop (quotient_const $ rel $ abs $ rep)
+    |> Syntax.check_term lthy
+
+  val typedef_quotient_thm_tac =
+    EVERY1 [
+      K (rewrite_goals_tac [abs_def, rep_def]),
+      rtac @{thm quot_type.Quotient},
+      rtac quot_type_thm]
+in
+  Goal.prove lthy [] [] goal
+    (K typedef_quotient_thm_tac)
+end
+
+
+(* main function for constructing a quotient type *)
+fun mk_quotient_type (((vs, qty_name, mx), (rty, rel)), equiv_thm) lthy =
+let
+  (* generates the typedef *)
+  val ((qty_full_name, typedef_info), lthy1) = typedef_make (vs, qty_name, mx, rel, rty) lthy
+
+  (* abs and rep functions from the typedef *)
+  val Abs_ty = #abs_type typedef_info
+  val Rep_ty = #rep_type typedef_info
+  val Abs_name = #Abs_name typedef_info
+  val Rep_name = #Rep_name typedef_info
+  val Abs_const = Const (Abs_name, Rep_ty --> Abs_ty)
+  val Rep_const = Const (Rep_name, Abs_ty --> Rep_ty)
+
+  (* more useful abs and rep definitions *)
+  val abs_const = Const (@{const_name "quot_type.abs"}, dummyT )
+  val rep_const = Const (@{const_name "quot_type.rep"}, dummyT )
+  val abs_trm = Syntax.check_term lthy1 (abs_const $ rel $ Abs_const)
+  val rep_trm = Syntax.check_term lthy1 (rep_const $ Rep_const)
+  val abs_name = Binding.prefix_name "abs_" qty_name
+  val rep_name = Binding.prefix_name "rep_" qty_name
+
+  val ((abs, abs_def), lthy2) = define (abs_name, NoSyn, abs_trm) lthy1
+  val ((rep, rep_def), lthy3) = define (rep_name, NoSyn, rep_trm) lthy2
+
+  (* quot_type theorem *)
+  val quot_thm = typedef_quot_type_thm (rel, Abs_const, Rep_const, equiv_thm, typedef_info) lthy3
+
+  (* quotient theorem *)
+  val quotient_thm = typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_thm) lthy3
+  val quotient_thm_name = Binding.prefix_name "Quotient_" qty_name
+
+  (* name equivalence theorem *)
+  val equiv_thm_name = Binding.suffix_name "_equivp" qty_name
+
+  (* storing the quot-info *)
+  fun qinfo phi = transform_quotdata phi
+    {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm}
+  val lthy4 = Local_Theory.declaration true
+    (fn phi => quotdata_update_gen qty_full_name (qinfo phi)) lthy3
+in
+  lthy4
+  |> note (quotient_thm_name, quotient_thm, [intern_attr quotient_rules_add])
+  ||>> note (equiv_thm_name, equiv_thm, [intern_attr equiv_rules_add])
+end
+
+
+(* sanity checks for the quotient type specifications *)
+fun sanity_check ((vs, qty_name, _), (rty, rel)) =
+let
+  val rty_tfreesT = map fst (Term.add_tfreesT rty [])
+  val rel_tfrees = map fst (Term.add_tfrees rel [])
+  val rel_frees = map fst (Term.add_frees rel [])
+  val rel_vars = Term.add_vars rel []
+  val rel_tvars = Term.add_tvars rel []
+  val qty_str = Binding.str_of qty_name ^ ": "
+
+  val illegal_rel_vars =
+    if null rel_vars andalso null rel_tvars then []
+    else [qty_str ^ "illegal schematic variable(s) in the relation."]
+
+  val dup_vs =
+    (case duplicates (op =) vs of
+       [] => []
+     | dups => [qty_str ^ "duplicate type variable(s) on the lhs: " ^ commas_quote dups])
+
+  val extra_rty_tfrees =
+    (case subtract (op =) vs rty_tfreesT of
+       [] => []
+     | extras => [qty_str ^ "extra type variable(s) on the lhs: " ^ commas_quote extras])
+
+  val extra_rel_tfrees =
+    (case subtract (op =) vs rel_tfrees of
+       [] => []
+     | extras => [qty_str ^ "extra type variable(s) in the relation: " ^ commas_quote extras])
+
+  val illegal_rel_frees =
+    (case rel_frees of
+      [] => []
+    | xs => [qty_str ^ "illegal variable(s) in the relation: " ^ commas_quote xs])
+
+  val errs = illegal_rel_vars @ dup_vs @ extra_rty_tfrees @ extra_rel_tfrees @ illegal_rel_frees
+in
+  if null errs then () else error (cat_lines errs)
+end
+
+(* check for existence of map functions *)
+fun map_check ctxt (_, (rty, _)) =
+let
+  val thy = ProofContext.theory_of ctxt
+
+  fun map_check_aux rty warns =
+    case rty of
+      Type (_, []) => warns
+    | Type (s, _) => if maps_defined thy s then warns else s::warns
+    | _ => warns
+
+  val warns = map_check_aux rty []
+in
+  if null warns then ()
+  else warning ("No map function defined for " ^ commas warns ^
+    ". This will cause problems later on.")
+end
+
+
+
+(*** interface and syntax setup ***)
+
+
+(* the ML-interface takes a list of 5-tuples consisting of:
+
+ - the name of the quotient type
+ - its free type variables (first argument)
+ - its mixfix annotation
+ - the type to be quotient
+ - the relation according to which the type is quotient
+
+ it opens a proof-state in which one has to show that the
+ relations are equivalence relations
+*)
+
+fun quotient_type quot_list lthy =
+let
+  (* sanity check *)
+  val _ = List.app sanity_check quot_list
+  val _ = List.app (map_check lthy) quot_list
+
+  fun mk_goal (rty, rel) =
+  let
+    val equivp_ty = ([rty, rty] ---> @{typ bool}) --> @{typ bool}
+  in
+    HOLogic.mk_Trueprop (Const (@{const_name equivp}, equivp_ty) $ rel)
+  end
+
+  val goals = map (mk_goal o snd) quot_list
+
+  fun after_qed thms lthy =
+    fold_map mk_quotient_type (quot_list ~~ thms) lthy |> snd
+in
+  theorem after_qed goals lthy
+end
+
+fun quotient_type_cmd specs lthy =
+let
+  fun parse_spec ((((vs, qty_name), mx), rty_str), rel_str) lthy =
+  let
+    (* new parsing with proper declaration *)
+    val rty = Syntax.read_typ lthy rty_str
+    val lthy1 = Variable.declare_typ rty lthy
+    val rel = 
+      Syntax.parse_term lthy1 rel_str
+      |> Syntax.type_constraint (rty --> rty --> @{typ bool}) 
+      |> Syntax.check_term lthy1 
+    val lthy2 = Variable.declare_term rel lthy1
+  in
+    (((vs, qty_name, mx), (rty, rel)), lthy2)
+  end
+
+  val (spec', lthy') = fold_map parse_spec specs lthy
+in
+  quotient_type spec' lthy'
+end
+
+local
+  structure P = OuterParse;
+in
+
+val quotspec_parser =
+  P.and_list1 ((P.type_args -- P.binding) -- P.opt_infix -- 
+    (P.$$$ "=" |-- P.typ) -- (P.$$$ "/" |-- P.term))
+end
+
+val _ = OuterKeyword.keyword "/"
+
+val _ =
+  OuterSyntax.local_theory_to_proof "quotient_type"
+    "quotient type definitions (require equivalence proofs)"
+       OuterKeyword.thy_goal (quotspec_parser >> quotient_type_cmd)
+
+end; (* structure *)
--- a/FIXME-TODO	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,69 +0,0 @@
-Highest Priority
-================
-
-- give examples for the new quantifier translations in regularization
-  (quotient_term.ML)
-
-
-Higher Priority
-===============
-
-
-- Also, in the interest of making nicer generated documentation, you
-  might want to change all your "section" headings in Quotient.thy to
-  "subsection", and add a "header" statement to the top of the file.
-  Otherwise, each "section" gets its own chapter in the generated pdf,
-  when the rest of HOL has one chapter per theory file (the chapter
-  title comes from the "header" statement).
-
-- If the constant definition gives the wrong definition
-  term, one gets a cryptic message about absrep_fun
-
-- Handle theorems that include Ball/Bex. For this, would 
-  it help if we introduced separate Bex and Ball constants 
-  for quotienting?
-
-- The user should be able to give quotient_respects and 
-  preserves theorems in a more natural form.
-
-Lower Priority
-==============
-
-- accept partial equivalence relations
-
-- think about what happens if things go wrong (like
-  theorem cannot be lifted) / proper diagnostic 
-  messages for the user
-
-- inductions from the datatype package have a strange
-  order of quantifiers in assumptions.
-
-- find clean ways how to write down the "mathematical"
-  procedure for a possible submission (Peter submitted 
-  his work only to TPHOLs 2005...we would have to go
-  maybe for the Journal of Formalised Mathematics)
-
-- add tests for adding theorems to the various thm lists
-
-- Maybe quotient and equiv theorems like the ones for
-  [QuotList, QuotOption, QuotPair...] could be automatically
-  proven?
-
-- Examples: Finite multiset.
-
-- The current syntax of the quotient_definition is
-
-      "qconst :: qty"
-      as "rconst"
-
-  Is it possible to have the more Isabelle-like
-  syntax
-   
-      qconst :: "qty"
-      as "rconst"
-
-  That means "qconst :: qty" is not read as a term, but
-  as two entities.
-
-- Restrict automatic translation to particular quotient types
-
--- a/IsaMakefile	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,46 +0,0 @@
-
-## targets
-
-default: Quot
-images: 
-
-all: Quot
-
-
-## global settings
-
-SRC = $(ISABELLE_HOME)/src
-OUT = $(ISABELLE_OUTPUT)
-LOG = $(OUT)/log
-
-USEDIR = $(ISABELLE_TOOL) usedir -v true -t true ##-D generated
-
-
-## Quot
-
-Quot: $(LOG)/HOL-Quot.gz
-
-$(LOG)/HOL-Quot.gz: Quot/ROOT.ML Quot/*.thy
-	@$(USEDIR) HOL-Nominal Quot
-
-paper: $(LOG)/HOL-Quot-Paper.gz
-
-$(LOG)/HOL-Quot-Paper.gz: Paper/ROOT.ML Paper/document/root.tex Paper/*.thy
-	@$(USEDIR) -D generated HOL Paper
-	$(ISATOOL) document -o pdf  Paper/generated
-	@cp Paper/document.pdf paper.pdf
-
-keywords:
-	mkdir -p tmp
-	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/Pure.gz tmp 
-	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/HOL.gz tmp
-	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/Pure-ProofGeneral.gz tmp
-	cp $(ISABELLE_HOME)/heaps/polyml-5.3.0_x86-linux/log/HOL-Nominal.gz tmp
-	cp $(LOG)/HOL-Nominal-Quot.gz tmp
-	isabelle keywords -k quot tmp/*
-	
-	
-## clean
-
-clean:
-	@rm -f $(LOG)/HOL-Quot.gz
--- a/Quot/Examples/AbsRepTest.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,240 +0,0 @@
-theory AbsRepTest
-imports "../Quotient" "../Quotient_List" "../Quotient_Option" "../Quotient_Sum" "../Quotient_Product" List
-begin
-
-
-(*
-ML_command "ProofContext.debug := false"
-ML_command "ProofContext.verbose := false"
-*)
-
-ML {* open Quotient_Term *}
-
-ML {*
-fun test_funs flag ctxt (rty, qty) =
-  (absrep_fun_chk flag ctxt (rty, qty)
-   |> Syntax.string_of_term ctxt
-   |> writeln;
-   equiv_relation_chk ctxt (rty, qty) 
-   |> Syntax.string_of_term ctxt
-   |> writeln)
-*}
-
-definition
-  erel1 (infixl "\<approx>1" 50)
-where
-  "erel1 \<equiv> \<lambda>xs ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
-
-quotient_type 
-  'a fset = "'a list" / erel1
-  apply(rule equivpI)
-  unfolding erel1_def reflp_def symp_def transp_def
-  by auto
-
-definition
-  erel2 (infixl "\<approx>2" 50)
-where
-  "erel2 \<equiv> \<lambda>(xs::('a * 'a) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
-
-quotient_type 
-  'a foo = "('a * 'a) list" / erel2
-  apply(rule equivpI)
-  unfolding erel2_def reflp_def symp_def transp_def
-  by auto
-
-definition
-  erel3 (infixl "\<approx>3" 50)
-where
-  "erel3 \<equiv> \<lambda>(xs::('a * int) list) ys. \<forall>e. e \<in> set xs \<longleftrightarrow> e \<in> set ys"
-
-quotient_type 
-  'a bar = "('a * int) list" / "erel3"
-  apply(rule equivpI)
-  unfolding erel3_def reflp_def symp_def transp_def
-  by auto
-
-fun
-  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infixl "\<approx>4" 50)
-where
-  "intrel (x, y) (u, v) = (x + v = u + y)"
-
-quotient_type myint = "nat \<times> nat" / intrel
-  by (auto simp add: equivp_def expand_fun_eq)
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "nat \<times> nat"}, 
-      @{typ "myint"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "('a * 'a) list"}, 
-      @{typ "'a foo"})
-*}
-
-ML {*
-test_funs RepF @{context} 
-     (@{typ "(('a * 'a) list * 'b)"}, 
-      @{typ "('a foo * 'b)"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "(('a list) * int) list"}, 
-      @{typ "('a fset) bar"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "('a list)"}, 
-      @{typ "('a fset)"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "('a list) list"}, 
-      @{typ "('a fset) fset"})
-*}
-
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "((nat * nat) list) list"}, 
-      @{typ "((myint) fset) fset"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "(('a * 'a) list) list"}, 
-      @{typ "(('a * 'a) fset) fset"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-      (@{typ "(nat * nat) list"}, 
-       @{typ "myint fset"})
-*}
-
-ML {*
-test_funs AbsF @{context} 
-     (@{typ "('a list) list \<Rightarrow> 'a list"}, 
-      @{typ "('a fset) fset \<Rightarrow> 'a fset"})
-*}
-
-lemma OO_sym_inv:
-  assumes sr: "symp r"
-  and     ss: "symp s"
-  shows "(r OO s) x y = (s OO r) y x"
-  using sr ss
-  unfolding symp_def
-  apply (metis pred_comp.intros pred_compE ss symp_def)
-  done
-
-lemma abs_o_rep:
-  assumes a: "Quotient r absf repf"
-  shows "absf o repf = id"
-  apply(rule ext)
-  apply(simp add: Quotient_abs_rep[OF a])
-  done
-
-lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
-  apply (rule eq_reflection)
-  apply auto
-  done
-
-lemma map_rel_cong: "b \<approx>1 ba \<Longrightarrow> map f b \<approx>1 map f ba"
-  unfolding erel1_def
-  apply(simp only: set_map set_in_eq)
-  done
-
-lemma quotient_compose_list_gen_pre:
-  assumes a: "equivp r2"
-  and b: "Quotient r2 abs2 rep2"
-  shows  "(list_rel r2 OOO op \<approx>1) r s =
-          ((list_rel r2 OOO op \<approx>1) r r \<and> (list_rel r2 OOO op \<approx>1) s s \<and>
-           abs_fset (map abs2 r) = abs_fset (map abs2 s))"
-  apply rule
-  apply rule
-  apply rule
-  apply (rule list_rel_refl)
-  apply (metis equivp_def a)
-  apply rule
-  apply (rule equivp_reflp[OF fset_equivp])
-  apply (rule list_rel_refl)
-  apply (metis equivp_def a)
-  apply(rule)
-  apply rule
-  apply (rule list_rel_refl)
-  apply (metis equivp_def a)
-  apply rule
-  apply (rule equivp_reflp[OF fset_equivp])
-  apply (rule list_rel_refl)
-  apply (metis equivp_def a)
-  apply (subgoal_tac "map abs2 r \<approx>1 map abs2 s")
-  apply (metis Quotient_rel[OF Quotient_fset])
-  apply (auto)[1]
-  apply (subgoal_tac "map abs2 r = map abs2 b")
-  prefer 2
-  apply (metis Quotient_rel[OF list_quotient[OF b]])
-  apply (subgoal_tac "map abs2 s = map abs2 ba")
-  prefer 2
-  apply (metis Quotient_rel[OF list_quotient[OF b]])
-  apply (simp add: map_rel_cong)
-  apply rule
-  apply (rule rep_abs_rsp[of "list_rel r2" "map abs2"])
-  apply (rule list_quotient)
-  apply (rule b)
-  apply (rule list_rel_refl)
-  apply (metis equivp_def a)
-  apply rule
-  prefer 2
-  apply (rule rep_abs_rsp_left[of "list_rel r2" "map abs2"])
-  apply (rule list_quotient)
-  apply (rule b)
-  apply (rule list_rel_refl)
-  apply (metis equivp_def a)
-  apply (erule conjE)+
-  apply (subgoal_tac "map abs2 r \<approx>1 map abs2 s")
-  apply (rule map_rel_cong)
-  apply (assumption)
-  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp a b)
-  done
-
-lemma quotient_compose_list_gen:
-  assumes a: "Quotient r2 abs2 rep2"
-  and     b: "equivp r2" (* reflp is not enough *)
-  shows  "Quotient ((list_rel r2) OOO (op \<approx>1))
-               (abs_fset \<circ> (map abs2)) ((map rep2) \<circ> rep_fset)"
-  unfolding Quotient_def comp_def
-  apply (rule)+
-  apply (simp add: abs_o_rep[OF a] id_simps Quotient_abs_rep[OF Quotient_fset])
-  apply (rule)
-  apply (rule)
-  apply (rule)
-  apply (rule list_rel_refl)
-  apply (metis b equivp_def)
-  apply (rule)
-  apply (rule equivp_reflp[OF fset_equivp])
-  apply (rule list_rel_refl)
-  apply (metis b equivp_def)
-  apply rule
-  apply rule
-  apply(rule quotient_compose_list_gen_pre[OF b a])
-  done
-
-(* This is the general statement but the types of abs2 and rep2
-   are wrong as can be seen in following exanples *)
-lemma quotient_compose_general:
-  assumes a2: "Quotient r1 abs1 rep1"
-  and         "Quotient r2 abs2 rep2"
-  shows  "Quotient ((list_rel r2) OOO r1)
-               (abs1 \<circ> (map abs2)) ((map rep2) \<circ> rep1)"
-sorry
-
-thm quotient_compose_list_gen[OF Quotient_fset fset_equivp]
-thm quotient_compose_general[OF Quotient_fset]
-(* Doesn't work: *)
-(* thm quotient_compose_general[OF Quotient_fset Quotient_fset] *)
-
-end
--- a/Quot/Examples/FSet.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,433 +0,0 @@
-theory FSet
-imports "../Quotient" "../Quotient_List" "../Quotient_Product" List
-begin
-
-inductive
-  list_eq (infix "\<approx>" 50)
-where
-  "a#b#xs \<approx> b#a#xs"
-| "[] \<approx> []"
-| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
-| "a#a#xs \<approx> a#xs"
-| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
-| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
-
-lemma list_eq_refl:
-  shows "xs \<approx> xs"
-  by (induct xs) (auto intro: list_eq.intros)
-
-lemma equivp_list_eq:
-  shows "equivp list_eq"
-  unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
-  apply(auto intro: list_eq.intros list_eq_refl)
-  done
-
-quotient_type 
-  'a fset = "'a list" / "list_eq"
-  by (rule equivp_list_eq)
-
-quotient_definition
-   "EMPTY :: 'a fset"
-is
-  "[]::'a list"
-
-quotient_definition
-   "INSERT :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
-  "op #"
-
-quotient_definition
-   "FUNION :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
-  "op @"
-
-fun
-  card1 :: "'a list \<Rightarrow> nat"
-where
-  card1_nil: "(card1 []) = 0"
-| card1_cons: "(card1 (x # xs)) = (if (x mem xs) then (card1 xs) else (Suc (card1 xs)))"
-
-quotient_definition
-   "CARD :: 'a fset \<Rightarrow> nat"
-is
-  "card1"
-
-quotient_definition
-  "fconcat :: ('a fset) fset \<Rightarrow> 'a fset"
-is
-  "concat"
-
-term concat
-term fconcat
-
-text {*
- Maybe make_const_def should require a theorem that says that the particular lifted function
- respects the relation. With it such a definition would be impossible:
- make_const_def CARD @{term "length"} NoSyn @{typ "'a list"} @{typ "'a fset"} #> snd
-*}
-
-lemma card1_0:
-  fixes a :: "'a list"
-  shows "(card1 a = 0) = (a = [])"
-  by (induct a) auto
-
-lemma not_mem_card1:
-  fixes x :: "'a"
-  fixes xs :: "'a list"
-  shows "(~(x mem xs)) = (card1 (x # xs) = Suc (card1 xs))"
-  by auto
-
-lemma mem_cons:
-  fixes x :: "'a"
-  fixes xs :: "'a list"
-  assumes a : "x mem xs"
-  shows "x # xs \<approx> xs"
-  using a by (induct xs) (auto intro: list_eq.intros )
-
-lemma card1_suc:
-  fixes xs :: "'a list"
-  fixes n :: "nat"
-  assumes c: "card1 xs = Suc n"
-  shows "\<exists>a ys. ~(a mem ys) \<and> xs \<approx> (a # ys)"
-  using c
-apply(induct xs)
-apply (metis Suc_neq_Zero card1_0)
-apply (metis FSet.card1_cons list_eq.intros(6) list_eq_refl mem_cons)
-done
-
-definition
-  rsp_fold
-where
-  "rsp_fold f = ((!u v. (f u v = f v u)) \<and> (!u v w. ((f u (f v w) = f (f u v) w))))"
-
-primrec
-  fold1
-where
-  "fold1 f (g :: 'a \<Rightarrow> 'b) (z :: 'b) [] = z"
-| "fold1 f g z (a # A) =
-     (if rsp_fold f
-     then (
-       if (a mem A) then (fold1 f g z A) else (f (g a) (fold1 f g z A))
-     ) else z)"
-
-lemma fs1_strong_cases:
-  fixes X :: "'a list"
-  shows "(X = []) \<or> (\<exists>a. \<exists> Y. (~(a mem Y) \<and> (X \<approx> a # Y)))"
-  apply (induct X)
-  apply (simp)
-  apply (metis List.member.simps(1) list_eq.intros(6) list_eq_refl mem_cons)
-  done
-
-quotient_definition
-   "IN :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
-is
-  "op mem"
-
-quotient_definition
-   "FOLD :: ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b fset \<Rightarrow> 'a"
-is
-  "fold1"
-
-quotient_definition
-  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
-is
-  "map"
-
-lemma mem_rsp:
-  fixes z
-  assumes a: "x \<approx> y"
-  shows "(z mem x) = (z mem y)"
-  using a by induct auto
-
-lemma ho_memb_rsp[quot_respect]:
-  "(op = ===> (op \<approx> ===> op =)) (op mem) (op mem)"
-  by (simp add: mem_rsp)
-
-lemma card1_rsp:
-  fixes a b :: "'a list"
-  assumes e: "a \<approx> b"
-  shows "card1 a = card1 b"
-  using e by induct (simp_all add: mem_rsp)
-
-lemma ho_card1_rsp[quot_respect]: 
-  "(op \<approx> ===> op =) card1 card1"
-  by (simp add: card1_rsp)
-
-lemma cons_rsp:
-  fixes z
-  assumes a: "xs \<approx> ys"
-  shows "(z # xs) \<approx> (z # ys)"
-  using a by (rule list_eq.intros(5))
-
-lemma ho_cons_rsp[quot_respect]:
-  "(op = ===> op \<approx> ===> op \<approx>) op # op #"
-  by (simp add: cons_rsp)
-
-lemma append_rsp_aux1:
-  assumes a : "l2 \<approx> r2 "
-  shows "(h @ l2) \<approx> (h @ r2)"
-using a
-apply(induct h)
-apply(auto intro: list_eq.intros(5))
-done
-
-lemma append_rsp_aux2:
-  assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "
-  shows "(l1 @ l2) \<approx> (r1 @ r2)"
-using a
-apply(induct arbitrary: l2 r2)
-apply(simp_all)
-apply(blast intro: list_eq.intros append_rsp_aux1)+
-done
-
-lemma append_rsp[quot_respect]:
-  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
-  by (auto simp add: append_rsp_aux2)
-
-lemma map_rsp:
-  assumes a: "a \<approx> b"
-  shows "map f a \<approx> map f b"
-  using a
-  apply (induct)
-  apply(auto intro: list_eq.intros)
-  done
-
-lemma ho_map_rsp[quot_respect]:
-  "(op = ===> op \<approx> ===> op \<approx>) map map"
-  by (simp add: map_rsp)
-
-lemma map_append:
-  "(map f (a @ b)) \<approx> (map f a) @ (map f b)"
- by simp (rule list_eq_refl)
-
-lemma ho_fold_rsp[quot_respect]:
-  "(op = ===> op = ===> op = ===> op \<approx> ===> op =) fold1 fold1"
-  apply (auto)
-  apply (case_tac "rsp_fold x")
-  prefer 2
-  apply (erule_tac list_eq.induct)
-  apply (simp_all)
-  apply (erule_tac list_eq.induct)
-  apply (simp_all)
-  apply (auto simp add: mem_rsp rsp_fold_def)
-done
-
-lemma list_equiv_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>"
-by (auto intro: list_eq.intros)
-
-lemma "IN x EMPTY = False"
-apply(lifting member.simps(1))
-done
-
-lemma "IN x (INSERT y xa) = (x = y \<or> IN x xa)"
-apply (lifting member.simps(2))
-done
-
-lemma "INSERT a (INSERT a x) = INSERT a x"
-apply (lifting list_eq.intros(4))
-done
-
-lemma "x = xa \<Longrightarrow> INSERT a x = INSERT a xa"
-apply (lifting list_eq.intros(5))
-done
-
-lemma "CARD x = Suc n \<Longrightarrow> (\<exists>a b. \<not> IN a b & x = INSERT a b)"
-apply (lifting card1_suc)
-done
-
-lemma "(\<not> IN x xa) = (CARD (INSERT x xa) = Suc (CARD xa))"
-apply (lifting not_mem_card1)
-done
-
-lemma "FOLD f g (z::'b) (INSERT a x) =
-  (if rsp_fold f then if IN a x then FOLD f g z x else f (g a) (FOLD f g z x) else z)"
-apply(lifting fold1.simps(2))
-done
-
-lemma "fmap f (FUNION (x::'b fset) (xa::'b fset)) = FUNION (fmap f x) (fmap f xa)"
-apply (lifting map_append)
-done
-
-lemma "FUNION (FUNION x xa) xb = FUNION x (FUNION xa xb)"
-apply (lifting append_assoc)
-done
-
-
-lemma "\<lbrakk>P EMPTY; \<And>a x. P x \<Longrightarrow> P (INSERT a x)\<rbrakk> \<Longrightarrow> P l"
-apply(lifting list.induct)
-done
-
-lemma list_induct_part:
-  assumes a: "P (x :: 'a list) ([] :: 'c list)"
-  assumes b: "\<And>e t. P x t \<Longrightarrow> P x (e # t)"
-  shows "P x l"
-  apply (rule_tac P="P x" in list.induct)
-  apply (rule a)
-  apply (rule b)
-  apply (assumption)
-  done
-
-lemma "P (x :: 'a list) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
-apply (lifting list_induct_part)
-done
-
-lemma "P (x :: 'a fset) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
-apply (lifting list_induct_part)
-done
-
-lemma "P (x :: 'a fset) ([] :: 'c list) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (e # t)) \<Longrightarrow> P x l"
-apply (lifting list_induct_part)
-done
-
-quotient_type 'a fset2 = "'a list" / "list_eq"
-  by (rule equivp_list_eq)
-
-quotient_definition
-   "EMPTY2 :: 'a fset2"
-is
-  "[]::'a list"
-
-quotient_definition
-   "INSERT2 :: 'a \<Rightarrow> 'a fset2 \<Rightarrow> 'a fset2"
-is
-  "op #"
-
-lemma "P (x :: 'a fset2) (EMPTY :: 'c fset) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT e t)) \<Longrightarrow> P x l"
-apply (lifting list_induct_part)
-done
-
-lemma "P (x :: 'a fset) (EMPTY2 :: 'c fset2) \<Longrightarrow> (\<And>e t. P x t \<Longrightarrow> P x (INSERT2 e t)) \<Longrightarrow> P x l"
-apply (lifting list_induct_part)
-done
-
-quotient_definition
-  "fset_rec :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
-is
-  "list_rec"
-
-quotient_definition
-  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
-is
-  "list_case"
-
-(* Probably not true without additional assumptions about the function *)
-lemma list_rec_rsp[quot_respect]:
-  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_rec list_rec"
-  apply (auto)
-  apply (erule_tac list_eq.induct)
-  apply (simp_all)
-  sorry
-
-lemma list_case_rsp[quot_respect]:
-  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
-  apply (auto)
-  sorry
-
-lemma "fset_rec (f1::'t) x (INSERT a xa) = x a xa (fset_rec f1 x xa)"
-apply (lifting list.recs(2))
-done
-
-lemma "fset_case (f1::'t) f2 (INSERT a xa) = f2 a xa"
-apply (lifting list.cases(2))
-done
-
-lemma ttt: "((op @) x ((op #) e [])) = (((op #) e x))"
-sorry
-
-lemma "(FUNION x (INSERT e EMPTY)) = ((INSERT e x))"
-apply (lifting ttt)
-done
-
-
-lemma ttt2: "(\<lambda>e. ((op @) x ((op #) e []))) = (\<lambda>e. ((op #) e x))"
-sorry
-
-lemma "(\<lambda>e. (FUNION x (INSERT e EMPTY))) = (\<lambda>e. (INSERT e x))"
-apply(lifting ttt2)
-apply(regularize)
-apply(rule impI)
-apply(simp)
-apply(rule allI)
-apply(rule list_eq_refl)
-done
-
-lemma ttt3: "(\<lambda>x. ((op @) x (e # []))) = (op #) e"
-sorry
-
-lemma "(\<lambda>x. (FUNION x (INSERT e EMPTY))) = INSERT e"
-apply(lifting ttt3)
-apply(regularize)
-apply(auto simp add: cons_rsp)
-done
-lemma hard: "(\<lambda>P. \<lambda>Q. P (Q (x::'a list))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a list)))"
-sorry
-
-lemma eq_imp_rel:
-  shows "equivp R \<Longrightarrow> a = b \<longrightarrow> R a b"
-  by (simp add: equivp_reflp)
-
-
-lemma hard_lift: "(\<lambda>P. \<lambda>Q. P (Q (x::'a fset))) = (\<lambda>P. \<lambda>Q. Q (P (x::'a fset)))"
-apply(lifting hard)
-apply(regularize)
-apply(rule fun_rel_id_asm)
-apply(subst babs_simp)
-apply(tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
-apply(rule fun_rel_id_asm)
-apply(rule impI)
-apply(rule mp[OF eq_imp_rel[OF fset_equivp]])
-apply(drule fun_cong)
-apply(drule fun_cong)
-apply(assumption)
-done
-
-lemma test: "All (\<lambda>(x::'a list, y). x = y)"
-sorry
-
-lemma  "All (\<lambda>(x::'a fset, y). x = y)"
-apply(lifting test)
-done
-
-lemma test2: "Ex (\<lambda>(x::'a list, y). x = y)"
-sorry
-
-lemma  "Ex (\<lambda>(x::'a fset, y). x = y)"
-apply(lifting test2)
-done
-
-lemma test3: "All (\<lambda> (x :: 'a list, y, z). x = y \<and> y = z)"
-sorry
-
-lemma "All (\<lambda> (x :: 'a fset, y, z). x = y \<and> y = z)"
-apply(lifting test3)
-done
-
-lemma test4: "\<forall> (x :: 'a list, y, z) \<in> Respects(
-  prod_rel (op \<approx>) (prod_rel (op \<approx>) (op \<approx>))
-). x = y \<and> y = z"
-sorry
-
-lemma "All (\<lambda> (x :: 'a fset, y, z). x = y \<and> y = z)"
-apply (lifting test4)
-sorry
-
-lemma test5: "\<forall> (x :: 'a list \<Rightarrow> 'a list, y) \<in> Respects(
-  prod_rel (op \<approx> ===> op \<approx>) (op \<approx> ===> op \<approx>)
-). (op \<approx> ===> op \<approx>) x y"
-sorry
-
-lemma "All (\<lambda> (x :: 'a fset \<Rightarrow> 'a fset, y). x = y)"
-apply (lifting test5)
-done
-
-lemma test6: "\<forall> (x :: 'a list \<Rightarrow> 'a list, y, z) \<in> Respects(
-  prod_rel (op \<approx> ===> op \<approx>) (prod_rel (op \<approx> ===> op \<approx>) (op \<approx> ===> op \<approx>))
-). (op \<approx> ===> op \<approx>) x y \<and> (op \<approx> ===> op \<approx>) y z"
-sorry
-
-lemma "All (\<lambda> (x :: 'a fset \<Rightarrow> 'a fset, y, z). x = y \<and> y = z)"
-apply (lifting test6)
-done
-
-end
--- a/Quot/Examples/FSet2.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,128 +0,0 @@
-theory FSet2
-imports "../Quotient" "../Quotient_List" List
-begin
-
-inductive
-  list_eq (infix "\<approx>" 50)
-where
-  "a#b#xs \<approx> b#a#xs"
-| "[] \<approx> []"
-| "xs \<approx> ys \<Longrightarrow> ys \<approx> xs"
-| "a#a#xs \<approx> a#xs"
-| "xs \<approx> ys \<Longrightarrow> a#xs \<approx> a#ys"
-| "\<lbrakk>xs1 \<approx> xs2; xs2 \<approx> xs3\<rbrakk> \<Longrightarrow> xs1 \<approx> xs3"
-
-lemma list_eq_refl:
-  shows "xs \<approx> xs"
-by (induct xs) (auto intro: list_eq.intros)
-
-lemma equivp_list_eq:
-  shows "equivp list_eq"
-unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
-by (auto intro: list_eq.intros list_eq_refl)
-
-quotient_type
-  'a fset = "'a list" / "list_eq"
-  by (rule equivp_list_eq)
-
-quotient_definition
-  fempty ("{||}")
-where
-  "fempty :: 'a fset"
-is
-  "[]"
-
-quotient_definition
-  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
-  "(op #)"
-
-lemma finsert_rsp[quot_respect]:
-  shows "(op = ===> op \<approx> ===> op \<approx>) (op #) (op #)"
-by (auto intro: list_eq.intros)
-
-quotient_definition
-  funion ("_ \<union>f _" [65,66] 65)
-where
-  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
-  "(op @)"
-
-lemma append_rsp_aux1:
-  assumes a : "l2 \<approx> r2 "
-  shows "(h @ l2) \<approx> (h @ r2)"
-using a
-apply(induct h)
-apply(auto intro: list_eq.intros(5))
-done
-
-lemma append_rsp_aux2:
-  assumes a : "l1 \<approx> r1" "l2 \<approx> r2 "
-  shows "(l1 @ l2) \<approx> (r1 @ r2)"
-using a
-apply(induct arbitrary: l2 r2)
-apply(simp_all)
-apply(blast intro: list_eq.intros append_rsp_aux1)+
-done
-
-lemma append_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
-  by (auto simp add: append_rsp_aux2)
-
-
-quotient_definition
-  fmem ("_ \<in>f _" [50, 51] 50)
-where
-  "fmem :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
-is
-  "(op mem)"
-
-lemma memb_rsp_aux:
-  assumes a: "x \<approx> y"
-  shows "(z mem x) = (z mem y)"
-  using a by induct auto
-
-lemma memb_rsp[quot_respect]:
-  shows "(op = ===> (op \<approx> ===> op =)) (op mem) (op mem)"
-  by (simp add: memb_rsp_aux)
-
-definition
-  fnot_mem :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ \<notin>f _" [50, 51] 50)
-where
-  "x \<notin>f S \<equiv> \<not>(x \<in>f S)"
-
-definition
-  "inter_list" :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
-where
-  "inter_list X Y \<equiv> [x \<leftarrow> X. x\<in>set Y]"
-
-quotient_definition
-  finter ("_ \<inter>f _" [70, 71] 70)
-where
-  "finter::'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
-  "inter_list"
-
-no_syntax
-  "@Finset"     :: "args => 'a fset"                       ("{|(_)|}")
-syntax
-  "@Finfset"     :: "args => 'a fset"                       ("{|(_)|}")
-translations
-  "{|x, xs|}"     == "CONST finsert x {|xs|}"
-  "{|x|}"         == "CONST finsert x {||}"
-
-
-subsection {* Empty sets *}
-
-lemma test:
-  shows "\<not>(x # xs \<approx> [])"
-sorry
-
-lemma finsert_not_empty[simp]: 
-  shows "finsert x S \<noteq> {||}"
-  by (lifting test)
-
-
-
-
-end;
--- a/Quot/Examples/FSet3.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,714 +0,0 @@
-theory FSet3
-imports "../Quotient" "../Quotient_List" List
-begin
-
-ML {*
-structure QuotientRules = Named_Thms
-  (val name = "quot_thm"
-   val description = "Quotient theorems.")
-*}
-
-ML {*
-open QuotientRules
-*}
-
-fun
-  list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50)
-where
-  "list_eq xs ys = (\<forall>x. x \<in> set xs \<longleftrightarrow> x \<in> set ys)"
-
-lemma list_eq_equivp:
-  shows "equivp list_eq"
-unfolding equivp_reflp_symp_transp reflp_def symp_def transp_def
-by auto
-
-(* FIXME-TODO: because of beta-reduction, one cannot give the *)
-(* FIXME-TODO: relation as a term or abbreviation             *) 
-quotient_type 
-  'a fset = "'a list" / "list_eq"
-by (rule list_eq_equivp)
-
-
-section {* empty fset, finsert and membership *} 
-
-quotient_definition
-  fempty  ("{||}")
-where
-  "fempty :: 'a fset"
-is "[]::'a list"
-
-quotient_definition
-  "finsert :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" 
-is "op #"
-
-syntax
-  "@Finset"     :: "args => 'a fset"  ("{|(_)|}")
-
-translations
-  "{|x, xs|}" == "CONST finsert x {|xs|}"
-  "{|x|}"     == "CONST finsert x {||}"
-
-definition 
-  memb :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
-where
-  "memb x xs \<equiv> x \<in> set xs"
-
-quotient_definition
-  fin ("_ |\<in>| _" [50, 51] 50)
-where
-  "fin :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool"
-is "memb"
-
-abbreviation
-  fnotin :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" ("_ |\<notin>| _" [50, 51] 50)
-where
-  "a |\<notin>| S \<equiv> \<not>(a |\<in>| S)"
-
-lemma memb_rsp[quot_respect]: 
-  shows "(op = ===> op \<approx> ===> op =) memb memb"
-by (auto simp add: memb_def)
-
-lemma nil_rsp[quot_respect]:
-  shows "[] \<approx> []"
-by simp
-
-lemma cons_rsp[quot_respect]: 
-  shows "(op = ===> op \<approx> ===> op \<approx>) op # op #"
-by simp
-
-
-section {* Augmenting a set -- @{const finsert} *}
-
-text {* raw section *}
-
-lemma nil_not_cons:
-  shows "\<not>[] \<approx> x # xs"
-  by auto
-
-lemma memb_cons_iff:
-  shows "memb x (y # xs) = (x = y \<or> memb x xs)"
-  by (induct xs) (auto simp add: memb_def)
-
-lemma memb_consI1:
-  shows "memb x (x # xs)"
-  by (simp add: memb_def)
-
-lemma memb_consI2:
-  shows "memb x xs \<Longrightarrow> memb x (y # xs)"
-  by (simp add: memb_def)
-
-lemma memb_absorb:
-  shows "memb x xs \<Longrightarrow> x # xs \<approx> xs"
-  by (induct xs) (auto simp add: memb_def id_simps)
-
-text {* lifted section *}
-
-lemma fin_finsert_iff[simp]:
-  "x |\<in>| finsert y S = (x = y \<or> x |\<in>| S)"
-by (lifting memb_cons_iff) 
-
-lemma
-  shows finsertI1: "x |\<in>| finsert x S"
-  and   finsertI2: "x |\<in>| S \<Longrightarrow> x |\<in>| finsert y S"
-  by (lifting memb_consI1, lifting memb_consI2)
-
-
-lemma finsert_absorb [simp]: 
-  shows "x |\<in>| S \<Longrightarrow> finsert x S = S"
-  by (lifting memb_absorb)
-
-
-section {* Singletons *}
-
-text {* raw section *}
-
-lemma singleton_list_eq:
-  shows "[x] \<approx> [y] \<longleftrightarrow> x = y"
-  by (simp add: id_simps) auto
-
-text {* lifted section *}
-
-lemma fempty_not_finsert[simp]:
-  shows "{||} \<noteq> finsert x S"
-  by (lifting nil_not_cons)
-
-lemma fsingleton_eq[simp]:
-  shows "{|x|} = {|y|} \<longleftrightarrow> x = y"
-  by (lifting singleton_list_eq)
-
-section {* Union *}
-
-quotient_definition
-  funion  (infixl "|\<union>|" 65)
-where
-  "funion :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-is
-  "op @"
-
-section {* Cardinality of finite sets *}
-
-fun
-  fcard_raw :: "'a list \<Rightarrow> nat"
-where
-  fcard_raw_nil:  "fcard_raw [] = 0"
-| fcard_raw_cons: "fcard_raw (x # xs) = (if memb x xs then fcard_raw xs else Suc (fcard_raw xs))"
-
-quotient_definition
-  "fcard :: 'a fset \<Rightarrow> nat" 
-is "fcard_raw"
-
-text {* raw section *}
-
-lemma fcard_raw_ge_0:
-  assumes a: "x \<in> set xs"
-  shows "0 < fcard_raw xs"
-using a
-by (induct xs) (auto simp add: memb_def)
-
-lemma fcard_raw_delete_one:
-  "fcard_raw ([x \<leftarrow> xs. x \<noteq> y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)"
-by (induct xs) (auto dest: fcard_raw_ge_0 simp add: memb_def)
-
-lemma fcard_raw_rsp_aux:
-  assumes a: "a \<approx> b"
-  shows "fcard_raw a = fcard_raw b"
-using a
-apply(induct a arbitrary: b)
-apply(auto simp add: memb_def)
-apply(metis)
-apply(drule_tac x="[x \<leftarrow> b. x \<noteq> a1]" in meta_spec)
-apply(simp add: fcard_raw_delete_one)
-apply(metis Suc_pred' fcard_raw_ge_0 fcard_raw_delete_one memb_def)
-done
-
-lemma fcard_raw_rsp[quot_respect]:
-  "(op \<approx> ===> op =) fcard_raw fcard_raw"
-  by (simp add: fcard_raw_rsp_aux)
-
-text {* lifted section *}
-
-lemma fcard_fempty [simp]:
-  shows "fcard {||} = 0"
-by (lifting fcard_raw_nil)
-
-lemma fcard_finsert_if [simp]:
-  shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))"
-by (lifting fcard_raw_cons)
-
-
-section {* Induction and Cases rules for finite sets *}
-
-lemma fset_exhaust[case_names fempty finsert, cases type: fset]:
-  shows "\<lbrakk>S = {||} \<Longrightarrow> P; \<And>x S'. S = finsert x S' \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
-by (lifting list.exhaust)
-
-lemma fset_induct[case_names fempty finsert]:
-  shows "\<lbrakk>P {||}; \<And>x S. P S \<Longrightarrow> P (finsert x S)\<rbrakk> \<Longrightarrow> P S"
-by (lifting list.induct)
-
-lemma fset_induct2[case_names fempty finsert, induct type: fset]:
-  assumes prem1: "P {||}" 
-  and     prem2: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (finsert x S)"
-  shows "P S"
-proof(induct S rule: fset_induct)
-  case fempty
-  show "P {||}" by (rule prem1)
-next
-  case (finsert x S)
-  have asm: "P S" by fact
-  show "P (finsert x S)"
-  proof(cases "x |\<in>| S")
-    case True
-    have "x |\<in>| S" by fact
-    then show "P (finsert x S)" using asm by simp
-  next
-    case False
-    have "x |\<notin>| S" by fact
-    then show "P (finsert x S)" using prem2 asm by simp
-  qed
-qed
-
-
-section {* fmap and fset comprehension *}
-
-quotient_definition
-  "fmap :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset"
-is
- "map"
-
-quotient_definition
-  "fconcat :: ('a fset) fset => 'a fset"
-is
- "concat"
-
-(*lemma fconcat_rsp[quot_respect]:
-  shows "((list_rel op \<approx>) ===> op \<approx>) concat concat"
-apply(auto)
-sorry
-*)
-
-(* PROBLEM: these lemmas needs to be restated, since  *)
-(* concat.simps(1) and concat.simps(2) contain the    *)
-(* type variables ?'a1.0 (which are turned into frees *)
-(* 'a_1                                               *)
-lemma concat1:
-  shows "concat [] \<approx> []"
-by (simp add: id_simps)
-
-lemma concat2:
-  shows "concat (x # xs) \<approx> x @ concat xs"
-by (simp add: id_simps)
-
-lemma concat_rsp[quot_respect]:
-  shows "(list_rel op \<approx> OOO op \<approx> ===> op \<approx>) concat concat"
-sorry
-
-lemma nil_rsp2[quot_respect]: "(list_rel op \<approx> OOO op \<approx>) [] []"
-  apply (metis FSet3.nil_rsp list_rel.simps(1) pred_comp.intros)
-  done
-
-lemma set_in_eq: "(\<forall>e. ((e \<in> A) \<longleftrightarrow> (e \<in> B))) \<equiv> A = B"
-  apply (rule eq_reflection)
-  apply auto
-  done
-
-lemma map_rel_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba"
-  unfolding list_eq.simps
-  apply(simp only: set_map set_in_eq)
-  done
-
-lemma quotient_compose_list_pre:
-  "(list_rel op \<approx> OOO op \<approx>) r s =
-  ((list_rel op \<approx> OOO op \<approx>) r r \<and> (list_rel op \<approx> OOO op \<approx>) s s \<and>
-  abs_fset (map abs_fset r) = abs_fset (map abs_fset s))"
-  apply rule
-  apply rule
-  apply rule
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply rule
-  apply (rule equivp_reflp[OF fset_equivp])
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply(rule)
-  apply rule
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply rule
-  apply (rule equivp_reflp[OF fset_equivp])
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply (subgoal_tac "map abs_fset r \<approx> map abs_fset s")
-  apply (metis Quotient_rel[OF Quotient_fset])
-  apply (auto simp only:)[1]
-  apply (subgoal_tac "map abs_fset r = map abs_fset b")
-  prefer 2
-  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
-  apply (subgoal_tac "map abs_fset s = map abs_fset ba")
-  prefer 2
-  apply (metis Quotient_rel[OF list_quotient[OF Quotient_fset]])
-  apply (simp only: map_rel_cong)
-  apply rule
-  apply (rule rep_abs_rsp[of "list_rel op \<approx>" "map abs_fset"])
-  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply rule
-  prefer 2
-  apply (rule rep_abs_rsp_left[of "list_rel op \<approx>" "map abs_fset"])
-  apply (tactic {* Quotient_Tacs.quotient_tac @{context} 1 *})
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply (erule conjE)+
-  apply (subgoal_tac "map abs_fset r \<approx> map abs_fset s")
-  prefer 2
-  apply (metis Quotient_def Quotient_fset equivp_reflp fset_equivp)
-  apply (rule map_rel_cong)
-  apply (assumption)
-  done
-
-lemma quotient_compose_list[quot_thm]:
-  shows  "Quotient ((list_rel op \<approx>) OOO (op \<approx>))
-    (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)"
-  unfolding Quotient_def comp_def
-  apply (rule)+
-  apply (simp add: abs_o_rep[OF Quotient_fset] id_simps Quotient_abs_rep[OF Quotient_fset])
-  apply (rule)
-  apply (rule)
-  apply (rule)
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply (rule)
-  apply (rule equivp_reflp[OF fset_equivp])
-  apply (rule list_rel_refl)
-  apply (metis equivp_def fset_equivp)
-  apply rule
-  apply rule
-  apply(rule quotient_compose_list_pre)
-  done
-
-lemma fconcat_empty:
-  shows "fconcat {||} = {||}"
-apply(lifting concat1)
-apply(cleaning)
-apply(simp add: comp_def)
-apply(fold fempty_def[simplified id_simps])
-apply(rule refl)
-done
-
-(* Should be true *)
-
-lemma insert_rsp2[quot_respect]:
-  "(op \<approx> ===> list_rel op \<approx> OOO op \<approx> ===> list_rel op \<approx> OOO op \<approx>) op # op #"
-apply auto
-apply (simp add: set_in_eq)
-sorry
-
-lemma append_rsp[quot_respect]:
-  "(op \<approx> ===> op \<approx> ===> op \<approx>) op @ op @"
-  by (auto)
-
-lemma fconcat_insert:
-  shows "fconcat (finsert x S) = x |\<union>| fconcat S"
-apply(lifting concat2)
-apply(cleaning)
-apply (simp add: finsert_def fconcat_def comp_def)
-apply cleaning
-done
-
-text {* raw section *}
-
-lemma map_rsp_aux:
-  assumes a: "a \<approx> b"
-  shows "map f a \<approx> map f b"
-  using a
-apply(induct a arbitrary: b)
-apply(auto)
-apply(metis rev_image_eqI)
-done
-
-lemma map_rsp[quot_respect]:
-  shows "(op = ===> op \<approx> ===> op \<approx>) map map"
-by (auto simp add: map_rsp_aux)
-
-
-text {* lifted section *}
-
-(* TBD *)
-
-text {* syntax for fset comprehensions (adapted from lists) *}
-
-nonterminals fsc_qual fsc_quals
-
-syntax
-"_fsetcompr" :: "'a \<Rightarrow> fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> 'a fset"  ("{|_ . __")
-"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ <- _")
-"_fsc_test" :: "bool \<Rightarrow> fsc_qual" ("_")
-"_fsc_end" :: "fsc_quals" ("|}")
-"_fsc_quals" :: "fsc_qual \<Rightarrow> fsc_quals \<Rightarrow> fsc_quals" (", __")
-"_fsc_abs" :: "'a => 'b fset => 'b fset"
-
-syntax (xsymbols)
-"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
-syntax (HTML output)
-"_fsc_gen" :: "'a \<Rightarrow> 'a fset \<Rightarrow> fsc_qual" ("_ \<leftarrow> _")
-
-parse_translation (advanced) {*
-let
-  val femptyC = Syntax.const @{const_name fempty};
-  val finsertC = Syntax.const @{const_name finsert};
-  val fmapC = Syntax.const @{const_name fmap};
-  val fconcatC = Syntax.const @{const_name fconcat};
-  val IfC = Syntax.const @{const_name If};
-  fun fsingl x = finsertC $ x $ femptyC;
-
-  fun pat_tr ctxt p e opti = (* %x. case x of p => e | _ => [] *)
-    let
-      val x = Free (Name.variant (fold Term.add_free_names [p, e] []) "x", dummyT);
-      val e = if opti then fsingl e else e;
-      val case1 = Syntax.const "_case1" $ p $ e;
-      val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN
-                                        $ femptyC;
-      val cs = Syntax.const "_case2" $ case1 $ case2
-      val ft = Datatype_Case.case_tr false Datatype.info_of_constr
-                 ctxt [x, cs]
-    in lambda x ft end;
-
-  fun abs_tr ctxt (p as Free(s,T)) e opti =
-        let val thy = ProofContext.theory_of ctxt;
-            val s' = Sign.intern_const thy s
-        in if Sign.declared_const thy s'
-           then (pat_tr ctxt p e opti, false)
-           else (lambda p e, true)
-        end
-    | abs_tr ctxt p e opti = (pat_tr ctxt p e opti, false);
-
-  fun fsc_tr ctxt [e, Const("_fsc_test",_) $ b, qs] =
-        let 
-          val res = case qs of 
-                      Const("_fsc_end",_) => fsingl e
-                    | Const("_fsc_quals",_)$ q $ qs => fsc_tr ctxt [e, q, qs];
-        in 
-          IfC $ b $ res $ femptyC 
-        end
-
-    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_end",_)] =
-         (case abs_tr ctxt p e true of
-            (f,true) => fmapC $ f $ es
-          | (f, false) => fconcatC $ (fmapC $ f $ es))
-       
-    | fsc_tr ctxt [e, Const("_fsc_gen",_) $ p $ es, Const("_fsc_quals",_) $ q $ qs] =
-        let
-          val e' = fsc_tr ctxt [e, q, qs];
-        in 
-          fconcatC $ (fmapC $ (fst (abs_tr ctxt p e' false)) $ es) 
-        end
-
-in [("_fsetcompr", fsc_tr)] end
-*}
-
-(* examles *)
-term "{|(x,y,z). b|}"
-term "{|x. x \<leftarrow> xs|}"
-term "{|(x,y,z). x\<leftarrow>xs|}"
-term "{|e x y. x\<leftarrow>xs, y\<leftarrow>ys|}"
-term "{|(x,y,z). x<a, x>b|}"
-term "{|(x,y,z). x\<leftarrow>xs, x>b|}"
-term "{|(x,y,z). x<a, x\<leftarrow>xs|}"
-term "{|(x,y). Cons True x \<leftarrow> xs|}"
-term "{|(x,y,z). Cons x [] \<leftarrow> xs|}"
-term "{|(x,y,z). x<a, x>b, x=d|}"
-term "{|(x,y,z). x<a, x>b, y\<leftarrow>ys|}"
-term "{|(x,y,z). x<a, x\<leftarrow>xs,y>b|}"
-term "{|(x,y,z). x<a, x\<leftarrow>xs, y\<leftarrow>ys|}"
-term "{|(x,y,z). x\<leftarrow>xs, x>b, y<a|}"
-term "{|(x,y,z). x\<leftarrow>xs, x>b, y\<leftarrow>ys|}"
-term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,y>x|}"
-term "{|(x,y,z). x\<leftarrow>xs, y\<leftarrow>ys,z\<leftarrow>zs|}"
-
-
-(* BELOW CONSTRUCTION SITE *)
-
-
-lemma no_mem_nil: 
-  "(\<forall>a. a \<notin> set A) = (A = [])"
-by (induct A) (auto)
-
-lemma none_mem_nil: 
-  "(\<forall>a. a \<notin> set A) = (A \<approx> [])"
-by simp
-
-lemma mem_cons: 
-  "a \<in> set A \<Longrightarrow> a # A \<approx> A"
-by auto
-
-lemma cons_left_comm:
-  "x # y # A \<approx> y # x # A"
-by (auto simp add: id_simps)
-
-lemma cons_left_idem:
-  "x # x # A \<approx> x # A"
-by (auto simp add: id_simps)
-
-lemma finite_set_raw_strong_cases:
-  "(X = []) \<or> (\<exists>a Y. ((a \<notin> set Y) \<and> (X \<approx> a # Y)))"
-  apply (induct X)
-  apply (simp)
-  apply (rule disjI2)
-  apply (erule disjE)
-  apply (rule_tac x="a" in exI)
-  apply (rule_tac x="[]" in exI)
-  apply (simp)
-  apply (erule exE)+
-  apply (case_tac "a = aa")
-  apply (rule_tac x="a" in exI)
-  apply (rule_tac x="Y" in exI)
-  apply (simp)
-  apply (rule_tac x="aa" in exI)
-  apply (rule_tac x="a # Y" in exI)
-  apply (auto)
-  done
-
-fun
-  delete_raw :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a list"
-where
-  "delete_raw [] x = []"
-| "delete_raw (a # A) x = (if (a = x) then delete_raw A x else a # (delete_raw A x))"
-
-lemma mem_delete_raw:
-  "x \<in> set (delete_raw A a) = (x \<in> set A \<and> \<not>(x = a))"
-  by (induct A arbitrary: x a) (auto)
-
-lemma mem_delete_raw_ident:
-  "\<not>(a \<in> set (delete_raw A a))"
-by (induct A) (auto)
-
-lemma not_mem_delete_raw_ident:
-  "b \<notin> set A \<Longrightarrow> (delete_raw A b = A)"
-by (induct A) (auto)
-
-lemma delete_raw_RSP:
-  "A \<approx> B \<Longrightarrow> delete_raw A a \<approx> delete_raw B a"
-apply(induct A arbitrary: B a)
-apply(auto)
-sorry
-
-lemma cons_delete_raw:
-  "a # (delete_raw A a) \<approx> (if a \<in> set A then A else (a # A))"
-sorry
-
-lemma mem_cons_delete_raw:
-    "a \<in> set A \<Longrightarrow> a # (delete_raw A a) \<approx> A"
-sorry
-
-lemma finite_set_raw_delete_raw_cases:
-    "X = [] \<or> (\<exists>a. a mem X \<and> X \<approx> a # delete_raw X a)"
-  by (induct X) (auto)
-
-
-
-
-
-lemma list2set_thm:
-  shows "set [] = {}"
-  and "set (h # t) = insert h (set t)"
-  by (auto)
-
-lemma list2set_RSP:
-  "A \<approx> B \<Longrightarrow> set A = set B"
-  by auto
-
-definition
-  rsp_fold
-where
-  "rsp_fold f = (\<forall>u v w. (f u (f v w) = f v (f u w)))"
-
-primrec
-  fold_raw :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b"
-where
-  "fold_raw f z [] = z"
-| "fold_raw f z (a # A) =
-     (if (rsp_fold f) then
-       if a mem A then fold_raw f z A
-       else f a (fold_raw f z A)
-     else z)"
-
-lemma mem_lcommuting_fold_raw:
-  "rsp_fold f \<Longrightarrow> h mem B \<Longrightarrow> fold_raw f z B = f h (fold_raw f z (delete_raw B h))"
-sorry
-
-lemma fold_rsp[quot_respect]:
-  "(op = ===> op = ===> op \<approx> ===> op =) fold_raw fold_raw"
-apply(auto)
-sorry
-
-primrec
-  inter_raw
-where
-  "inter_raw [] B = []"
-| "inter_raw (a # A) B = (if a mem B then a # inter_raw A B else inter_raw A B)"
-
-lemma mem_inter_raw:
-  "x mem (inter_raw A B) = x mem A \<and> x mem B"
-sorry
-
-lemma inter_raw_RSP:
-  "A1 \<approx> A2 \<and> B1 \<approx> B2 \<Longrightarrow> (inter_raw A1 B1) \<approx> (inter_raw A2 B2)"
-sorry
-
-
-(* LIFTING DEFS *)
-
-
-section {* Constants on the Quotient Type *} 
-
-
-quotient_definition
-  "fdelete :: 'a fset \<Rightarrow> 'a \<Rightarrow> 'a fset" 
-  is "delete_raw"
-
-quotient_definition
-  finter ("_ \<inter>f _" [70, 71] 70)
-where
-  "finter :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset"
-  is "inter_raw"
-
-quotient_definition
-  "ffold :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" 
-  is "fold_raw"
-
-quotient_definition
-  "fset_to_set :: 'a fset \<Rightarrow> 'a set" 
-  is "set"
-
-
-section {* Lifted Theorems *}
-
-thm list.cases (* ??? *)
-
-thm cons_left_comm
-lemma "finsert a (finsert b S) = finsert b (finsert a S)"
-by (lifting cons_left_comm)
-
-thm cons_left_idem
-lemma "finsert a (finsert a S) = finsert a S"
-by (lifting cons_left_idem)
-
-(* thm MEM:
-  MEM x [] = F
-  MEM x (h::t) = (x=h) \/ MEM x t *)
-thm none_mem_nil
-(*lemma "(\<forall>a. a \<notin>f A) = (A = fempty)"*)
-
-thm mem_cons
-thm finite_set_raw_strong_cases
-(*thm card_raw.simps*)
-(*thm not_mem_card_raw*)
-(*thm card_raw_suc*)
-
-lemma "fcard X = Suc n \<Longrightarrow> (\<exists>a S. a \<notin>f S & X = finsert a S)"
-(*by (lifting card_raw_suc)*)
-sorry
-
-(*thm card_raw_cons_gt_0
-thm mem_card_raw_gt_0
-thm not_nil_equiv_cons
-*)
-thm delete_raw.simps
-(*thm mem_delete_raw*)
-(*thm card_raw_delete_raw*)
-thm cons_delete_raw
-thm mem_cons_delete_raw
-thm finite_set_raw_delete_raw_cases
-thm append.simps
-(* MEM_APPEND: MEM e (APPEND l1 l2) = MEM e l1 \/ MEM e l2 *)
-thm inter_raw.simps
-thm mem_inter_raw
-thm fold_raw.simps
-thm list2set_thm
-thm list_eq_def
-thm list.induct
-lemma "\<lbrakk>P fempty; \<And>a x. P x \<Longrightarrow> P (finsert a x)\<rbrakk> \<Longrightarrow> P l"
-by (lifting list.induct)
-
-(* We also have map and some properties of it in FSet *)
-(* and the following which still lifts ok *)
-lemma "funion (funion x xa) xb = funion x (funion xa xb)"
-by (lifting append_assoc)
-
-quotient_definition
-  "fset_case :: 'a \<Rightarrow> ('b \<Rightarrow> 'b fset \<Rightarrow> 'a) \<Rightarrow> 'b fset \<Rightarrow> 'a"
-is
-  "list_case"
-
-(* NOT SURE IF TRUE *)
-lemma list_case_rsp[quot_respect]:
-  "(op = ===> (op = ===> op \<approx> ===> op =) ===> op \<approx> ===> op =) list_case list_case"
-  apply (auto)
-  sorry
-
-lemma "fset_case (f1::'t) f2 (finsert a xa) = f2 a xa"
-apply (lifting list.cases(2))
-done
-
-end
--- a/Quot/Examples/IntEx.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,277 +0,0 @@
-theory IntEx
-imports "../Quotient_Product" "../Quotient_List"
-begin
-
-fun
-  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
-where
-  "intrel (x, y) (u, v) = (x + v = u + y)"
-
-quotient_type 
-  my_int = "nat \<times> nat" / intrel
-  apply(unfold equivp_def)
-  apply(auto simp add: mem_def expand_fun_eq)
-  done
-
-quotient_definition
-  "ZERO :: my_int"
-is
-  "(0::nat, 0::nat)"
-
-quotient_definition
-  "ONE :: my_int"
-is
-  "(1::nat, 0::nat)"
-
-fun
-  my_plus :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
-where
-  "my_plus (x, y) (u, v) = (x + u, y + v)"
-
-quotient_definition
-  "PLUS :: my_int \<Rightarrow> my_int \<Rightarrow> my_int"
-is
-  "my_plus"
-
-fun
-  my_neg :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
-where
-  "my_neg (x, y) = (y, x)"
-
-quotient_definition
-  "NEG :: my_int \<Rightarrow> my_int"
-is
-  "my_neg"
-
-definition
-  MINUS :: "my_int \<Rightarrow> my_int \<Rightarrow> my_int"
-where
-  "MINUS z w = PLUS z (NEG w)"
-
-fun
-  my_mult :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
-where
-  "my_mult (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
-
-quotient_definition
-  "MULT :: my_int \<Rightarrow> my_int \<Rightarrow> my_int"
-is
-  "my_mult"
-
-
-(* NOT SURE WETHER THIS DEFINITION IS CORRECT *)
-fun
-  my_le :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
-where
-  "my_le (x, y) (u, v) = (x+v \<le> u+y)"
-
-quotient_definition
-   "LE :: my_int \<Rightarrow> my_int \<Rightarrow> bool"
-is
-  "my_le"
-
-term LE
-thm LE_def
-
-
-definition
-  LESS :: "my_int \<Rightarrow> my_int \<Rightarrow> bool"
-where
-  "LESS z w = (LE z w \<and> z \<noteq> w)"
-
-term LESS
-thm LESS_def
-
-definition
-  ABS :: "my_int \<Rightarrow> my_int"
-where
-  "ABS i = (if (LESS i ZERO) then (NEG i) else i)"
-
-definition
-  SIGN :: "my_int \<Rightarrow> my_int"
-where
- "SIGN i = (if i = ZERO then ZERO else if (LESS ZERO i) then ONE else (NEG ONE))"
-
-print_quotconsts
-
-lemma plus_sym_pre:
-  shows "my_plus a b \<approx> my_plus b a"
-  apply(cases a)
-  apply(cases b)
-  apply(auto)
-  done
-
-lemma plus_rsp[quot_respect]:
-  shows "(intrel ===> intrel ===> intrel) my_plus my_plus"
-by (simp)
-
-lemma neg_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx>) my_neg my_neg"
-by simp
-
-lemma test1: "my_plus a b = my_plus a b"
-apply(rule refl)
-done
-
-lemma "PLUS a b = PLUS a b"
-apply(lifting_setup test1)
-apply(regularize)
-apply(injection)
-apply(cleaning)
-done
-
-thm lambda_prs
-
-lemma test2: "my_plus a = my_plus a"
-apply(rule refl)
-done
-
-lemma "PLUS a = PLUS a"
-apply(lifting_setup test2)
-apply(rule impI)
-apply(rule ballI)
-apply(rule apply_rsp[OF Quotient_my_int plus_rsp])
-apply(simp only: in_respects)
-apply(injection)
-apply(cleaning)
-done
-
-lemma test3: "my_plus = my_plus"
-apply(rule refl)
-done
-
-lemma "PLUS = PLUS"
-apply(lifting_setup test3)
-apply(rule impI)
-apply(rule plus_rsp)
-apply(injection)
-apply(cleaning)
-done
-
-
-lemma "PLUS a b = PLUS b a"
-apply(lifting plus_sym_pre)
-done
-
-lemma plus_assoc_pre:
-  shows "my_plus (my_plus i j) k \<approx> my_plus i (my_plus j k)"
-  apply (cases i)
-  apply (cases j)
-  apply (cases k)
-  apply (simp)
-  done
-
-lemma plus_assoc: "PLUS (PLUS x xa) xb = PLUS x (PLUS xa xb)"
-apply(lifting plus_assoc_pre)
-done
-
-lemma int_induct_raw:
-  assumes a: "P (0::nat, 0)"
-  and     b: "\<And>i. P i \<Longrightarrow> P (my_plus i (1,0))"
-  and     c: "\<And>i. P i \<Longrightarrow> P (my_plus i (my_neg (1,0)))"
-  shows      "P x"
-  apply(case_tac x) apply(simp)
-  apply(rule_tac x="b" in spec)
-  apply(rule_tac Nat.induct)
-  apply(rule allI)
-  apply(rule_tac Nat.induct)
-  using a b c apply(auto)
-  done
-
-lemma int_induct:
-  assumes a: "P ZERO"
-  and     b: "\<And>i. P i \<Longrightarrow> P (PLUS i ONE)"
-  and     c: "\<And>i. P i \<Longrightarrow> P (PLUS i (NEG ONE))"
-  shows      "P x"
-  using a b c
-  by (lifting int_induct_raw)
-
-lemma ex1tst: "Bex1_rel (op \<approx>) (\<lambda>x :: nat \<times> nat. x \<approx> x)"
-sorry
-
-lemma ex1tst': "\<exists>!(x::my_int). x = x"
-apply(lifting ex1tst)
-done
-
-
-lemma ho_tst: "foldl my_plus x [] = x"
-apply simp
-done
-
-
-term foldl
-lemma "foldl PLUS x [] = x"
-apply(lifting ho_tst)
-done
-
-lemma ho_tst2: "foldl my_plus x (h # t) \<approx> my_plus h (foldl my_plus x t)"
-sorry
-
-lemma "foldl PLUS x (h # t) = PLUS h (foldl PLUS x t)"
-apply(lifting ho_tst2)
-done
-
-lemma ho_tst3: "foldl f (s::nat \<times> nat) ([]::(nat \<times> nat) list) = s"
-by simp
-
-lemma "foldl f (x::my_int) ([]::my_int list) = x"
-apply(lifting ho_tst3)
-done
-
-lemma lam_tst: "(\<lambda>x. (x, x)) y = (y, (y :: nat \<times> nat))"
-by simp
-
-(* Don't know how to keep the goal non-contracted... *)
-lemma "(\<lambda>x. (x, x)) (y::my_int) = (y, y)"
-apply(lifting lam_tst)
-done
-
-lemma lam_tst2: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
-by simp
-
-lemma
-  shows "equivp (op \<approx>)"
-  and "equivp ((op \<approx>) ===> (op \<approx>))"
-(* Nitpick finds a counterexample! *)
-oops
-
-lemma lam_tst3a: "(\<lambda>(y :: nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat). x)"
-by auto
-
-lemma id_rsp:
-  shows "(R ===> R) id id"
-by simp
-
-lemma lam_tst3a_reg: "(op \<approx> ===> op \<approx>) (Babs (Respects op \<approx>) (\<lambda>y. y)) (Babs (Respects op \<approx>) (\<lambda>x. x))"
-apply (rule babs_rsp[OF Quotient_my_int])
-apply (simp add: id_rsp)
-done
-
-lemma "(\<lambda>(y :: my_int). y) = (\<lambda>(x :: my_int). x)"
-apply(lifting lam_tst3a)
-apply(rule impI)
-apply(rule lam_tst3a_reg)
-done
-
-lemma lam_tst3b: "(\<lambda>(y :: nat \<times> nat \<Rightarrow> nat \<times> nat). y) = (\<lambda>(x :: nat \<times> nat \<Rightarrow> nat \<times> nat). x)"
-by auto
-
-lemma "(\<lambda>(y :: my_int => my_int). y) = (\<lambda>(x :: my_int => my_int). x)"
-apply(lifting lam_tst3b)
-apply(rule impI)
-apply(rule babs_rsp[OF fun_quotient[OF Quotient_my_int Quotient_my_int]])
-apply(simp add: id_rsp)
-done
-
-lemma lam_tst4: "map (\<lambda>x. my_plus x (0,0)) l = l"
-apply (induct l)
-apply simp
-apply (case_tac a)
-apply simp
-done
-
-lemma "map (\<lambda>x. PLUS x ZERO) l = l"
-apply(lifting lam_tst4)
-done
-
-end
--- a/Quot/Examples/IntEx2.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,445 +0,0 @@
-theory IntEx2
-imports "../Quotient" "../Quotient_Product" Nat
-(*uses
-  ("Tools/numeral.ML")
-  ("Tools/numeral_syntax.ML")
-  ("Tools/int_arith.ML")*)
-begin
-
-fun
-  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" (infix "\<approx>" 50)
-where
-  "intrel (x, y) (u, v) = (x + v = u + y)"
-
-quotient_type int = "nat \<times> nat" / intrel
-  unfolding equivp_def
-  by (auto simp add: mem_def expand_fun_eq)
-
-instantiation int :: "{zero, one, plus, uminus, minus, times, ord, abs, sgn}"
-begin
-
-ML {* @{term "0 \<Colon> int"} *}
-
-quotient_definition
-  "0 \<Colon> int" is "(0\<Colon>nat, 0\<Colon>nat)"
-
-quotient_definition
-  "1 \<Colon> int" is "(1\<Colon>nat, 0\<Colon>nat)"
-
-fun
-  plus_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
-where
-  "plus_raw (x, y) (u, v) = (x + u, y + v)"
-
-quotient_definition
-  "(op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int)" is "plus_raw"
-
-fun
-  uminus_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
-where
-  "uminus_raw (x, y) = (y, x)"
-
-quotient_definition
-  "(uminus \<Colon> (int \<Rightarrow> int))" is "uminus_raw"
-
-definition
-  minus_int_def [code del]:  "z - w = z + (-w\<Colon>int)"
-
-fun
-  mult_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"
-where
-  "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
-
-quotient_definition
-  mult_int_def: "(op *) :: (int \<Rightarrow> int \<Rightarrow> int)" is "mult_raw"
-
-fun
-  le_raw :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool"
-where
-  "le_raw (x, y) (u, v) = (x+v \<le> u+y)"
-
-quotient_definition
-  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" is "le_raw"
-
-definition
-  less_int_def [code del]: "(z\<Colon>int) < w = (z \<le> w \<and> z \<noteq> w)"
-
-definition
-  zabs_def: "\<bar>i\<Colon>int\<bar> = (if i < 0 then - i else i)"
-
-definition
-  zsgn_def: "sgn (i\<Colon>int) = (if i = 0 then 0 else if 0 < i then 1 else - 1)"
-
-instance ..
-
-end
-
-lemma plus_raw_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) plus_raw plus_raw"
-by auto
-
-lemma uminus_raw_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx>) uminus_raw uminus_raw"
-  by auto
-
-lemma mult_raw_fst:
-  assumes a: "x \<approx> z"
-  shows "mult_raw x y \<approx> mult_raw z y"
-using a
-apply(cases x, cases y, cases z)
-apply(auto simp add: mult_raw.simps intrel.simps)
-apply(rename_tac u v w x y z)
-apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
-apply(simp add: mult_ac)
-apply(simp add: add_mult_distrib [symmetric])
-done
-
-lemma mult_raw_snd:
-  assumes a: "x \<approx> z"
-  shows "mult_raw y x \<approx> mult_raw y z"
-using a
-apply(cases x, cases y, cases z)
-apply(auto simp add: mult_raw.simps intrel.simps)
-apply(rename_tac u v w x y z)
-apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
-apply(simp add: mult_ac)
-apply(simp add: add_mult_distrib [symmetric])
-done
-
-lemma mult_raw_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx> ===> op \<approx>) mult_raw mult_raw"
-apply(simp only: fun_rel_def)
-apply(rule allI | rule impI)+
-apply(rule equivp_transp[OF int_equivp])
-apply(rule mult_raw_fst)
-apply(assumption)
-apply(rule mult_raw_snd)
-apply(assumption)
-done
-
-lemma le_raw_rsp[quot_respect]:
-  shows "(op \<approx> ===> op \<approx> ===> op =) le_raw le_raw"
-by auto
-
-lemma plus_assoc_raw:
-  shows "plus_raw (plus_raw i j) k \<approx> plus_raw i (plus_raw j k)"
-by (cases i, cases j, cases k) (simp)
-
-lemma plus_sym_raw:
-  shows "plus_raw i j \<approx> plus_raw j i"
-by (cases i, cases j) (simp)
-
-lemma plus_zero_raw:
-  shows "plus_raw  (0, 0) i \<approx> i"
-by (cases i) (simp)
-
-lemma plus_minus_zero_raw:
-  shows "plus_raw (uminus_raw i) i \<approx> (0, 0)"
-by (cases i) (simp)
-
-lemma times_assoc_raw:
-  shows "mult_raw (mult_raw i j) k \<approx> mult_raw i (mult_raw j k)"
-by (cases i, cases j, cases k) 
-   (simp add: algebra_simps)
-
-lemma times_sym_raw:
-  shows "mult_raw i j \<approx> mult_raw j i"
-by (cases i, cases j) (simp add: algebra_simps)
-
-lemma times_one_raw:
-  shows "mult_raw  (1, 0) i \<approx> i"
-by (cases i) (simp)
-
-lemma times_plus_comm_raw:
-  shows "mult_raw (plus_raw i j) k \<approx> plus_raw (mult_raw i k) (mult_raw j k)"
-by (cases i, cases j, cases k) 
-   (simp add: algebra_simps)
-
-lemma one_zero_distinct:
-  shows "\<not> (0, 0) \<approx> ((1::nat), (0::nat))"
-  by simp
-
-text{* The integers form a @{text comm_ring_1}*}
-
-instance int :: comm_ring_1
-proof
-  fix i j k :: int
-  show "(i + j) + k = i + (j + k)"
-    by (lifting plus_assoc_raw)
-  show "i + j = j + i" 
-    by (lifting plus_sym_raw)
-  show "0 + i = (i::int)"
-    by (lifting plus_zero_raw)
-  show "- i + i = 0"
-    by (lifting plus_minus_zero_raw)
-  show "i - j = i + - j"
-    by (simp add: minus_int_def)
-  show "(i * j) * k = i * (j * k)"
-    by (lifting times_assoc_raw)
-  show "i * j = j * i"
-    by (lifting times_sym_raw)
-  show "1 * i = i"
-    by (lifting times_one_raw)
-  show "(i + j) * k = i * k + j * k"
-    by (lifting times_plus_comm_raw)
-  show "0 \<noteq> (1::int)"
-    by (lifting one_zero_distinct)
-qed
-
-lemma plus_raw_rsp_aux:
-  assumes a: "a \<approx> b" "c \<approx> d"
-  shows "plus_raw a c \<approx> plus_raw b d"
-using a
-by (cases a, cases b, cases c, cases d) 
-   (simp)
-
-lemma add:
-     "(abs_int (x,y)) + (abs_int (u,v)) =
-      (abs_int (x + u, y + v))"
-apply(simp add: plus_int_def id_simps)
-apply(fold plus_raw.simps)
-apply(rule Quotient_rel_abs[OF Quotient_int])
-apply(rule plus_raw_rsp_aux)
-apply(simp_all add: rep_abs_rsp_left[OF Quotient_int])
-done
-
-definition int_of_nat_raw: 
-  "int_of_nat_raw m = (m :: nat, 0 :: nat)"
-
-quotient_definition
-  "int_of_nat :: nat \<Rightarrow> int" is "int_of_nat_raw"
-
-lemma[quot_respect]: 
-  shows "(op = ===> op \<approx>) int_of_nat_raw int_of_nat_raw"
-by (simp add: equivp_reflp[OF int_equivp])
-
-lemma int_of_nat:
-  shows "of_nat m = int_of_nat m"
-apply (induct m)
-apply (simp_all add: zero_int_def one_int_def int_of_nat_def int_of_nat_raw add)
-done
-
-lemma le_antisym_raw:
-  shows "le_raw i j \<Longrightarrow> le_raw j i \<Longrightarrow> i \<approx> j"
-by (cases i, cases j) (simp)
-
-lemma le_refl_raw:
-  shows "le_raw i i"
-by (cases i) (simp)
-
-lemma le_trans_raw:
-  shows "le_raw i j \<Longrightarrow> le_raw j k \<Longrightarrow> le_raw i k"
-by (cases i, cases j, cases k) (simp)
-
-lemma le_cases_raw:
-  shows "le_raw i j \<or> le_raw j i"
-by (cases i, cases j) 
-   (simp add: linorder_linear)
-
-instance int :: linorder
-proof
-  fix i j k :: int
-  show antisym: "i \<le> j \<Longrightarrow> j \<le> i \<Longrightarrow> i = j"
-    by (lifting le_antisym_raw)
-  show "(i < j) = (i \<le> j \<and> \<not> j \<le> i)"
-    by (auto simp add: less_int_def dest: antisym) 
-  show "i \<le> i"
-    by (lifting le_refl_raw)
-  show "i \<le> j \<Longrightarrow> j \<le> k \<Longrightarrow> i \<le> k"
-    by (lifting le_trans_raw)
-  show "i \<le> j \<or> j \<le> i"
-    by (lifting le_cases_raw)
-qed
-
-instantiation int :: distrib_lattice
-begin
-
-definition
-  "(inf \<Colon> int \<Rightarrow> int \<Rightarrow> int) = min"
-
-definition
-  "(sup \<Colon> int \<Rightarrow> int \<Rightarrow> int) = max"
-
-instance
-  by intro_classes
-    (auto simp add: inf_int_def sup_int_def min_max.sup_inf_distrib1)
-
-end
-
-lemma le_plus_raw:
-  shows "le_raw i j \<Longrightarrow> le_raw (plus_raw k i) (plus_raw k j)"
-by (cases i, cases j, cases k) (simp)
-
-
-instance int :: ordered_cancel_ab_semigroup_add
-proof
-  fix i j k :: int
-  show "i \<le> j \<Longrightarrow> k + i \<le> k + j"
-    by (lifting le_plus_raw)
-qed
-
-abbreviation
-  "less_raw i j \<equiv> le_raw i j \<and> \<not>(i \<approx> j)"
-
-lemma zmult_zless_mono2_lemma:
-  fixes i j::int
-  and   k::nat
-  shows "i < j \<Longrightarrow> 0 < k \<Longrightarrow> of_nat k * i < of_nat k * j"
-apply(induct "k")
-apply(simp)
-apply(case_tac "k = 0")
-apply(simp_all add: left_distrib add_strict_mono)
-done
-
-lemma zero_le_imp_eq_int_raw:
-  fixes k::"(nat \<times> nat)"
-  shows "less_raw (0, 0) k \<Longrightarrow> (\<exists>n > 0. k \<approx> int_of_nat_raw n)"
-apply(cases k)
-apply(simp add:int_of_nat_raw)
-apply(auto)
-apply(rule_tac i="b" and j="a" in less_Suc_induct)
-apply(auto)
-done
-
-lemma zero_le_imp_eq_int:
-  fixes k::int
-  shows "0 < k \<Longrightarrow> \<exists>n > 0. k = of_nat n"
-  unfolding less_int_def int_of_nat
-  by (lifting zero_le_imp_eq_int_raw)
-
-lemma zmult_zless_mono2: 
-  fixes i j k::int
-  assumes a: "i < j" "0 < k"
-  shows "k * i < k * j"
-using a
-using a
-apply(drule_tac zero_le_imp_eq_int)
-apply(auto simp add: zmult_zless_mono2_lemma)
-done
-
-text{*The integers form an ordered integral domain*}
-instance int :: linordered_idom
-proof
-  fix i j k :: int
-  show "i < j \<Longrightarrow> 0 < k \<Longrightarrow> k * i < k * j"
-    by (rule zmult_zless_mono2)
-  show "\<bar>i\<bar> = (if i < 0 then -i else i)"
-    by (simp only: zabs_def)
-  show "sgn (i\<Colon>int) = (if i=0 then 0 else if 0<i then 1 else - 1)"
-    by (simp only: zsgn_def)
-qed
-
-lemmas int_distrib =
-  left_distrib [of "z1::int" "z2" "w", standard]
-  right_distrib [of "w::int" "z1" "z2", standard]
-  left_diff_distrib [of "z1::int" "z2" "w", standard]
-  right_diff_distrib [of "w::int" "z1" "z2", standard]
-
-
-subsection {* Embedding of the Integers into any @{text ring_1}: @{text of_int}*}
-
-(*
-context ring_1
-begin
-
- 
-definition 
-  of_int :: "int \<Rightarrow> 'a" 
-where
-  "of_int 
-*)
-
-
-subsection {* Binary representation *}
-
-text {*
-  This formalization defines binary arithmetic in terms of the integers
-  rather than using a datatype. This avoids multiple representations (leading
-  zeroes, etc.)  See @{text "ZF/Tools/twos-compl.ML"}, function @{text
-  int_of_binary}, for the numerical interpretation.
-
-  The representation expects that @{text "(m mod 2)"} is 0 or 1,
-  even if m is negative;
-  For instance, @{text "-5 div 2 = -3"} and @{text "-5 mod 2 = 1"}; thus
-  @{text "-5 = (-3)*2 + 1"}.
-  
-  This two's complement binary representation derives from the paper 
-  "An Efficient Representation of Arithmetic for Term Rewriting" by
-  Dave Cohen and Phil Watson, Rewriting Techniques and Applications,
-  Springer LNCS 488 (240-251), 1991.
-*}
-
-subsubsection {* The constructors @{term Bit0}, @{term Bit1}, @{term Pls} and @{term Min} *}
-
-definition
-  Pls :: int where
-  [code del]: "Pls = 0"
-
-definition
-  Min :: int where
-  [code del]: "Min = - 1"
-
-definition
-  Bit0 :: "int \<Rightarrow> int" where
-  [code del]: "Bit0 k = k + k"
-
-definition
-  Bit1 :: "int \<Rightarrow> int" where
-  [code del]: "Bit1 k = 1 + k + k"
-
-class number = -- {* for numeric types: nat, int, real, \dots *}
-  fixes number_of :: "int \<Rightarrow> 'a"
-
-(*use "~~/src/HOL/Tools/numeral.ML"
-
-syntax
-  "_Numeral" :: "num_const \<Rightarrow> 'a"    ("_")
-
-use "~~/src/HOL/Tools/numeral_syntax.ML"
-
-setup NumeralSyntax.setup
-
-abbreviation
-  "Numeral0 \<equiv> number_of Pls"
-
-abbreviation
-  "Numeral1 \<equiv> number_of (Bit1 Pls)"
-
-lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)"
-  -- {* Unfold all @{text let}s involving constants *}
-  unfolding Let_def ..
-
-definition
-  succ :: "int \<Rightarrow> int" where
-  [code del]: "succ k = k + 1"
-
-definition
-  pred :: "int \<Rightarrow> int" where
-  [code del]: "pred k = k - 1"
-
-lemmas
-  max_number_of [simp] = max_def
-    [of "number_of u" "number_of v", standard, simp]
-and
-  min_number_of [simp] = min_def 
-    [of "number_of u" "number_of v", standard, simp]
-  -- {* unfolding @{text minx} and @{text max} on numerals *}
-
-lemmas numeral_simps = 
-  succ_def pred_def Pls_def Min_def Bit0_def Bit1_def
-
-text {* Removal of leading zeroes *}
-
-lemma Bit0_Pls [simp, code_post]:
-  "Bit0 Pls = Pls"
-  unfolding numeral_simps by simp
-
-lemma Bit1_Min [simp, code_post]:
-  "Bit1 Min = Min"
-  unfolding numeral_simps by simp
-
-lemmas normalize_bin_simps =
-  Bit0_Pls Bit1_Min
-*)
-
-end
--- a/Quot/Examples/LFex.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,297 +0,0 @@
-theory LFex
-imports Nominal "../Quotient_List"
-begin
-
-atom_decl name ident
-
-nominal_datatype kind = 
-    Type
-  | KPi "ty" "name" "kind"
-and ty =  
-    TConst "ident"
-  | TApp "ty" "trm"
-  | TPi "ty" "name" "ty"
-and trm = 
-    Const "ident"
-  | Var "name"
-  | App "trm" "trm"
-  | Lam "ty" "name" "trm" 
-
-function
-    fv_kind :: "kind \<Rightarrow> name set"
-and fv_ty   :: "ty \<Rightarrow> name set"
-and fv_trm  :: "trm \<Rightarrow> name set"
-where
-  "fv_kind (Type) = {}"
-| "fv_kind (KPi A x K) = (fv_ty A) \<union> ((fv_kind K) - {x})"
-| "fv_ty (TConst i) = {}"
-| "fv_ty (TApp A M) = (fv_ty A) \<union> (fv_trm M)"
-| "fv_ty (TPi A x B) = (fv_ty A) \<union> ((fv_ty B) - {x})"
-| "fv_trm (Const i) = {}"
-| "fv_trm (Var x) = {x}"
-| "fv_trm (App M N) = (fv_trm M) \<union> (fv_trm N)"
-| "fv_trm (Lam A x M) = (fv_ty A) \<union> ((fv_trm M) - {x})"
-sorry
-
-termination fv_kind sorry
-
-inductive
-    akind :: "kind \<Rightarrow> kind \<Rightarrow> bool" ("_ \<approx>ki _" [100, 100] 100)
-and aty   :: "ty \<Rightarrow> ty \<Rightarrow> bool"     ("_ \<approx>ty _" [100, 100] 100)
-and atrm  :: "trm \<Rightarrow> trm \<Rightarrow> bool"   ("_ \<approx>tr _" [100, 100] 100)
-where
-  a1:  "(Type) \<approx>ki (Type)"
-| a21: "\<lbrakk>A \<approx>ty A'; K \<approx>ki K'\<rbrakk> \<Longrightarrow> (KPi A x K) \<approx>ki (KPi A' x K')"
-| a22: "\<lbrakk>A \<approx>ty A'; K \<approx>ki ([(x,x')]\<bullet>K'); x \<notin> (fv_ty A'); x \<notin> ((fv_kind K') - {x'})\<rbrakk> 
-        \<Longrightarrow> (KPi A x K) \<approx>ki (KPi A' x' K')"
-| a3:  "i = j \<Longrightarrow> (TConst i) \<approx>ty (TConst j)"
-| a4:  "\<lbrakk>A \<approx>ty A'; M \<approx>tr M'\<rbrakk> \<Longrightarrow> (TApp A M) \<approx>ty (TApp A' M')"
-| a51: "\<lbrakk>A \<approx>ty A'; B \<approx>ty B'\<rbrakk> \<Longrightarrow> (TPi A x B) \<approx>ty (TPi A' x B')"
-| a52: "\<lbrakk>A \<approx>ty A'; B \<approx>ty ([(x,x')]\<bullet>B'); x \<notin> (fv_ty B'); x \<notin> ((fv_ty B') - {x'})\<rbrakk> 
-        \<Longrightarrow> (TPi A x B) \<approx>ty (TPi A' x' B')"
-| a6:  "i = j \<Longrightarrow> (Const i) \<approx>trm (Const j)"
-| a7:  "x = y \<Longrightarrow> (Var x) \<approx>trm (Var y)"
-| a8:  "\<lbrakk>M \<approx>trm M'; N \<approx>tr N'\<rbrakk> \<Longrightarrow> (App M N) \<approx>tr (App M' N')"
-| a91: "\<lbrakk>A \<approx>ty A'; M \<approx>tr M'\<rbrakk> \<Longrightarrow> (Lam A x M) \<approx>tr (Lam A' x M')"
-| a92: "\<lbrakk>A \<approx>ty A'; M \<approx>tr ([(x,x')]\<bullet>M'); x \<notin> (fv_ty B'); x \<notin> ((fv_trm M') - {x'})\<rbrakk> 
-        \<Longrightarrow> (Lam A x M) \<approx>tr (Lam A' x' M')"
-
-lemma al_refl:
-  fixes K::"kind" 
-  and   A::"ty"
-  and   M::"trm"
-  shows "K \<approx>ki K"
-  and   "A \<approx>ty A"
-  and   "M \<approx>tr M"
-  apply(induct K and A and M rule: kind_ty_trm.inducts)
-  apply(auto intro: akind_aty_atrm.intros)
-  done
-
-lemma alpha_equivps:
-  shows "equivp akind"
-  and   "equivp aty"
-  and   "equivp atrm"
-sorry
-
-quotient_type KIND = kind / akind
-  by (rule alpha_equivps)
-
-quotient_type 
-    TY = ty / aty and   
-    TRM = trm / atrm
-  by (auto intro: alpha_equivps)
-
-quotient_definition
-   "TYP :: KIND"
-is
-  "Type"
-
-quotient_definition
-   "KPI :: TY \<Rightarrow> name \<Rightarrow> KIND \<Rightarrow> KIND"
-is
-  "KPi"
-
-quotient_definition
-   "TCONST :: ident \<Rightarrow> TY"
-is
-  "TConst"
-
-quotient_definition
-   "TAPP :: TY \<Rightarrow> TRM \<Rightarrow> TY"
-is
-  "TApp"
-
-quotient_definition
-   "TPI :: TY \<Rightarrow> name \<Rightarrow> TY \<Rightarrow> TY"
-is
-  "TPi"
-
-(* FIXME: does not work with CONST *)
-quotient_definition
-   "CONS :: ident \<Rightarrow> TRM"
-is
-  "Const"
-
-quotient_definition
-   "VAR :: name \<Rightarrow> TRM"
-is
-  "Var"
-
-quotient_definition
-   "APP :: TRM \<Rightarrow> TRM \<Rightarrow> TRM"
-is
-  "App"
-
-quotient_definition
-   "LAM :: TY \<Rightarrow> name \<Rightarrow> TRM \<Rightarrow> TRM"
-is
-  "Lam"
-
-thm TYP_def
-thm KPI_def
-thm TCONST_def
-thm TAPP_def
-thm TPI_def
-thm VAR_def
-thm CONS_def
-thm APP_def
-thm LAM_def
-
-(* FIXME: print out a warning if the type contains a liftet type, like kind \<Rightarrow> name set *)
-quotient_definition
-   "FV_kind :: KIND \<Rightarrow> name set"
-is
-  "fv_kind"
-
-quotient_definition
-   "FV_ty :: TY \<Rightarrow> name set"
-is
-  "fv_ty"
-
-quotient_definition
-   "FV_trm :: TRM \<Rightarrow> name set"
-is
-  "fv_trm"
-
-thm FV_kind_def
-thm FV_ty_def
-thm FV_trm_def
-
-(* FIXME: does not work yet *)
-overloading
-    perm_kind \<equiv> "perm :: 'x prm \<Rightarrow> KIND \<Rightarrow> KIND"   (unchecked)
-    perm_ty   \<equiv> "perm :: 'x prm \<Rightarrow> TY \<Rightarrow> TY"       (unchecked)
-    perm_trm  \<equiv> "perm :: 'x prm \<Rightarrow> TRM \<Rightarrow> TRM"     (unchecked) 
-begin
-
-quotient_definition
-   "perm_kind :: 'x prm \<Rightarrow> KIND \<Rightarrow> KIND"
-is
-  "(perm::'x prm \<Rightarrow> kind \<Rightarrow> kind)"
-
-quotient_definition
-   "perm_ty :: 'x prm \<Rightarrow> TY \<Rightarrow> TY"
-is
-  "(perm::'x prm \<Rightarrow> ty \<Rightarrow> ty)"
-
-quotient_definition
-   "perm_trm :: 'x prm \<Rightarrow> TRM \<Rightarrow> TRM"
-is
-  "(perm::'x prm \<Rightarrow> trm \<Rightarrow> trm)"
-
-end
-
-(* TODO/FIXME: Think whether these RSP theorems are true. *)
-lemma kpi_rsp[quot_respect]: 
-  "(aty ===> op = ===> akind ===> akind) KPi KPi" sorry
-lemma tconst_rsp[quot_respect]: 
-  "(op = ===> aty) TConst TConst" sorry
-lemma tapp_rsp[quot_respect]: 
-  "(aty ===> atrm ===> aty) TApp TApp" sorry
-lemma tpi_rsp[quot_respect]: 
-  "(aty ===> op = ===> aty ===> aty) TPi TPi" sorry
-lemma var_rsp[quot_respect]: 
-  "(op = ===> atrm) Var Var" sorry
-lemma app_rsp[quot_respect]: 
-  "(atrm ===> atrm ===> atrm) App App" sorry
-lemma const_rsp[quot_respect]: 
-  "(op = ===> atrm) Const Const" sorry
-lemma lam_rsp[quot_respect]: 
-  "(aty ===> op = ===> atrm ===> atrm) Lam Lam" sorry
-
-lemma perm_kind_rsp[quot_respect]: 
-  "(op = ===> akind ===> akind) op \<bullet> op \<bullet>" sorry
-lemma perm_ty_rsp[quot_respect]: 
-  "(op = ===> aty ===> aty) op \<bullet> op \<bullet>" sorry
-lemma perm_trm_rsp[quot_respect]: 
-  "(op = ===> atrm ===> atrm) op \<bullet> op \<bullet>" sorry
-
-lemma fv_ty_rsp[quot_respect]: 
-  "(aty ===> op =) fv_ty fv_ty" sorry
-lemma fv_kind_rsp[quot_respect]: 
-  "(akind ===> op =) fv_kind fv_kind" sorry
-lemma fv_trm_rsp[quot_respect]: 
-  "(atrm ===> op =) fv_trm fv_trm" sorry
-
-
-thm akind_aty_atrm.induct
-thm kind_ty_trm.induct
-
-
-lemma 
-  assumes a0:
-  "P1 TYP TYP"
-  and a1: 
-  "\<And>A A' K K' x. \<lbrakk>(A::TY) = A'; P2 A A'; (K::KIND) = K'; P1 K K'\<rbrakk> 
-  \<Longrightarrow> P1 (KPI A x K) (KPI A' x K')"
-  and a2:    
-  "\<And>A A' K K' x x'. \<lbrakk>(A ::TY) = A'; P2 A A'; (K :: KIND) = ([(x, x')] \<bullet> K'); P1 K ([(x, x')] \<bullet> K'); 
-    x \<notin> FV_ty A'; x \<notin> FV_kind K' - {x'}\<rbrakk> \<Longrightarrow> P1 (KPI A x K) (KPI A' x' K')"
-  and a3: 
-  "\<And>i j. i = j \<Longrightarrow> P2 (TCONST i) (TCONST j)"
-  and a4:
-  "\<And>A A' M M'. \<lbrakk>(A ::TY) = A'; P2 A A'; (M :: TRM) = M'; P3 M M'\<rbrakk> \<Longrightarrow> P2 (TAPP A M) (TAPP A' M')"
-  and a5:
-  "\<And>A A' B B' x. \<lbrakk>(A ::TY) = A'; P2 A A'; (B ::TY) = B'; P2 B B'\<rbrakk> \<Longrightarrow> P2 (TPI A x B) (TPI A' x B')"
-  and a6:
-  "\<And>A A' B x x' B'. \<lbrakk>(A ::TY) = A'; P2 A A'; (B ::TY) = ([(x, x')] \<bullet> B'); P2 B ([(x, x')] \<bullet> B'); 
-  x \<notin> FV_ty B'; x \<notin> FV_ty B' - {x'}\<rbrakk> \<Longrightarrow> P2 (TPI A x B) (TPI A' x' B')"
-  and a7:
-  "\<And>i j m. i = j \<Longrightarrow> P3 (CONS i) (m (CONS j))"
-  and a8:
-  "\<And>x y m. x = y \<Longrightarrow> P3 (VAR x) (m (VAR y))"
-  and a9:
-  "\<And>M m M' N N'. \<lbrakk>(M :: TRM) = m M'; P3 M (m M'); (N :: TRM) = N'; P3 N N'\<rbrakk> \<Longrightarrow> P3 (APP M N) (APP M' N')"
-  and a10: 
-  "\<And>A A' M M' x. \<lbrakk>(A ::TY) = A'; P2 A A'; (M :: TRM) = M'; P3 M M'\<rbrakk> \<Longrightarrow> P3 (LAM A x M) (LAM A' x M')"
-  and a11:
-  "\<And>A A' M x x' M' B'. \<lbrakk>(A ::TY) = A'; P2 A A'; (M :: TRM) = ([(x, x')] \<bullet> M'); P3 M ([(x, x')] \<bullet> M'); 
-  x \<notin> FV_ty B'; x \<notin> FV_trm M' - {x'}\<rbrakk> \<Longrightarrow> P3 (LAM A x M) (LAM A' x' M')"
-  shows "((x1 :: KIND) = x2 \<longrightarrow> P1 x1 x2) \<and>
-         ((x3 ::TY) = x4 \<longrightarrow> P2 x3 x4) \<and> 
-         ((x5 :: TRM) = x6 \<longrightarrow> P3 x5 x6)"
-using a0 a1 a2 a3 a4 a5 a6 a7 a8 a9 a10 a11
-apply(lifting akind_aty_atrm.induct)
-(*
-Profiling:
-ML_prf {* fun ith i =  (#concl (fst (Subgoal.focus @{context} i (#goal (Isar.goal ()))))) *}
-ML_prf {* profile 2 Seq.list_of ((clean_tac @{context} quot defs 1) (ith 3)) *}
-ML_prf {* profile 2 Seq.list_of ((regularize_tac @{context} @{thms alpha_equivps} 1) (ith 1)) *}
-ML_prf {* PolyML.profiling 1 *}
-ML_prf {* profile 2 Seq.list_of ((all_inj_repabs_tac @{context} quot rel_refl trans2 1) (#goal (Isar.goal ()))) *}
-*)
- done
-
-(* Does not work:
-lemma
-  assumes a0: "P1 TYP"
-  and     a1: "\<And>ty name kind. \<lbrakk>P2 ty; P1 kind\<rbrakk> \<Longrightarrow> P1 (KPI ty name kind)"
-  and     a2: "\<And>id. P2 (TCONST id)"
-  and     a3: "\<And>ty trm. \<lbrakk>P2 ty; P3 trm\<rbrakk> \<Longrightarrow> P2 (TAPP ty trm)"
-  and     a4: "\<And>ty1 name ty2. \<lbrakk>P2 ty1; P2 ty2\<rbrakk> \<Longrightarrow> P2 (TPI ty1 name ty2)"
-  and     a5: "\<And>id. P3 (CONS id)"
-  and     a6: "\<And>name. P3 (VAR name)"
-  and     a7: "\<And>trm1 trm2. \<lbrakk>P3 trm1; P3 trm2\<rbrakk> \<Longrightarrow> P3 (APP trm1 trm2)"
-  and     a8: "\<And>ty name trm. \<lbrakk>P2 ty; P3 trm\<rbrakk> \<Longrightarrow> P3 (LAM ty name trm)"
-  shows "P1 mkind \<and> P2 mty \<and> P3 mtrm"
-using a0 a1 a2 a3 a4 a5 a6 a7 a8
-*)
-
-
-lemma "\<lbrakk>P TYP;
-  \<And>ty name kind. \<lbrakk>Q ty; P kind\<rbrakk> \<Longrightarrow> P (KPI ty name kind);
-  \<And>id. Q (TCONST id);
-  \<And>ty trm. \<lbrakk>Q ty; R trm\<rbrakk> \<Longrightarrow> Q (TAPP ty trm);
-  \<And>ty1 name ty2. \<lbrakk>Q ty1; Q ty2\<rbrakk> \<Longrightarrow> Q (TPI ty1 name ty2);
-  \<And>id. R (CONS id); \<And>name. R (VAR name);
-  \<And>trm1 trm2. \<lbrakk>R trm1; R trm2\<rbrakk> \<Longrightarrow> R (APP trm1 trm2);
-  \<And>ty name trm. \<lbrakk>Q ty; R trm\<rbrakk> \<Longrightarrow> R (LAM ty name trm)\<rbrakk>
-  \<Longrightarrow> P mkind \<and> Q mty \<and> R mtrm"
-apply(lifting kind_ty_trm.induct)
-done
-
-end
-
-
-
-
--- a/Quot/Examples/LamEx.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,636 +0,0 @@
-theory LamEx
-imports Nominal "../Quotient" "../Quotient_List"
-begin
-
-atom_decl name
-
-datatype rlam =
-  rVar "name"
-| rApp "rlam" "rlam"
-| rLam "name" "rlam"
-
-fun
-  rfv :: "rlam \<Rightarrow> name set"
-where
-  rfv_var: "rfv (rVar a) = {a}"
-| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
-| rfv_lam: "rfv (rLam a t) = (rfv t) - {a}"
-
-overloading
-  perm_rlam \<equiv> "perm :: 'x prm \<Rightarrow> rlam \<Rightarrow> rlam"   (unchecked)
-begin
-
-fun
-  perm_rlam
-where
-  "perm_rlam pi (rVar a) = rVar (pi \<bullet> a)"
-| "perm_rlam pi (rApp t1 t2) = rApp (perm_rlam pi t1) (perm_rlam pi t2)"
-| "perm_rlam pi (rLam a t) = rLam (pi \<bullet> a) (perm_rlam pi t)"
-
-end
-
-declare perm_rlam.simps[eqvt]
-
-instance rlam::pt_name
-  apply(default)
-  apply(induct_tac [!] x rule: rlam.induct)
-  apply(simp_all add: pt_name2 pt_name3)
-  done
-
-instance rlam::fs_name
-  apply(default)
-  apply(induct_tac [!] x rule: rlam.induct)
-  apply(simp add: supp_def)
-  apply(fold supp_def)
-  apply(simp add: supp_atm)
-  apply(simp add: supp_def Collect_imp_eq Collect_neg_eq)
-  apply(simp add: supp_def)
-  apply(simp add: supp_def Collect_imp_eq Collect_neg_eq[symmetric])
-  apply(fold supp_def)
-  apply(simp add: supp_atm)
-  done
-
-declare set_diff_eqvt[eqvt]
-
-lemma rfv_eqvt[eqvt]:
-  fixes pi::"name prm"
-  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
-apply(induct t)
-apply(simp_all)
-apply(simp add: perm_set_eq)
-apply(simp add: union_eqvt)
-apply(simp add: set_diff_eqvt)
-apply(simp add: perm_set_eq)
-done
-
-inductive
-    alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)"
-| a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2"
-| a3: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s \<and> (pi \<bullet> a) = b)
-       \<Longrightarrow> rLam a t \<approx> rLam b s"
-
-
-(* should be automatic with new version of eqvt-machinery *)
-lemma alpha_eqvt:
-  fixes pi::"name prm"
-  shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)"
-apply(induct rule: alpha.induct)
-apply(simp add: a1)
-apply(simp add: a2)
-apply(simp)
-apply(rule a3)
-apply(erule conjE)
-apply(erule exE)
-apply(erule conjE)
-apply(rule_tac x="pi \<bullet> pia" in exI)
-apply(rule conjI)
-apply(rule_tac pi1="rev pi" in perm_bij[THEN iffD1])
-apply(perm_simp add: eqvts)
-apply(rule conjI)
-apply(rule_tac pi1="rev pi" in pt_fresh_star_bij(1)[OF pt_name_inst at_name_inst, THEN iffD1])
-apply(perm_simp add: eqvts)
-apply(rule conjI)
-apply(subst perm_compose[symmetric])
-apply(simp)
-apply(subst perm_compose[symmetric])
-apply(simp)
-done
-
-lemma alpha_refl:
-  shows "t \<approx> t"
-apply(induct t rule: rlam.induct)
-apply(simp add: a1)
-apply(simp add: a2)
-apply(rule a3)
-apply(rule_tac x="[]" in exI)
-apply(simp_all add: fresh_star_def fresh_list_nil)
-done
-
-lemma alpha_sym:
-  shows "t \<approx> s \<Longrightarrow> s \<approx> t"
-apply(induct rule: alpha.induct)
-apply(simp add: a1)
-apply(simp add: a2)
-apply(rule a3)
-apply(erule exE)
-apply(rule_tac x="rev pi" in exI)
-apply(simp)
-apply(simp add: fresh_star_def fresh_list_rev)
-apply(rule conjI)
-apply(erule conjE)+
-apply(rotate_tac 3)
-apply(drule_tac pi="rev pi" in alpha_eqvt)
-apply(perm_simp)
-apply(rule pt_bij2[OF pt_name_inst at_name_inst])
-apply(simp)
-done
-
-lemma alpha_trans:
-  shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
-apply(induct arbitrary: t3 rule: alpha.induct)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(simp add: a1)
-apply(rotate_tac 4)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(simp add: a2)
-apply(rotate_tac 1)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(erule conjE)+
-apply(erule exE)+
-apply(erule conjE)+
-apply(rule a3)
-apply(rule_tac x="pia @ pi" in exI)
-apply(simp add: fresh_star_def fresh_list_append)
-apply(simp add: pt_name2)
-apply(drule_tac x="rev pia \<bullet> sa" in spec)
-apply(drule mp)
-apply(rotate_tac 8)
-apply(drule_tac pi="rev pia" in alpha_eqvt)
-apply(perm_simp)
-apply(rotate_tac 11)
-apply(drule_tac pi="pia" in alpha_eqvt)
-apply(perm_simp)
-done
-
-lemma alpha_equivp:
-  shows "equivp alpha"
-apply(rule equivpI)
-unfolding reflp_def symp_def transp_def
-apply(auto intro: alpha_refl alpha_sym alpha_trans)
-done
-
-lemma alpha_rfv:
-  shows "t \<approx> s \<Longrightarrow> rfv t = rfv s"
-apply(induct rule: alpha.induct)
-apply(simp)
-apply(simp)
-apply(simp)
-done
-
-quotient_type lam = rlam / alpha
-  by (rule alpha_equivp)
-
-
-quotient_definition
-  "Var :: name \<Rightarrow> lam"
-is
-  "rVar"
-
-quotient_definition
-   "App :: lam \<Rightarrow> lam \<Rightarrow> lam"
-is
-  "rApp"
-
-quotient_definition
-  "Lam :: name \<Rightarrow> lam \<Rightarrow> lam"
-is
-  "rLam"
-
-quotient_definition
-  "fv :: lam \<Rightarrow> name set"
-is
-  "rfv"
-
-(* definition of overloaded permutation function *)
-(* for the lifted type lam                       *)
-overloading
-  perm_lam \<equiv> "perm :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"   (unchecked)
-begin
-
-quotient_definition
-  "perm_lam :: 'x prm \<Rightarrow> lam \<Rightarrow> lam"
-is
-  "perm::'x prm \<Rightarrow> rlam \<Rightarrow> rlam"
-
-end
-
-lemma perm_rsp[quot_respect]:
-  "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>"
-  apply(auto)
-  (* this is propably true if some type conditions are imposed ;o) *)
-  sorry
-
-lemma fresh_rsp:
-  "(op = ===> alpha ===> op =) fresh fresh"
-  apply(auto)
-  (* this is probably only true if some type conditions are imposed *)
-  sorry
-
-lemma rVar_rsp[quot_respect]:
-  "(op = ===> alpha) rVar rVar"
-  by (auto intro: a1)
-
-lemma rApp_rsp[quot_respect]: "(alpha ===> alpha ===> alpha) rApp rApp"
-  by (auto intro: a2)
-
-lemma rLam_rsp[quot_respect]: "(op = ===> alpha ===> alpha) rLam rLam"
-  apply(auto)
-  apply(rule a3)
-  apply(rule_tac x="[]" in exI)
-  unfolding fresh_star_def
-  apply(simp add: fresh_list_nil)
-  apply(simp add: alpha_rfv)
-  done
-
-lemma rfv_rsp[quot_respect]: 
-  "(alpha ===> op =) rfv rfv"
-apply(simp add: alpha_rfv)
-done
-
-section {* lifted theorems *}
-
-lemma lam_induct:
-  "\<lbrakk>\<And>name. P (Var name);
-    \<And>lam1 lam2. \<lbrakk>P lam1; P lam2\<rbrakk> \<Longrightarrow> P (App lam1 lam2);
-    \<And>name lam. P lam \<Longrightarrow> P (Lam name lam)\<rbrakk> 
-    \<Longrightarrow> P lam"
-  by (lifting rlam.induct)
-
-ML {* show_all_types := true *}
-
-lemma perm_lam [simp]:
-  fixes pi::"'a prm"
-  shows "pi \<bullet> Var a = Var (pi \<bullet> a)"
-  and   "pi \<bullet> App t1 t2 = App (pi \<bullet> t1) (pi \<bullet> t2)"
-  and   "pi \<bullet> Lam a t = Lam (pi \<bullet> a) (pi \<bullet> t)"
-apply(lifting perm_rlam.simps)
-ML_prf {*
-  List.last (map (symmetric o #def) (Quotient_Info.qconsts_dest @{context}));
-  List.last (map (Thm.varifyT o symmetric o #def) (Quotient_Info.qconsts_dest @{context}))
-*}
-done
-
-instance lam::pt_name
-apply(default)
-apply(induct_tac [!] x rule: lam_induct)
-apply(simp_all add: pt_name2 pt_name3)
-done
-
-lemma fv_lam [simp]: 
-  shows "fv (Var a) = {a}"
-  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
-  and   "fv (Lam a t) = fv t - {a}"
-apply(lifting rfv_var rfv_app rfv_lam)
-done
-
-
-lemma a1: 
-  "a = b \<Longrightarrow> Var a = Var b"
-  by  (lifting a1)
-
-lemma a2: 
-  "\<lbrakk>x = xa; xb = xc\<rbrakk> \<Longrightarrow> App x xb = App xa xc"
-  by  (lifting a2)
-
-lemma a3: 
-  "\<lbrakk>\<exists>pi::name prm. (fv t - {a} = fv s - {b} \<and> (fv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) = s \<and> (pi \<bullet> a) = b)\<rbrakk> 
-   \<Longrightarrow> Lam a t = Lam b s"
-  by  (lifting a3)
-
-lemma alpha_cases: 
-  "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
-    \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P;
-    \<And>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s; 
-         \<exists>pi::name prm. fv t - {a} = fv s - {b} \<and> (fv t - {a}) \<sharp>* pi \<and> (pi \<bullet> t) = s \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> P\<rbrakk>
-    \<Longrightarrow> P"
-  by (lifting alpha.cases)
-
-lemma alpha_induct: 
-  "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b);
-    \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc);
-     \<And>t a s b.
-        \<lbrakk>\<exists>pi::name prm. fv t - {a} = fv s - {b} \<and>
-         (fv t - {a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s) \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
-    \<Longrightarrow> qxb qx qxa"
-  by (lifting alpha.induct)
-
-lemma lam_inject [simp]: 
-  shows "(Var a = Var b) = (a = b)"
-  and   "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)"
-apply(lifting rlam.inject(1) rlam.inject(2))
-apply(auto)
-apply(drule alpha.cases)
-apply(simp_all)
-apply(simp add: alpha.a1)
-apply(drule alpha.cases)
-apply(simp_all)
-apply(drule alpha.cases)
-apply(simp_all)
-apply(rule alpha.a2)
-apply(simp_all)
-done
-
-lemma rlam_distinct:
-  shows "\<not>(rVar nam \<approx> rApp rlam1' rlam2')"
-  and   "\<not>(rApp rlam1' rlam2' \<approx> rVar nam)"
-  and   "\<not>(rVar nam \<approx> rLam nam' rlam')"
-  and   "\<not>(rLam nam' rlam' \<approx> rVar nam)"
-  and   "\<not>(rApp rlam1 rlam2 \<approx> rLam nam' rlam')"
-  and   "\<not>(rLam nam' rlam' \<approx> rApp rlam1 rlam2)"
-apply auto
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-apply(erule alpha.cases)
-apply simp_all
-done
-
-lemma lam_distinct[simp]:
-  shows "Var nam \<noteq> App lam1' lam2'"
-  and   "App lam1' lam2' \<noteq> Var nam"
-  and   "Var nam \<noteq> Lam nam' lam'"
-  and   "Lam nam' lam' \<noteq> Var nam"
-  and   "App lam1 lam2 \<noteq> Lam nam' lam'"
-  and   "Lam nam' lam' \<noteq> App lam1 lam2"
-apply(lifting rlam_distinct(1) rlam_distinct(2) rlam_distinct(3) rlam_distinct(4) rlam_distinct(5) rlam_distinct(6))
-done
-
-lemma var_supp1:
-  shows "(supp (Var a)) = ((supp a)::name set)"
-  by (simp add: supp_def)
-
-lemma var_supp:
-  shows "(supp (Var a)) = {a::name}"
-  using var_supp1 by (simp add: supp_atm)
-
-lemma app_supp:
-  shows "supp (App t1 t2) = (supp t1) \<union> ((supp t2)::name set)"
-apply(simp only: perm_lam supp_def lam_inject)
-apply(simp add: Collect_imp_eq Collect_neg_eq)
-done
-
-lemma lam_supp:
-  shows "supp (Lam x t) = ((supp ([x].t))::name set)"
-apply(simp add: supp_def)
-apply(simp add: abs_perm)
-sorry
-
-
-instance lam::fs_name
-apply(default)
-apply(induct_tac x rule: lam_induct)
-apply(simp add: var_supp)
-apply(simp add: app_supp)
-apply(simp add: lam_supp abs_supp)
-done
-
-lemma fresh_lam:
-  "(a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> a \<sharp> t)"
-apply(simp add: fresh_def)
-apply(simp add: lam_supp abs_supp)
-apply(auto)
-done
-
-lemma lam_induct_strong:
-  fixes a::"'a::fs_name"
-  assumes a1: "\<And>name b. P b (Var name)"
-  and     a2: "\<And>lam1 lam2 b. \<lbrakk>\<And>c. P c lam1; \<And>c. P c lam2\<rbrakk> \<Longrightarrow> P b (App lam1 lam2)"
-  and     a3: "\<And>name lam b. \<lbrakk>\<And>c. P c lam; name \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
-  shows "P a lam"
-proof -
-  have "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" 
-  proof (induct lam rule: lam_induct)
-    case (1 name pi)
-    show "P a (pi \<bullet> Var name)"
-      apply (simp)
-      apply (rule a1)
-      done
-  next
-    case (2 lam1 lam2 pi)
-    have b1: "\<And>(pi::name prm) a. P a (pi \<bullet> lam1)" by fact
-    have b2: "\<And>(pi::name prm) a. P a (pi \<bullet> lam2)" by fact
-    show "P a (pi \<bullet> App lam1 lam2)"
-      apply (simp)
-      apply (rule a2)
-      apply (rule b1)
-      apply (rule b2)
-      done
-  next
-    case (3 name lam pi a)
-    have b: "\<And>(pi::name prm) a. P a (pi \<bullet> lam)" by fact
-    obtain c::name where fr: "c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
-      apply(rule exists_fresh[of "(a, pi\<bullet>name, pi\<bullet>lam)"])
-      apply(simp_all add: fs_name1)
-      done
-    from b fr have p: "P a (Lam c (([(c, pi\<bullet>name)]@pi)\<bullet>lam))" 
-      apply -
-      apply(rule a3)
-      apply(blast)
-      apply(simp)
-      done
-    have eq: "[(c, pi\<bullet>name)] \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
-      apply(rule perm_fresh_fresh)
-      using fr
-      apply(simp add: fresh_lam)
-      apply(simp add: fresh_lam)
-      done
-    show "P a (pi \<bullet> Lam name lam)" 
-      apply (simp)
-      apply(subst eq[symmetric])
-      using p
-      apply(simp only: perm_lam pt_name2 swap_simps)
-      done
-  qed
-  then have "P a (([]::name prm) \<bullet> lam)" by blast
-  then show "P a lam" by simp 
-qed
-
-
-lemma var_fresh:
-  fixes a::"name"
-  shows "(a \<sharp> (Var b)) = (a \<sharp> b)"
-  apply(simp add: fresh_def)
-  apply(simp add: var_supp1)
-  done
-
-(* lemma hom_reg: *)
-
-lemma rlam_rec_eqvt:
-  fixes pi::"name prm"
-  and   f1::"name \<Rightarrow> ('a::pt_name)"
-  shows "(pi\<bullet>rlam_rec f1 f2 f3 t) = rlam_rec (pi\<bullet>f1) (pi\<bullet>f2) (pi\<bullet>f3) (pi\<bullet>t)"
-apply(induct t)
-apply(simp_all)
-apply(simp add: perm_fun_def)
-apply(perm_simp)
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-back
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(simp)
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-back
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(subst pt_fun_app_eq[OF pt_name_inst at_name_inst])
-apply(simp)
-done
- 
-
-lemma rlam_rec_respects:
-  assumes f1: "f_var \<in> Respects (op= ===> op=)"
-  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
-  and     f3: "f_lam \<in> Respects (op= ===> alpha ===> op= ===> op=)"
-  shows "rlam_rec f_var f_app f_lam \<in> Respects (alpha ===> op =)"
-apply(simp add: mem_def)
-apply(simp add: Respects_def)
-apply(rule allI)
-apply(rule allI)
-apply(rule impI)
-apply(erule alpha.induct)
-apply(simp)
-apply(simp)
-using f2
-apply(simp add: mem_def)
-apply(simp add: Respects_def)
-using f3[simplified mem_def Respects_def]
-apply(simp)
-apply(case_tac "a=b")
-apply(clarify)
-apply(simp)
-(* probably true *)
-sorry
-
-function
-  term1_hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
-                (rlam \<Rightarrow> rlam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
-                ((name \<Rightarrow> rlam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> rlam \<Rightarrow> 'a"
-where
-  "term1_hom var app abs' (rVar x) = (var x)"
-| "term1_hom var app abs' (rApp t u) =
-     app t u (term1_hom var app abs' t) (term1_hom var app abs' u)"
-| "term1_hom var app abs' (rLam x u) =
-     abs' (\<lambda>y. [(x, y)] \<bullet> u) (\<lambda>y. term1_hom var app abs' ([(x, y)] \<bullet> u))"
-apply(pat_completeness)
-apply(auto)
-done
-
-lemma pi_size:
-  fixes pi::"name prm"
-  and   t::"rlam"
-  shows "size (pi \<bullet> t) = size t"
-apply(induct t)
-apply(auto)
-done
-
-termination term1_hom
-  apply(relation "measure (\<lambda>(f1, f2, f3, t). size t)")
-apply(auto simp add: pi_size)
-done
-
-lemma lam_exhaust:
-  "\<lbrakk>\<And>name. y = Var name \<Longrightarrow> P; \<And>rlam1 rlam2. y = App rlam1 rlam2 \<Longrightarrow> P; \<And>name rlam. y = Lam name rlam \<Longrightarrow> P\<rbrakk>
-    \<Longrightarrow> P"
-apply(lifting rlam.exhaust)
-done
-
-(* THIS IS NOT TRUE, but it lets prove the existence of the hom function *)
-lemma lam_inject':
-  "(Lam a x = Lam b y) = ((\<lambda>c. [(a, c)] \<bullet> x) = (\<lambda>c. [(b, c)] \<bullet> y))"
-sorry
-
-function
-  hom :: "(name \<Rightarrow> 'a) \<Rightarrow>
-                (lam \<Rightarrow> lam \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow>
-                ((name \<Rightarrow> lam) \<Rightarrow> (name \<Rightarrow> 'a) \<Rightarrow> 'a) \<Rightarrow> lam \<Rightarrow> 'a"
-where
-  "hom f_var f_app f_lam (Var x) = f_var x"
-| "hom f_var f_app f_lam (App l r) = f_app l r (hom f_var f_app f_lam l) (hom f_var f_app f_lam r)"
-| "hom f_var f_app f_lam (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom f_var f_app f_lam ([(a,b)] \<bullet> x))"
-defer
-apply(simp_all add: lam_inject') (* inject, distinct *)
-apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
-apply(rule refl)
-apply(rule ext)
-apply(tactic {* Cong_Tac.cong_tac @{thm cong} 1 *})
-apply simp_all
-apply(erule conjE)+
-apply(rule_tac x="b" in cong)
-apply simp_all
-apply auto
-apply(rule_tac y="b" in lam_exhaust)
-apply simp_all
-apply auto
-apply meson
-apply(simp_all add: lam_inject')
-apply metis
-done
-
-termination hom
-  apply -
-(*
-ML_prf {* Size.size_thms @{theory} "LamEx.lam" *}
-*)
-sorry
-
-thm hom.simps
-
-lemma term1_hom_rsp:
-  "\<lbrakk>(alpha ===> alpha ===> op =) f_app f_app; ((op = ===> alpha) ===> op =) f_lam f_lam\<rbrakk>
-       \<Longrightarrow> (alpha ===> op =) (term1_hom f_var f_app f_lam) (term1_hom f_var f_app f_lam)"
-apply(simp)
-apply(rule allI)+
-apply(rule impI)
-apply(erule alpha.induct)
-apply(auto)[1]
-apply(auto)[1]
-apply(simp)
-apply(erule conjE)+
-apply(erule exE)+
-apply(erule conjE)+
-apply(clarify)
-sorry
-
-lemma hom: "
-\<forall>f_var. \<forall>f_app \<in> Respects(alpha ===> alpha ===> op =).
-\<forall>f_lam \<in> Respects((op = ===> alpha) ===> op =).
-\<exists>hom\<in>Respects (alpha ===> op =). 
-    ((\<forall>x. hom (rVar x) = f_var x) \<and>
-     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
-     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
-apply(rule allI)
-apply(rule ballI)+
-apply(rule_tac x="term1_hom f_var f_app f_lam" in bexI)
-apply(simp_all)
-apply(simp only: in_respects)
-apply(rule term1_hom_rsp)
-apply(assumption)+
-done
-
-lemma hom':
-"\<exists>hom.
-  ((\<forall>x. hom (Var x) = f_var x) \<and>
-   (\<forall>l r. hom (App l r) = f_app l r (hom l) (hom r)) \<and>
-   (\<forall>x a. hom (Lam a x) = f_lam (\<lambda>b. ([(a,b)] \<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
-apply (lifting hom)
-done
-
-(* test test
-lemma raw_hom_correct: 
-  assumes f1: "f_var \<in> Respects (op= ===> op=)"
-  and     f2: "f_app \<in> Respects (alpha ===> alpha ===> op= ===> op= ===> op=)"
-  and     f3: "f_lam \<in> Respects ((op= ===> alpha) ===> (op= ===> op=) ===> op=)"
-  shows "\<exists>!hom\<in>Respects (alpha ===> op =). 
-    ((\<forall>x. hom (rVar x) = f_var x) \<and>
-     (\<forall>l r. hom (rApp l r) = f_app l r (hom l) (hom r)) \<and>
-     (\<forall>x a. hom (rLam a x) = f_lam (\<lambda>b. ([(a,b)]\<bullet> x)) (\<lambda>b. hom ([(a,b)] \<bullet> x))))"
-unfolding Bex1_def
-apply(rule ex1I)
-sorry
-*)
-
-
-end
-
--- a/Quot/Examples/LarryDatatype.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,394 +0,0 @@
-theory LarryDatatype
-imports Main "../Quotient" "../Quotient_Syntax"
-begin
-
-subsection{*Defining the Free Algebra*}
-
-datatype
-  freemsg = NONCE  nat
-        | MPAIR  freemsg freemsg
-        | CRYPT  nat freemsg  
-        | DECRYPT  nat freemsg
-
-inductive 
-  msgrel::"freemsg \<Rightarrow> freemsg \<Rightarrow> bool" (infixl "\<sim>" 50)
-where 
-  CD:    "CRYPT K (DECRYPT K X) \<sim> X"
-| DC:    "DECRYPT K (CRYPT K X) \<sim> X"
-| NONCE: "NONCE N \<sim> NONCE N"
-| MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
-| CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
-| DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
-| SYM:   "X \<sim> Y \<Longrightarrow> Y \<sim> X"
-| TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
-
-lemmas msgrel.intros[intro]
-
-text{*Proving that it is an equivalence relation*}
-
-lemma msgrel_refl: "X \<sim> X"
-by (induct X, (blast intro: msgrel.intros)+)
-
-theorem equiv_msgrel: "equivp msgrel"
-proof (rule equivpI)
-  show "reflp msgrel" by (simp add: reflp_def msgrel_refl)
-  show "symp msgrel" by (simp add: symp_def, blast intro: msgrel.SYM)
-  show "transp msgrel" by (simp add: transp_def, blast intro: msgrel.TRANS)
-qed
-
-subsection{*Some Functions on the Free Algebra*}
-
-subsubsection{*The Set of Nonces*}
-
-fun
-  freenonces :: "freemsg \<Rightarrow> nat set"
-where
-  "freenonces (NONCE N) = {N}"
-| "freenonces (MPAIR X Y) = freenonces X \<union> freenonces Y"
-| "freenonces (CRYPT K X) = freenonces X"
-| "freenonces (DECRYPT K X) = freenonces X"
-
-theorem msgrel_imp_eq_freenonces: 
-  assumes a: "U \<sim> V"
-  shows "freenonces U = freenonces V"
-  using a by (induct) (auto) 
-
-subsubsection{*The Left Projection*}
-
-text{*A function to return the left part of the top pair in a message.  It will
-be lifted to the initial algrebra, to serve as an example of that process.*}
-fun
-  freeleft :: "freemsg \<Rightarrow> freemsg"
-where
-  "freeleft (NONCE N) = NONCE N"
-| "freeleft (MPAIR X Y) = X"
-| "freeleft (CRYPT K X) = freeleft X"
-| "freeleft (DECRYPT K X) = freeleft X"
-
-text{*This theorem lets us prove that the left function respects the
-equivalence relation.  It also helps us prove that MPair
-  (the abstract constructor) is injective*}
-lemma msgrel_imp_eqv_freeleft_aux:
-  shows "freeleft U \<sim> freeleft U"
-  by (induct rule: freeleft.induct) (auto)
-
-theorem msgrel_imp_eqv_freeleft:
-  assumes a: "U \<sim> V" 
-  shows "freeleft U \<sim> freeleft V"
-  using a
-  by (induct) (auto intro: msgrel_imp_eqv_freeleft_aux)
-
-subsubsection{*The Right Projection*}
-
-text{*A function to return the right part of the top pair in a message.*}
-fun
-  freeright :: "freemsg \<Rightarrow> freemsg"
-where
-  "freeright (NONCE N) = NONCE N"
-| "freeright (MPAIR X Y) = Y"
-| "freeright (CRYPT K X) = freeright X"
-| "freeright (DECRYPT K X) = freeright X"
-
-text{*This theorem lets us prove that the right function respects the
-equivalence relation.  It also helps us prove that MPair
-  (the abstract constructor) is injective*}
-lemma msgrel_imp_eqv_freeright_aux:
-  shows "freeright U \<sim> freeright U"
-  by (induct rule: freeright.induct) (auto)
-
-theorem msgrel_imp_eqv_freeright:
-  assumes a: "U \<sim> V" 
-  shows "freeright U \<sim> freeright V"
-  using a
-  by (induct) (auto intro: msgrel_imp_eqv_freeright_aux)
-
-subsubsection{*The Discriminator for Constructors*}
-
-text{*A function to distinguish nonces, mpairs and encryptions*}
-fun 
-  freediscrim :: "freemsg \<Rightarrow> int"
-where
-   "freediscrim (NONCE N) = 0"
- | "freediscrim (MPAIR X Y) = 1"
- | "freediscrim (CRYPT K X) = freediscrim X + 2"
- | "freediscrim (DECRYPT K X) = freediscrim X - 2"
-
-text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
-theorem msgrel_imp_eq_freediscrim:
-  assumes a: "U \<sim> V"
-  shows "freediscrim U = freediscrim V"
-  using a by (induct) (auto)
-
-subsection{*The Initial Algebra: A Quotiented Message Type*}
-
-quotient_type msg = freemsg / msgrel
-  by (rule equiv_msgrel)
-
-text{*The abstract message constructors*}
-
-quotient_definition
-  "Nonce :: nat \<Rightarrow> msg"
-is
-  "NONCE"
-
-quotient_definition
-  "MPair :: msg \<Rightarrow> msg \<Rightarrow> msg"
-is
-  "MPAIR"
-
-quotient_definition
-  "Crypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
-is
-  "CRYPT"
-
-quotient_definition
-  "Decrypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
-is
-  "DECRYPT"
-
-lemma [quot_respect]:
-  shows "(op = ===> op \<sim> ===> op \<sim>) CRYPT CRYPT"
-by (auto intro: CRYPT)
-
-lemma [quot_respect]:
-  shows "(op = ===> op \<sim> ===> op \<sim>) DECRYPT DECRYPT"
-by (auto intro: DECRYPT)
-
-text{*Establishing these two equations is the point of the whole exercise*}
-theorem CD_eq [simp]: 
-  shows "Crypt K (Decrypt K X) = X"
-  by (lifting CD)
-
-theorem DC_eq [simp]: 
-  shows "Decrypt K (Crypt K X) = X"
-  by (lifting DC)
-
-subsection{*The Abstract Function to Return the Set of Nonces*}
-
-quotient_definition
-   "nonces:: msg \<Rightarrow> nat set"
-is
-  "freenonces"
-
-text{*Now prove the four equations for @{term nonces}*}
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op =) freenonces freenonces"
-  by (simp add: msgrel_imp_eq_freenonces)
-
-lemma [quot_respect]:
-  shows "(op = ===> op \<sim>) NONCE NONCE"
-  by (simp add: NONCE)
-
-lemma nonces_Nonce [simp]: 
-  shows "nonces (Nonce N) = {N}"
-  by (lifting freenonces.simps(1))
- 
-lemma [quot_respect]:
-  shows " (op \<sim> ===> op \<sim> ===> op \<sim>) MPAIR MPAIR"
-  by (simp add: MPAIR)
-
-lemma nonces_MPair [simp]: 
-  shows "nonces (MPair X Y) = nonces X \<union> nonces Y"
-  by (lifting freenonces.simps(2))
-
-lemma nonces_Crypt [simp]: 
-  shows "nonces (Crypt K X) = nonces X"
-  by (lifting freenonces.simps(3))
-
-lemma nonces_Decrypt [simp]: 
-  shows "nonces (Decrypt K X) = nonces X"
-  by (lifting freenonces.simps(4))
-
-subsection{*The Abstract Function to Return the Left Part*}
-
-quotient_definition
-  "left:: msg \<Rightarrow> msg"
-is
-  "freeleft"
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op \<sim>) freeleft freeleft"
-  by (simp add: msgrel_imp_eqv_freeleft)
-
-lemma left_Nonce [simp]: 
-  shows "left (Nonce N) = Nonce N"
-  by (lifting freeleft.simps(1))
-
-lemma left_MPair [simp]: 
-  shows "left (MPair X Y) = X"
-  by (lifting freeleft.simps(2))
-
-lemma left_Crypt [simp]: 
-  shows "left (Crypt K X) = left X"
-  by (lifting freeleft.simps(3))
-
-lemma left_Decrypt [simp]: 
-  shows "left (Decrypt K X) = left X"
-  by (lifting freeleft.simps(4))
-
-subsection{*The Abstract Function to Return the Right Part*}
-
-quotient_definition
-  "right:: msg \<Rightarrow> msg"
-is
-  "freeright"
-
-text{*Now prove the four equations for @{term right}*}
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op \<sim>) freeright freeright"
-  by (simp add: msgrel_imp_eqv_freeright)
-
-lemma right_Nonce [simp]: 
-  shows "right (Nonce N) = Nonce N"
-  by (lifting freeright.simps(1))
-
-lemma right_MPair [simp]: 
-  shows "right (MPair X Y) = Y"
-  by (lifting freeright.simps(2))
-
-lemma right_Crypt [simp]: 
-  shows "right (Crypt K X) = right X"
-  by (lifting freeright.simps(3))
-
-lemma right_Decrypt [simp]: 
-  shows "right (Decrypt K X) = right X"
-  by (lifting freeright.simps(4))
-
-subsection{*Injectivity Properties of Some Constructors*}
-
-lemma NONCE_imp_eq: 
-  shows "NONCE m \<sim> NONCE n \<Longrightarrow> m = n"
-  by (drule msgrel_imp_eq_freenonces, simp)
-
-text{*Can also be proved using the function @{term nonces}*}
-lemma Nonce_Nonce_eq [iff]: 
-  shows "(Nonce m = Nonce n) = (m = n)"
-proof
-  assume "Nonce m = Nonce n"
-  then show "m = n" by (lifting NONCE_imp_eq)
-next
-  assume "m = n" 
-  then show "Nonce m = Nonce n" by simp
-qed
-
-lemma MPAIR_imp_eqv_left: 
-  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> X \<sim> X'"
-  by (drule msgrel_imp_eqv_freeleft) (simp)
-
-lemma MPair_imp_eq_left: 
-  assumes eq: "MPair X Y = MPair X' Y'" 
-  shows "X = X'"
-  using eq by (lifting MPAIR_imp_eqv_left)
-
-lemma MPAIR_imp_eqv_right: 
-  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> Y \<sim> Y'"
-  by (drule msgrel_imp_eqv_freeright) (simp)
-
-lemma MPair_imp_eq_right: 
-  shows "MPair X Y = MPair X' Y' \<Longrightarrow> Y = Y'"
-  by (lifting  MPAIR_imp_eqv_right)
-
-theorem MPair_MPair_eq [iff]: 
-  shows "(MPair X Y = MPair X' Y') = (X=X' & Y=Y')" 
-  by (blast dest: MPair_imp_eq_left MPair_imp_eq_right)
-
-lemma NONCE_neqv_MPAIR: 
-  shows "\<not>(NONCE m \<sim> MPAIR X Y)"
-  by (auto dest: msgrel_imp_eq_freediscrim)
-
-theorem Nonce_neq_MPair [iff]: 
-  shows "Nonce N \<noteq> MPair X Y"
-  by (lifting NONCE_neqv_MPAIR)
-
-text{*Example suggested by a referee*}
-
-lemma CRYPT_NONCE_neq_NONCE:
-  shows "\<not>(CRYPT K (NONCE M) \<sim> NONCE N)"
-  by (auto dest: msgrel_imp_eq_freediscrim)
-
-theorem Crypt_Nonce_neq_Nonce: 
-  shows "Crypt K (Nonce M) \<noteq> Nonce N"
-  by (lifting CRYPT_NONCE_neq_NONCE)
-
-text{*...and many similar results*}
-lemma CRYPT2_NONCE_neq_NONCE: 
-  shows "\<not>(CRYPT K (CRYPT K' (NONCE M)) \<sim> NONCE N)"
-  by (auto dest: msgrel_imp_eq_freediscrim)  
-
-theorem Crypt2_Nonce_neq_Nonce: 
-  shows "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"
-  by (lifting CRYPT2_NONCE_neq_NONCE) 
-
-theorem Crypt_Crypt_eq [iff]: 
-  shows "(Crypt K X = Crypt K X') = (X=X')" 
-proof
-  assume "Crypt K X = Crypt K X'"
-  hence "Decrypt K (Crypt K X) = Decrypt K (Crypt K X')" by simp
-  thus "X = X'" by simp
-next
-  assume "X = X'"
-  thus "Crypt K X = Crypt K X'" by simp
-qed
-
-theorem Decrypt_Decrypt_eq [iff]: 
-  shows "(Decrypt K X = Decrypt K X') = (X=X')" 
-proof
-  assume "Decrypt K X = Decrypt K X'"
-  hence "Crypt K (Decrypt K X) = Crypt K (Decrypt K X')" by simp
-  thus "X = X'" by simp
-next
-  assume "X = X'"
-  thus "Decrypt K X = Decrypt K X'" by simp
-qed
-
-lemma msg_induct_aux:
-  shows "\<lbrakk>\<And>N. P (Nonce N);
-          \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y);
-          \<And>K X. P X \<Longrightarrow> P (Crypt K X);
-          \<And>K X. P X \<Longrightarrow> P (Decrypt K X)\<rbrakk> \<Longrightarrow> P msg"
-  by (lifting freemsg.induct)
-
-lemma msg_induct [case_names Nonce MPair Crypt Decrypt, cases type: msg]:
-  assumes N: "\<And>N. P (Nonce N)"
-      and M: "\<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y)"
-      and C: "\<And>K X. P X \<Longrightarrow> P (Crypt K X)"
-      and D: "\<And>K X. P X \<Longrightarrow> P (Decrypt K X)"
-  shows "P msg"
-  using N M C D by (rule msg_induct_aux)
-
-subsection{*The Abstract Discriminator*}
-
-text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
-need this function in order to prove discrimination theorems.*}
-
-quotient_definition
-  "discrim:: msg \<Rightarrow> int"
-is
-  "freediscrim"
-
-text{*Now prove the four equations for @{term discrim}*}
-
-lemma [quot_respect]:
-  shows "(op \<sim> ===> op =) freediscrim freediscrim"
-  by (auto simp add: msgrel_imp_eq_freediscrim)
-
-lemma discrim_Nonce [simp]: 
-  shows "discrim (Nonce N) = 0"
-  by (lifting freediscrim.simps(1))
-
-lemma discrim_MPair [simp]: 
-  shows "discrim (MPair X Y) = 1"
-  by (lifting freediscrim.simps(2))
-
-lemma discrim_Crypt [simp]: 
-  shows "discrim (Crypt K X) = discrim X + 2"
-  by (lifting freediscrim.simps(3))
-
-lemma discrim_Decrypt [simp]: 
-  shows "discrim (Decrypt K X) = discrim X - 2"
-  by (lifting freediscrim.simps(4))
-
-end
-
--- a/Quot/Examples/LarryInt.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,395 +0,0 @@
-
-header{*The Integers as Equivalence Classes over Pairs of Natural Numbers*}
-
-theory LarryInt
-imports Nat "../Quotient" "../Quotient_Product"
-begin
-
-fun
-  intrel :: "(nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> bool" 
-where
-  "intrel (x, y) (u, v) = (x + v = u + y)"
-
-quotient_type int = "nat \<times> nat" / intrel
-  by (auto simp add: equivp_def expand_fun_eq)
-
-instantiation int :: "{zero, one, plus, uminus, minus, times, ord}"
-begin
-
-quotient_definition
-  Zero_int_def: "0::int" is "(0::nat, 0::nat)"
-
-quotient_definition
-  One_int_def: "1::int" is "(1::nat, 0::nat)"
-
-definition
-  "add_raw \<equiv> \<lambda>(x, y) (u, v). (x + (u::nat), y + (v::nat))"
-
-quotient_definition
-  "(op +) :: int \<Rightarrow> int \<Rightarrow> int" 
-is
-  "add_raw"
-
-definition
-  "uminus_raw \<equiv> \<lambda>(x::nat, y::nat). (y, x)"
-
-quotient_definition
-  "uminus :: int \<Rightarrow> int" 
-is
-  "uminus_raw"
-
-fun
-  mult_raw::"nat \<times> nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
-where
-  "mult_raw (x, y) (u, v) = (x*u + y*v, x*v + y*u)"
-
-quotient_definition
-  "(op *) :: int \<Rightarrow> int \<Rightarrow> int" 
-is
-  "mult_raw"
-
-definition
-  "le_raw \<equiv> \<lambda>(x, y) (u, v). (x+v \<le> u+(y::nat))"
-
-quotient_definition 
-  le_int_def: "(op \<le>) :: int \<Rightarrow> int \<Rightarrow> bool" 
-is
-  "le_raw"
-
-definition
-  less_int_def: "z < (w::int) \<equiv> (z \<le> w & z \<noteq> w)"
-
-definition
-  diff_int_def:  "z - (w::int) \<equiv> z + (-w)"
-
-instance ..
-
-end
-
-subsection{*Construction of the Integers*}
-
-lemma zminus_zminus_raw:
-  "uminus_raw (uminus_raw z) = z"
-  by (cases z) (simp add: uminus_raw_def)
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel) uminus_raw uminus_raw"
-  by (simp add: uminus_raw_def)
-  
-lemma zminus_zminus:
-  fixes z::"int"
-  shows "- (- z) = z"
-  by(lifting zminus_zminus_raw)
-
-lemma zminus_0_raw:
-  shows "uminus_raw (0, 0) = (0, 0::nat)"
-  by (simp add: uminus_raw_def)
-
-lemma zminus_0: 
-  shows "- 0 = (0::int)"
-  by (lifting zminus_0_raw)
-
-subsection{*Integer Addition*}
-
-lemma zminus_zadd_distrib_raw:
-  shows "uminus_raw (add_raw z w) = add_raw (uminus_raw z) (uminus_raw w)"
-by (cases z, cases w)
-   (auto simp add: add_raw_def uminus_raw_def)
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel ===> intrel) add_raw add_raw"
-by (simp add: add_raw_def)
-
-lemma zminus_zadd_distrib: 
-  fixes z w::"int"
-  shows "- (z + w) = (- z) + (- w)"
-  by(lifting zminus_zadd_distrib_raw)
-
-lemma zadd_commute_raw:
-  shows "add_raw z w = add_raw w z"
-by (cases z, cases w)
-   (simp add: add_raw_def)
-
-lemma zadd_commute:
-  fixes z w::"int"
-  shows "(z::int) + w = w + z"
-  by (lifting zadd_commute_raw)
-
-lemma zadd_assoc_raw:
-  shows "add_raw (add_raw z1 z2) z3 = add_raw z1 (add_raw z2 z3)"
-  by (cases z1, cases z2, cases z3) (simp add: add_raw_def)
-
-lemma zadd_assoc: 
-  fixes z1 z2 z3::"int"
-  shows "(z1 + z2) + z3 = z1 + (z2 + z3)"
-  by (lifting zadd_assoc_raw)
-
-lemma zadd_0_raw:
-  shows "add_raw (0, 0) z = z"
-  by (simp add: add_raw_def)
-
-
-text {*also for the instance declaration int :: plus_ac0*}
-lemma zadd_0: 
-  fixes z::"int"
-  shows "0 + z = z"
-  by (lifting zadd_0_raw)
-
-lemma zadd_zminus_inverse_raw:
-  shows "intrel (add_raw (uminus_raw z) z) (0, 0)"
-  by (cases z) (simp add: add_raw_def uminus_raw_def)
-
-lemma zadd_zminus_inverse2: 
-  fixes z::"int"
-  shows "(- z) + z = 0"
-  by (lifting zadd_zminus_inverse_raw)
-
-subsection{*Integer Multiplication*}
-
-lemma zmult_zminus_raw:
-  shows "mult_raw (uminus_raw z) w = uminus_raw (mult_raw z w)"
-apply(cases z, cases w)
-apply(simp add: uminus_raw_def)
-done
-
-lemma mult_raw_fst:
-  assumes a: "intrel x z"
-  shows "intrel (mult_raw x y) (mult_raw z y)"
-using a
-apply(cases x, cases y, cases z)
-apply(auto simp add: mult_raw.simps intrel.simps)
-apply(rename_tac u v w x y z)
-apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
-apply(simp add: mult_ac)
-apply(simp add: add_mult_distrib [symmetric])
-done
-
-lemma mult_raw_snd:
-  assumes a: "intrel x z"
-  shows "intrel (mult_raw y x) (mult_raw y z)"
-using a
-apply(cases x, cases y, cases z)
-apply(auto simp add: mult_raw.simps intrel.simps)
-apply(rename_tac u v w x y z)
-apply(subgoal_tac "u*w + z*w = y*w + v*w  &  u*x + z*x = y*x + v*x")
-apply(simp add: mult_ac)
-apply(simp add: add_mult_distrib [symmetric])
-done
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel ===> intrel) mult_raw mult_raw"
-apply(simp only: fun_rel_def)
-apply(rule allI | rule impI)+
-apply(rule equivp_transp[OF int_equivp])
-apply(rule mult_raw_fst)
-apply(assumption)
-apply(rule mult_raw_snd)
-apply(assumption)
-done
-
-lemma zmult_zminus: 
-  fixes z w::"int"
-  shows "(- z) * w = - (z * w)"
-  by (lifting zmult_zminus_raw)
-
-lemma zmult_commute_raw: 
-  shows "mult_raw z w = mult_raw w z"
-apply(cases z, cases w)
-apply(simp add: add_ac mult_ac)
-done
-
-lemma zmult_commute: 
-  fixes z w::"int"
-  shows "z * w = w * z"
-  by (lifting zmult_commute_raw)
-
-lemma zmult_assoc_raw:
-  shows "mult_raw (mult_raw z1 z2) z3 = mult_raw z1 (mult_raw z2 z3)"
-apply(cases z1, cases z2, cases z3)
-apply(simp add: add_mult_distrib2 mult_ac)
-done
-
-lemma zmult_assoc: 
-  fixes z1 z2 z3::"int"
-  shows "(z1 * z2) * z3 = z1 * (z2 * z3)"
-  by (lifting zmult_assoc_raw)
-
-lemma zadd_mult_distrib_raw:
-  shows "mult_raw (add_raw z1 z2) w = add_raw (mult_raw z1 w) (mult_raw z2 w)"
-apply(cases z1, cases z2, cases w)
-apply(simp add: add_mult_distrib2 mult_ac add_raw_def)
-done
-
-lemma zadd_zmult_distrib: 
-  fixes z1 z2 w::"int"
-  shows "(z1 + z2) * w = (z1 * w) + (z2 * w)"
-  by(lifting zadd_mult_distrib_raw)
-
-lemma zadd_zmult_distrib2: 
-  fixes w z1 z2::"int"
-  shows "w * (z1 + z2) = (w * z1) + (w * z2)"
-  by (simp add: zmult_commute [of w] zadd_zmult_distrib)
-
-lemma zdiff_zmult_distrib: 
-  fixes w z1 z2::"int"
-  shows "(z1 - z2) * w = (z1 * w) - (z2 * w)"
-  by (simp add: diff_int_def zadd_zmult_distrib zmult_zminus)
-
-lemma zdiff_zmult_distrib2: 
-  fixes w z1 z2::"int"
-  shows "w * (z1 - z2) = (w * z1) - (w * z2)"
-  by (simp add: zmult_commute [of w] zdiff_zmult_distrib)
-
-lemmas int_distrib =
-  zadd_zmult_distrib zadd_zmult_distrib2
-  zdiff_zmult_distrib zdiff_zmult_distrib2
-
-lemma zmult_1_raw:
-  shows "mult_raw (1, 0) z = z"
-  by (cases z) (auto)
-
-lemma zmult_1:
-  fixes z::"int"
-  shows "1 * z = z"
-  by (lifting zmult_1_raw)
-
-lemma zmult_1_right: 
-  fixes z::"int"
-  shows "z * 1 = z"
-  by (rule trans [OF zmult_commute zmult_1])
-
-lemma zero_not_one:
-  shows "\<not>(intrel (0, 0) (1::nat, 0::nat))"
-  by auto
-
-text{*The Integers Form A Ring*}
-instance int :: comm_ring_1
-proof
-  fix i j k :: int
-  show "(i + j) + k = i + (j + k)" by (simp add: zadd_assoc)
-  show "i + j = j + i" by (simp add: zadd_commute)
-  show "0 + i = i" by (rule zadd_0)
-  show "- i + i = 0" by (rule zadd_zminus_inverse2)
-  show "i - j = i + (-j)" by (simp add: diff_int_def)
-  show "(i * j) * k = i * (j * k)" by (rule zmult_assoc)
-  show "i * j = j * i" by (rule zmult_commute)
-  show "1 * i = i" by (rule zmult_1) 
-  show "(i + j) * k = i * k + j * k" by (simp add: int_distrib)
-  show "0 \<noteq> (1::int)" by (lifting zero_not_one)
-qed
-
-
-subsection{*The @{text "\<le>"} Ordering*}
-
-lemma zle_refl_raw:
-  shows "le_raw w w"
-  by (cases w) (simp add: le_raw_def)
-
-lemma [quot_respect]:
-  shows "(intrel ===> intrel ===> op =) le_raw le_raw"
-  by (auto) (simp_all add: le_raw_def)
-
-lemma zle_refl: 
-  fixes w::"int"
-  shows "w \<le> w"
-  by (lifting zle_refl_raw)
-
-
-lemma zle_trans_raw:
-  shows "\<lbrakk>le_raw i j; le_raw j k\<rbrakk> \<Longrightarrow> le_raw i k"
-apply(cases i, cases j, cases k)
-apply(auto simp add: le_raw_def)
-done
-
-lemma zle_trans: 
-  fixes i j k::"int"
-  shows "\<lbrakk>i \<le> j; j \<le> k\<rbrakk> \<Longrightarrow> i \<le> k"
-  by (lifting zle_trans_raw)
-
-lemma zle_anti_sym_raw:
-  shows "\<lbrakk>le_raw z w; le_raw w z\<rbrakk> \<Longrightarrow> intrel z w"
-apply(cases z, cases w)
-apply(auto iff: le_raw_def)
-done
-
-lemma zle_anti_sym: 
-  fixes z w::"int"
-  shows "\<lbrakk>z \<le> w; w \<le> z\<rbrakk> \<Longrightarrow> z = w"
-  by (lifting zle_anti_sym_raw)
-
-
-(* Axiom 'order_less_le' of class 'order': *)
-lemma zless_le: 
-  fixes w z::"int"
-  shows "(w < z) = (w \<le> z & w \<noteq> z)"
-  by (simp add: less_int_def)
-
-instance int :: order
-apply (default)
-apply(auto simp add: zless_le zle_anti_sym)[1]
-apply(rule zle_refl)
-apply(erule zle_trans, assumption)
-apply(erule zle_anti_sym, assumption)
-done
-
-(* Axiom 'linorder_linear' of class 'linorder': *)
-
-lemma zle_linear_raw:
-  shows "le_raw z w \<or> le_raw w z"
-apply(cases w, cases z)
-apply(auto iff: le_raw_def)
-done
-
-lemma zle_linear: 
-  fixes z w::"int"
-  shows "z \<le> w \<or> w \<le> z"
-  by (lifting zle_linear_raw)
-
-instance int :: linorder
-apply(default)
-apply(rule zle_linear)
-done
-
-lemma zadd_left_mono_raw:
-  shows "le_raw i j \<Longrightarrow> le_raw (add_raw k i) (add_raw k j)"
-apply(cases k)
-apply(auto simp add: add_raw_def le_raw_def)
-done
-
-lemma zadd_left_mono: 
-  fixes i j::"int"
-  shows "i \<le> j \<Longrightarrow> k + i \<le> k + j"
-  by (lifting zadd_left_mono_raw)
-
-
-subsection{*Magnitide of an Integer, as a Natural Number: @{term nat}*}
-
-definition
-  "nat_raw \<equiv> \<lambda>(x, y).x - (y::nat)"
-
-quotient_definition
-  "nat2::int \<Rightarrow> nat"
-is
-  "nat_raw"
-
-abbreviation
-  "less_raw x y \<equiv> (le_raw x y \<and> \<not>(intrel x y))"
-
-lemma nat_le_eq_zle_raw:
-  shows "less_raw (0, 0) w \<or> le_raw (0, 0) z \<Longrightarrow> (nat_raw w \<le> nat_raw z) = (le_raw w z)"
-  apply (cases w)
-  apply (cases z)
-  apply (simp add: nat_raw_def le_raw_def)
-  by auto
-
-lemma [quot_respect]:
-  shows "(intrel ===> op =) nat_raw nat_raw"
-  by (auto iff: nat_raw_def)
-
-lemma nat_le_eq_zle: 
-  fixes w z::"int"
-  shows "0 < w \<or> 0 \<le> z \<Longrightarrow> (nat2 w \<le> nat2 z) = (w\<le>z)"
-  unfolding less_int_def
-  by (lifting nat_le_eq_zle_raw)
-
-end
--- a/Quot/Examples/SigmaEx.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,253 +0,0 @@
-theory SigmaEx
-imports Nominal "../Quotient" "../Quotient_List" "../Quotient_Product"
-begin
-
-atom_decl name
-
-datatype robj =
-  rVar "name"
-| rObj "(string \<times> rmethod) list"
-| rInv "robj" "string"
-| rUpd "robj" "string" "rmethod"
-and rmethod =
-  rSig "name" "robj"
-
-inductive
-    alpha_obj :: "robj \<Rightarrow> robj \<Rightarrow> bool" ("_ \<approx>o _" [100, 100] 100)
-and alpha_method :: "rmethod \<Rightarrow> rmethod \<Rightarrow> bool" ("_ \<approx>m _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (rVar a) \<approx>o (rVar b)"
-| a2: "rObj [] \<approx>o rObj []"
-| a3: "rObj t1 \<approx>o rObj t2 \<Longrightarrow> m1 \<approx>m r2 \<Longrightarrow> rObj ((l1, m1) # t1) \<approx>o rObj ((l2, m2) # t2)"
-| a4: "x \<approx>o y \<Longrightarrow> rInv x l1 \<approx>o rInv y l2"
-| a5: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx>o s \<and> (pi \<bullet> a) = b)
-       \<Longrightarrow> rSig a t \<approx>m rSig b s"
-
-lemma alpha_equivps:
-  shows "equivp alpha_obj"
-  and   "equivp alpha_method"
-sorry
-
-quotient_type
-    obj = robj / alpha_obj
-and method = rmethod / alpha_method
-  by (auto intro: alpha_equivps)
-
-quotient_definition
-  "Var :: name \<Rightarrow> obj"
-is
-  "rVar"
-
-quotient_definition
-  "Obj :: (string \<times> method) list \<Rightarrow> obj"
-is
-  "rObj"
-
-quotient_definition
-  "Inv :: obj \<Rightarrow> string \<Rightarrow> obj"
-is
-  "rInv"
-
-quotient_definition
-  "Upd :: obj \<Rightarrow> string \<Rightarrow> method \<Rightarrow> obj"
-is
-  "rUpd"
-
-quotient_definition
-  "Sig :: name \<Rightarrow> obj \<Rightarrow> method"
-is
-  "rSig"
-
-overloading
-  perm_obj \<equiv> "perm :: 'x prm \<Rightarrow> obj \<Rightarrow> obj" (unchecked)
-  perm_method   \<equiv> "perm :: 'x prm \<Rightarrow> method \<Rightarrow> method" (unchecked)
-begin
-
-quotient_definition
-  "perm_obj :: 'x prm \<Rightarrow> obj \<Rightarrow> obj"
-is
-  "(perm::'x prm \<Rightarrow> robj \<Rightarrow> robj)"
-
-quotient_definition
-  "perm_method :: 'x prm \<Rightarrow> method \<Rightarrow> method"
-is
-  "(perm::'x prm \<Rightarrow> rmethod \<Rightarrow> rmethod)"
-
-end
-
-
-
-lemma tolift:
-"\<forall> fvar.
- \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
- \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
- \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<forall> fnil.
- \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
- \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
-
- Ex1 (\<lambda>x.
-(x \<in> (Respects (prod_rel (alpha_obj ===> op =)
-     (prod_rel (list_rel (prod_rel (op =) alpha_method) ===> op =)
-       (prod_rel ((prod_rel (op =) alpha_method) ===> op =)
-         (alpha_method ===> op =)
-       )
-     )))) \<and>
-(\<lambda> (hom_o\<Colon>robj \<Rightarrow> 'a, hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b, hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c, hom_m\<Colon>rmethod \<Rightarrow> 'd).
-
-((\<forall>x. hom_o (rVar x) = fvar x) \<and>
- (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
- (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
- (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
- (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
- (hom_d [] = fnil) \<and>
- (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
- (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
-)) x) "
-sorry
-
-lemma test_to: "Ex1 (\<lambda>x. (x \<in> (Respects alpha_obj)) \<and> P x)"
-ML_prf {* prop_of (#goal (Isar.goal ())) *}
-sorry
-lemma test_tod: "Ex1 (P :: obj \<Rightarrow> bool)"
-apply (lifting test_to)
-done
-
-
-
-
-(*syntax
-  "_expttrn"        :: "pttrn => bool => bool"      ("(3\<exists>\<exists> _./ _)" [0, 10] 10)
-
-translations
-  "\<exists>\<exists> x. P"   == "Ex (%x. P)"
-*)
-
-lemma rvar_rsp[quot_respect]: "(op = ===> alpha_obj) rVar rVar"
-  by (simp add: a1)
-
-lemma robj_rsp[quot_respect]: "(list_rel (prod_rel op = alpha_method) ===> alpha_obj) rObj rObj"
-sorry
-lemma rinv_rsp[quot_respect]: "(alpha_obj ===> op = ===> alpha_obj) rInv rInv"
-sorry
-lemma rupd_rsp[quot_respect]: "(alpha_obj ===> op = ===> alpha_method ===> alpha_obj) rUpd rUpd"
-sorry
-lemma rsig_rsp[quot_respect]: "(op = ===> alpha_obj ===> alpha_method) rSig rSig"
-sorry
-lemma operm_rsp[quot_respect]: "(op = ===> alpha_obj ===> alpha_obj) op \<bullet> op \<bullet>"
-sorry
-
-  
-lemma bex1_bex1reg: "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"
-apply (simp add: Ex1_def Bex1_rel_def in_respects)
-apply clarify
-apply auto
-apply (rule bexI)
-apply assumption
-apply (simp add: in_respects)
-apply (simp add: in_respects)
-apply auto
-done
-
-lemma liftd: "
-Ex1 (\<lambda>(hom_o, hom_d, hom_e, hom_m).
-
- (\<forall>x. hom_o (Var x) = fvar x) \<and>
- (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
- (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
- (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
- (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
- (hom_d [] = fnil) \<and>
- (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
- (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
-)"
-apply (lifting tolift)
-done
-
-lemma tolift':
-"\<forall> fvar.
- \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
- \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
- \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<forall> fnil.
- \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
- \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
-
- \<exists> hom_o\<Colon>robj \<Rightarrow> 'a \<in> Respects (alpha_obj ===> op =).
- \<exists> hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b \<in> Respects (list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<exists> hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c \<in> Respects ((prod_rel (op =) alpha_method) ===> op =).
- \<exists> hom_m\<Colon>rmethod \<Rightarrow> 'd \<in> Respects (alpha_method ===> op =).
-(
- (\<forall>x. hom_o (rVar x) = fvar x) \<and>
- (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
- (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
- (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
- (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
- (hom_d [] = fnil) \<and>
- (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
- (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
-)"
-sorry
-
-lemma liftd': "
-\<exists>hom_o. \<exists>hom_d. \<exists>hom_e. \<exists>hom_m.
-(
- (\<forall>x. hom_o (Var x) = fvar x) \<and>
- (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
- (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
- (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
- (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
- (hom_d [] = fnil) \<and>
- (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
- (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
-)"
-apply (lifting tolift')
-done
-
-lemma tolift'':
-"\<forall> fvar.
- \<forall> fobj\<in>Respects (op = ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<forall> fnvk\<in>Respects (op = ===> alpha_obj ===> op =).
- \<forall> fupd\<in>Respects (op = ===> op = ===> alpha_obj ===> op = ===> alpha_method ===> op =).
- \<forall> fcns\<in>Respects (op = ===> op = ===> prod_rel (op =) alpha_method ===> list_rel (prod_rel (op =) alpha_method) ===> op =).
- \<forall> fnil.
- \<forall> fpar\<in>Respects (op = ===> op = ===> alpha_method ===> op =).
- \<forall> fsgm\<in>Respects (op = ===> (op = ===> alpha_obj) ===> op =).
-
- Bex1_rel (alpha_obj ===> op =) (\<lambda>hom_o\<Colon>robj \<Rightarrow> 'a .
- Bex1_rel (list_rel (prod_rel (op =) alpha_method) ===> op =) (\<lambda>hom_d\<Colon>(char list \<times> rmethod) list \<Rightarrow> 'b.
- Bex1_rel ((prod_rel (op =) alpha_method) ===> op =) (\<lambda>hom_e\<Colon>char list \<times> rmethod \<Rightarrow> 'c.
- Bex1_rel (alpha_method ===> op =) (\<lambda>hom_m\<Colon>rmethod \<Rightarrow> 'd.
-(
- (\<forall>x. hom_o (rVar x) = fvar x) \<and>
- (\<forall>d. hom_o (rObj d) = fobj (hom_d d) d) \<and>
- (\<forall>a l. hom_o (rInv a l) = fnvk (hom_o a) a l) \<and>
- (\<forall>a l m. hom_o (rUpd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
- (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
- (hom_d [] = fnil) \<and>
- (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
- (\<forall>x a. hom_m (rSig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
-)
-))))"
-sorry
-
-lemma liftd'': "
-\<exists>!hom_o. \<exists>!hom_d. \<exists>!hom_e. \<exists>!hom_m.
-(
- (\<forall>x. hom_o (Var x) = fvar x) \<and>
- (\<forall>d. hom_o (Obj d) = fobj (hom_d d) d) \<and>
- (\<forall>a l. hom_o (Inv a l) = fnvk (hom_o a) a l) \<and>
- (\<forall>a l m. hom_o (Upd a l m) = fupd (hom_o a) (hom_m m) a l m) \<and>
- (\<forall>e d. hom_d (e # d) = fcns (hom_e e) (hom_d d) e d) \<and>
- (hom_d [] = fnil) \<and>
- (\<forall>l m. hom_e (l, m) = fpar (hom_m m) l m) \<and>
- (\<forall>x a. hom_m (Sig x a) = fsgm (\<lambda>y. hom_o ([(x, y)] \<bullet> a)) (\<lambda>y. [(x, y)] \<bullet> a))
-)"
-apply (lifting tolift'')
-done
-
-
-end
-
--- a/Quot/Examples/Terms.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,271 +0,0 @@
-theory Terms
-imports Nominal "../Quotient" "../Quotient_List"
-begin
-
-atom_decl name
-
-text {* primrec seems to be genarally faster than fun *}
-
-section {*** lets with binding patterns ***}
-
-datatype trm1 =
-  Vr1 "name"
-| Ap1 "trm1" "trm1"
-| Lm1 "name" "trm1"        --"name is bound in trm1"
-| Lt1 "bp" "trm1" "trm1"   --"all variables in bp are bound in the 2nd trm1"
-and bp =
-  BUnit
-| BVr "name"
-| BPr "bp" "bp"
-
-(* to be given by the user *)
-primrec 
-  bv1
-where
-  "bv1 (BUnit) = {}"
-| "bv1 (BVr x) = {x}"
-| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)"
-
-(* needs to be calculated by the package *)
-primrec 
-  fv_trm1 and fv_bp
-where
-  "fv_trm1 (Vr1 x) = {x}"
-| "fv_trm1 (Ap1 t1 t2) = (fv_trm1 t1) \<union> (fv_trm1 t2)"
-| "fv_trm1 (Lm1 x t) = (fv_trm1 t) - {x}"
-| "fv_trm1 (Lt1 bp t1 t2) = (fv_trm1 t1) \<union> (fv_trm1 t2 - bv1 bp)"
-| "fv_bp (BUnit) = {}"
-| "fv_bp (BVr x) = {x}"
-| "fv_bp (BPr b1 b2) = (fv_bp b1) \<union> (fv_bp b2)"
-
-(* needs to be stated by the package *)
-overloading
-  perm_trm1 \<equiv> "perm :: 'x prm \<Rightarrow> trm1 \<Rightarrow> trm1"   (unchecked)
-  perm_bp \<equiv> "perm :: 'x prm \<Rightarrow> bp \<Rightarrow> bp" (unchecked)
-begin
-
-primrec
-  perm_trm1 and perm_bp
-where
-  "perm_trm1 pi (Vr1 a) = Vr1 (pi \<bullet> a)"
-| "perm_trm1 pi (Ap1 t1 t2) = Ap1 (perm_trm1 pi t1) (perm_trm1 pi t2)"
-| "perm_trm1 pi (Lm1 a t) = Lm1 (pi \<bullet> a) (perm_trm1 pi t)"
-| "perm_trm1 pi (Lt1 bp t1 t2) = Lt1 (perm_bp pi bp) (perm_trm1 pi t1) (perm_trm1 pi t2)"
-| "perm_bp pi (BUnit) = BUnit"
-| "perm_bp pi (BVr a) = BVr (pi \<bullet> a)"
-| "perm_bp pi (BPr bp1 bp2) = BPr (perm_bp pi bp1) (perm_bp pi bp2)"
-
-end
-
-inductive
-  alpha1 :: "trm1 \<Rightarrow> trm1 \<Rightarrow> bool" ("_ \<approx>1 _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (Vr1 a) \<approx>1 (Vr1 b)"
-| a2: "\<lbrakk>t1 \<approx>1 t2; s1 \<approx>1 s2\<rbrakk> \<Longrightarrow> Ap1 t1 s1 \<approx>1 Ap1 t2 s2"
-| a3: "\<exists>pi::name prm. (fv_trm1 t - {a} = fv_trm1 s - {b} \<and> 
-                      (fv_trm1 t - {a})\<sharp>* pi \<and> 
-                      (pi \<bullet> t) \<approx>1 s \<and> (pi \<bullet> a) = b)
-       \<Longrightarrow> Lm1 a t \<approx>1 Lm1 b s"
-| a4: "\<exists>pi::name prm.(
-         t1 \<approx>1 t2 \<and>
-         (fv_trm1 s1 - fv_bp b1 = fv_trm1 s2 - fv_bp b2) \<and>
-         (fv_trm1 s1 - fv_bp b1) \<sharp>* pi \<and>
-         (pi \<bullet> s1 = s2)                    (* Optional: \<and> (pi \<bullet> b1 = b2) *)
-       ) \<Longrightarrow> Lt1 b1 t1 s1 \<approx>1 Lt1 b2 t2 s2"
-
-lemma alpha1_equivp: "equivp alpha1" sorry
-
-quotient_type qtrm1 = trm1 / alpha1
-  by (rule alpha1_equivp)
-
-
-section {*** lets with single assignments ***}
-
-datatype trm2 =
-  Vr2 "name"
-| Ap2 "trm2" "trm2"
-| Lm2 "name" "trm2"
-| Lt2 "assign" "trm2"
-and assign =
-  As "name" "trm2"
-
-(* to be given by the user *)
-primrec 
-  bv2
-where
-  "bv2 (As x t) = {x}"
-
-(* needs to be calculated by the package *)
-primrec
-  fv_trm2 and fv_assign
-where
-  "fv_trm2 (Vr2 x) = {x}"
-| "fv_trm2 (Ap2 t1 t2) = (fv_trm2 t1) \<union> (fv_trm2 t2)"
-| "fv_trm2 (Lm2 x t) = (fv_trm2 t) - {x}"
-| "fv_trm2 (Lt2 as t) = (fv_trm2 t - bv2 as) \<union> (fv_assign as)"
-| "fv_assign (As x t) = (fv_trm2 t)"
-
-(* needs to be stated by the package *)
-overloading
-  perm_trm2 \<equiv> "perm :: 'x prm \<Rightarrow> trm2 \<Rightarrow> trm2"   (unchecked)
-  perm_assign \<equiv> "perm :: 'x prm \<Rightarrow> assign \<Rightarrow> assign" (unchecked)
-begin
-
-primrec
-  perm_trm2 and perm_assign
-where
-  "perm_trm2 pi (Vr2 a) = Vr2 (pi \<bullet> a)"
-| "perm_trm2 pi (Ap2 t1 t2) = Ap2 (perm_trm2 pi t1) (perm_trm2 pi t2)"
-| "perm_trm2 pi (Lm2 a t) = Lm2 (pi \<bullet> a) (perm_trm2 pi t)"
-| "perm_trm2 pi (Lt2 as t) = Lt2 (perm_assign pi as) (perm_trm2 pi t)"
-| "perm_assign pi (As a t) = As (pi \<bullet> a) (perm_trm2 pi t)"
-
-end
-
-inductive
-  alpha2 :: "trm2 \<Rightarrow> trm2 \<Rightarrow> bool" ("_ \<approx>2 _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (Vr2 a) \<approx>2 (Vr2 b)"
-| a2: "\<lbrakk>t1 \<approx>2 t2; s1 \<approx>2 s2\<rbrakk> \<Longrightarrow> Ap2 t1 s1 \<approx>2 Ap2 t2 s2"
-| a3: "\<exists>pi::name prm. (fv_trm2 t - {a} = fv_trm2 s - {b} \<and> 
-                      (fv_trm2 t - {a})\<sharp>* pi \<and> 
-                      (pi \<bullet> t) \<approx>2 s \<and> 
-                      (pi \<bullet> a) = b)
-       \<Longrightarrow> Lm2 a t \<approx>2 Lm2 b s"
-| a4: "\<exists>pi::name prm. (
-         fv_trm2 t1 - fv_assign b1 = fv_trm2 t2 - fv_assign b2 \<and>
-         (fv_trm2 t1 - fv_assign b1) \<sharp>* pi \<and>
-         pi \<bullet> t1 = t2       (* \<and> (pi \<bullet> b1 = b2) *)
-       ) \<Longrightarrow> Lt2 b1 t1 \<approx>2 Lt2 b2 t2"
-
-lemma alpha2_equivp: "equivp alpha2" sorry
-
-quotient_type qtrm2 = trm2 / alpha2
-  by (rule alpha2_equivp)
-
-section {*** lets with many assignments ***}
-
-datatype trm3 =
-  Vr3 "name"
-| Ap3 "trm3" "trm3"
-| Lm3 "name" "trm3"
-| Lt3 "assigns" "trm3"
-and assigns =
-  ANil
-| ACons "name" "trm3" "assigns"
-
-(* to be given by the user *)
-primrec 
-  bv3
-where
-  "bv3 ANil = {}"
-| "bv3 (ACons x t as) = {x} \<union> (bv3 as)"
-
-primrec
-  fv_trm3 and fv_assigns
-where
-  "fv_trm3 (Vr3 x) = {x}"
-| "fv_trm3 (Ap3 t1 t2) = (fv_trm3 t1) \<union> (fv_trm3 t2)"
-| "fv_trm3 (Lm3 x t) = (fv_trm3 t) - {x}"
-| "fv_trm3 (Lt3 as t) = (fv_trm3 t - bv3 as) \<union> (fv_assigns as)"
-| "fv_assigns (ANil) = {}"
-| "fv_assigns (ACons x t as) = (fv_trm3 t) \<union> (fv_assigns as)"
-
-(* needs to be stated by the package *)
-overloading
-  perm_trm3 \<equiv> "perm :: 'x prm \<Rightarrow> trm3 \<Rightarrow> trm3"   (unchecked)
-  perm_assigns \<equiv> "perm :: 'x prm \<Rightarrow> assigns \<Rightarrow> assigns" (unchecked)
-begin
-
-primrec
-  perm_trm3 and perm_assigns
-where
-  "perm_trm3 pi (Vr3 a) = Vr3 (pi \<bullet> a)"
-| "perm_trm3 pi (Ap3 t1 t2) = Ap3 (perm_trm3 pi t1) (perm_trm3 pi t2)"
-| "perm_trm3 pi (Lm3 a t) = Lm3 (pi \<bullet> a) (perm_trm3 pi t)"
-| "perm_trm3 pi (Lt3 as t) = Lt3 (perm_assigns pi as) (perm_trm3 pi t)"
-| "perm_assigns pi (ANil) = ANil"
-| "perm_assigns pi (ACons a t as) = ACons (pi \<bullet> a) (perm_trm3 pi t) (perm_assigns pi as)"
-
-end
-
-inductive
-  alpha3 :: "trm3 \<Rightarrow> trm3 \<Rightarrow> bool" ("_ \<approx>3 _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (Vr3 a) \<approx>3 (Vr3 b)"
-| a2: "\<lbrakk>t1 \<approx>3 t2; s1 \<approx>3 s2\<rbrakk> \<Longrightarrow> Ap3 t1 s1 \<approx>3 Ap3 t2 s2"
-| a3: "\<exists>pi::name prm. (fv_trm3 t - {a} = fv_trm3 s - {b} \<and> 
-                      (fv_trm3 t - {a})\<sharp>* pi \<and> 
-                      (pi \<bullet> t) \<approx>3 s \<and> 
-                      (pi \<bullet> a) = b)
-       \<Longrightarrow> Lm3 a t \<approx>3 Lm3 b s"
-| a4: "\<exists>pi::name prm. (
-         fv_trm3 t1 - fv_assigns b1 = fv_trm3 t2 - fv_assigns b2 \<and>
-         (fv_trm3 t1 - fv_assigns b1) \<sharp>* pi \<and>
-         pi \<bullet> t1 = t2      (* \<and> (pi \<bullet> b1 = b2)  *)
-       ) \<Longrightarrow> Lt3 b1 t1 \<approx>3 Lt3 b2 t2"
-
-lemma alpha3_equivp: "equivp alpha3" sorry
-
-quotient_type qtrm3 = trm3 / alpha3
-  by (rule alpha3_equivp)
-
-
-section {*** lam with indirect list recursion ***}
-
-datatype trm4 =
-  Vr4 "name"
-| Ap4 "trm4" "trm4 list"
-| Lm4 "name" "trm4"
-
-thm trm4.recs
-
-primrec
-  fv_trm4 and fv_trm4_list
-where
-  "fv_trm4 (Vr4 x) = {x}"
-| "fv_trm4 (Ap4 t ts) = (fv_trm4 t) \<union> (fv_trm4_list ts)"
-| "fv_trm4 (Lm4 x t) = (fv_trm4 t) - {x}"
-| "fv_trm4_list ([]) = {}"
-| "fv_trm4_list (t#ts) = (fv_trm4 t) \<union> (fv_trm4_list ts)"
-
-
-(* needs to be stated by the package *)
-(* there cannot be a clause for lists, as *) 
-(* permutations are  already defined in Nominal (also functions, options, and so on) *)
-overloading
-  perm_trm4 \<equiv> "perm :: 'x prm \<Rightarrow> trm4 \<Rightarrow> trm4"   (unchecked)
-begin
-
-primrec
-  perm_trm4 
-where
-  "perm_trm4 pi (Vr4 a) = Vr4 (pi \<bullet> a)"
-| "perm_trm4 pi (Ap4 t ts) = Ap4 (perm_trm4 pi t) (pi \<bullet> ts)"
-| "perm_trm4 pi (Lm4 a t) = Lm4 (pi \<bullet> a) (perm_trm4 pi t)"
-
-end
-
-inductive
-    alpha4 :: "trm4 \<Rightarrow> trm4 \<Rightarrow> bool" ("_ \<approx>4 _" [100, 100] 100)
-and alpha4list :: "trm4 list \<Rightarrow> trm4 list \<Rightarrow> bool" ("_ \<approx>4list _" [100, 100] 100) 
-where
-  a1: "a = b \<Longrightarrow> (Vr4 a) \<approx>4 (Vr4 b)"
-| a2: "\<lbrakk>t1 \<approx>4 t2; s1 \<approx>4list s2\<rbrakk> \<Longrightarrow> Ap4 t1 s1 \<approx>4 Ap4 t2 s2"
-| a4: "\<exists>pi::name prm. (fv_trm4 t - {a} = fv_trm4 s - {b} \<and> 
-                      (fv_trm4 t - {a})\<sharp>* pi \<and> 
-                      (pi \<bullet> t) \<approx>4 s \<and> 
-                      (pi \<bullet> a) = b)
-       \<Longrightarrow> Lm4 a t \<approx>4 Lm4 b s"
-| a5: "[] \<approx>4list []"
-| a6: "\<lbrakk>t \<approx>4 s; ts \<approx>4list ss\<rbrakk> \<Longrightarrow> (t#ts) \<approx>4list (s#ss)"
-
-lemma alpha4_equivp: "equivp alpha4" sorry
-lemma alpha4list_equivp: "equivp alpha4list" sorry
-
-quotient_type 
-  qtrm4 = trm4 / alpha4 and
-  qtrm4list = "trm4 list" / alpha4list
-  by (simp_all add: alpha4_equivp alpha4list_equivp)
-
-end
--- a/Quot/Quotient.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,797 +0,0 @@
-(*  Title:      Quotient.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-*)
-
-theory Quotient
-imports Plain ATP_Linkup
-uses
-  ("quotient_info.ML")
-  ("quotient_typ.ML")
-  ("quotient_def.ML")
-  ("quotient_term.ML")
-  ("quotient_tacs.ML")
-begin
-
-
-text {*
-  Basic definition for equivalence relations
-  that are represented by predicates.
-*}
-
-definition
-  "equivp E \<equiv> \<forall>x y. E x y = (E x = E y)"
-
-definition
-  "reflp E \<equiv> \<forall>x. E x x"
-
-definition
-  "symp E \<equiv> \<forall>x y. E x y \<longrightarrow> E y x"
-
-definition
-  "transp E \<equiv> \<forall>x y z. E x y \<and> E y z \<longrightarrow> E x z"
-
-lemma equivp_reflp_symp_transp:
-  shows "equivp E = (reflp E \<and> symp E \<and> transp E)"
-  unfolding equivp_def reflp_def symp_def transp_def expand_fun_eq
-  by blast
-
-lemma equivp_reflp:
-  shows "equivp E \<Longrightarrow> E x x"
-  by (simp only: equivp_reflp_symp_transp reflp_def)
-
-lemma equivp_symp:
-  shows "equivp E \<Longrightarrow> E x y \<Longrightarrow> E y x"
-  by (metis equivp_reflp_symp_transp symp_def)
-
-lemma equivp_transp:
-  shows "equivp E \<Longrightarrow> E x y \<Longrightarrow> E y z \<Longrightarrow> E x z"
-  by (metis equivp_reflp_symp_transp transp_def)
-
-lemma equivpI:
-  assumes "reflp R" "symp R" "transp R"
-  shows "equivp R"
-  using assms by (simp add: equivp_reflp_symp_transp)
-
-lemma identity_equivp:
-  shows "equivp (op =)"
-  unfolding equivp_def
-  by auto
-
-text {* Partial equivalences: not yet used anywhere *}
-
-definition
-  "part_equivp E \<equiv> (\<exists>x. E x x) \<and> (\<forall>x y. E x y = (E x x \<and> E y y \<and> (E x = E y)))"
-
-lemma equivp_implies_part_equivp:
-  assumes a: "equivp E"
-  shows "part_equivp E"
-  using a
-  unfolding equivp_def part_equivp_def
-  by auto
-
-text {* Composition of Relations *}
-
-abbreviation
-  rel_conj (infixr "OOO" 75)
-where
-  "r1 OOO r2 \<equiv> r1 OO r2 OO r1"
-
-lemma eq_comp_r:
-  shows "((op =) OOO R) = R"
-  by (auto simp add: expand_fun_eq)
-
-section {* Respects predicate *}
-
-definition
-  Respects
-where
-  "Respects R x \<equiv> R x x"
-
-lemma in_respects:
-  shows "(x \<in> Respects R) = R x x"
-  unfolding mem_def Respects_def
-  by simp
-
-section {* Function map and function relation *}
-
-definition
-  fun_map (infixr "--->" 55)
-where
-[simp]: "fun_map f g h x = g (h (f x))"
-
-definition
-  fun_rel (infixr "===>" 55)
-where
-[simp]: "fun_rel E1 E2 f g = (\<forall>x y. E1 x y \<longrightarrow> E2 (f x) (g y))"
-
-
-lemma fun_map_id:
-  shows "(id ---> id) = id"
-  by (simp add: expand_fun_eq id_def)
-
-lemma fun_rel_eq:
-  shows "((op =) ===> (op =)) = (op =)"
-  by (simp add: expand_fun_eq)
-
-lemma fun_rel_id:
-  assumes a: "\<And>x y. R1 x y \<Longrightarrow> R2 (f x) (g y)"
-  shows "(R1 ===> R2) f g"
-  using a by simp
-
-lemma fun_rel_id_asm:
-  assumes a: "\<And>x y. R1 x y \<Longrightarrow> (A \<longrightarrow> R2 (f x) (g y))"
-  shows "A \<longrightarrow> (R1 ===> R2) f g"
-  using a by auto
-
-
-section {* Quotient Predicate *}
-
-definition
-  "Quotient E Abs Rep \<equiv>
-     (\<forall>a. Abs (Rep a) = a) \<and> (\<forall>a. E (Rep a) (Rep a)) \<and>
-     (\<forall>r s. E r s = (E r r \<and> E s s \<and> (Abs r = Abs s)))"
-
-lemma Quotient_abs_rep:
-  assumes a: "Quotient E Abs Rep"
-  shows "Abs (Rep a) = a"
-  using a
-  unfolding Quotient_def
-  by simp
-
-lemma Quotient_rep_reflp:
-  assumes a: "Quotient E Abs Rep"
-  shows "E (Rep a) (Rep a)"
-  using a
-  unfolding Quotient_def
-  by blast
-
-lemma Quotient_rel:
-  assumes a: "Quotient E Abs Rep"
-  shows " E r s = (E r r \<and> E s s \<and> (Abs r = Abs s))"
-  using a
-  unfolding Quotient_def
-  by blast
-
-lemma Quotient_rel_rep:
-  assumes a: "Quotient R Abs Rep"
-  shows "R (Rep a) (Rep b) = (a = b)"
-  using a
-  unfolding Quotient_def
-  by metis
-
-lemma Quotient_rep_abs:
-  assumes a: "Quotient R Abs Rep"
-  shows "R r r \<Longrightarrow> R (Rep (Abs r)) r"
-  using a unfolding Quotient_def
-  by blast
-
-lemma Quotient_rel_abs:
-  assumes a: "Quotient E Abs Rep"
-  shows "E r s \<Longrightarrow> Abs r = Abs s"
-  using a unfolding Quotient_def
-  by blast
-
-lemma Quotient_symp:
-  assumes a: "Quotient E Abs Rep"
-  shows "symp E"
-  using a unfolding Quotient_def symp_def
-  by metis
-
-lemma Quotient_transp:
-  assumes a: "Quotient E Abs Rep"
-  shows "transp E"
-  using a unfolding Quotient_def transp_def
-  by metis
-
-lemma identity_quotient:
-  shows "Quotient (op =) id id"
-  unfolding Quotient_def id_def
-  by blast
-
-lemma fun_quotient:
-  assumes q1: "Quotient R1 abs1 rep1"
-  and     q2: "Quotient R2 abs2 rep2"
-  shows "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)"
-proof -
-  have "\<forall>a. (rep1 ---> abs2) ((abs1 ---> rep2) a) = a"
-    using q1 q2
-    unfolding Quotient_def
-    unfolding expand_fun_eq
-    by simp
-  moreover
-  have "\<forall>a. (R1 ===> R2) ((abs1 ---> rep2) a) ((abs1 ---> rep2) a)"
-    using q1 q2
-    unfolding Quotient_def
-    by (simp (no_asm)) (metis)
-  moreover
-  have "\<forall>r s. (R1 ===> R2) r s = ((R1 ===> R2) r r \<and> (R1 ===> R2) s s \<and>
-        (rep1 ---> abs2) r  = (rep1 ---> abs2) s)"
-    unfolding expand_fun_eq
-    apply(auto)
-    using q1 q2 unfolding Quotient_def
-    apply(metis)
-    using q1 q2 unfolding Quotient_def
-    apply(metis)
-    using q1 q2 unfolding Quotient_def
-    apply(metis)
-    using q1 q2 unfolding Quotient_def
-    apply(metis)
-    done
-  ultimately
-  show "Quotient (R1 ===> R2) (rep1 ---> abs2) (abs1 ---> rep2)"
-    unfolding Quotient_def by blast
-qed
-
-lemma abs_o_rep:
-  assumes a: "Quotient R Abs Rep"
-  shows "Abs o Rep = id"
-  unfolding expand_fun_eq
-  by (simp add: Quotient_abs_rep[OF a])
-
-lemma equals_rsp:
-  assumes q: "Quotient R Abs Rep"
-  and     a: "R xa xb" "R ya yb"
-  shows "R xa ya = R xb yb"
-  using a Quotient_symp[OF q] Quotient_transp[OF q]
-  unfolding symp_def transp_def
-  by blast
-
-lemma lambda_prs:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "(Rep1 ---> Abs2) (\<lambda>x. Rep2 (f (Abs1 x))) = (\<lambda>x. f x)"
-  unfolding expand_fun_eq
-  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
-  by simp
-
-lemma lambda_prs1:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "(Rep1 ---> Abs2) (\<lambda>x. (Abs1 ---> Rep2) f x) = (\<lambda>x. f x)"
-  unfolding expand_fun_eq
-  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2]
-  by simp
-
-lemma rep_abs_rsp:
-  assumes q: "Quotient R Abs Rep"
-  and     a: "R x1 x2"
-  shows "R x1 (Rep (Abs x2))"
-  using a Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]
-  by metis
-
-lemma rep_abs_rsp_left:
-  assumes q: "Quotient R Abs Rep"
-  and     a: "R x1 x2"
-  shows "R (Rep (Abs x1)) x2"
-  using a Quotient_rel[OF q] Quotient_abs_rep[OF q] Quotient_rep_reflp[OF q]
-  by metis
-
-text{*
-  In the following theorem R1 can be instantiated with anything,
-  but we know some of the types of the Rep and Abs functions;
-  so by solving Quotient assumptions we can get a unique R1 that
-  will be provable; which is why we need to use apply_rsp and
-  not the primed version *}
-
-lemma apply_rsp:
-  fixes f g::"'a \<Rightarrow> 'c"
-  assumes q: "Quotient R1 Abs1 Rep1"
-  and     a: "(R1 ===> R2) f g" "R1 x y"
-  shows "R2 (f x) (g y)"
-  using a by simp
-
-lemma apply_rsp':
-  assumes a: "(R1 ===> R2) f g" "R1 x y"
-  shows "R2 (f x) (g y)"
-  using a by simp
-
-section {* lemmas for regularisation of ball and bex *}
-
-lemma ball_reg_eqv:
-  fixes P :: "'a \<Rightarrow> bool"
-  assumes a: "equivp R"
-  shows "Ball (Respects R) P = (All P)"
-  using a
-  unfolding equivp_def
-  by (auto simp add: in_respects)
-
-lemma bex_reg_eqv:
-  fixes P :: "'a \<Rightarrow> bool"
-  assumes a: "equivp R"
-  shows "Bex (Respects R) P = (Ex P)"
-  using a
-  unfolding equivp_def
-  by (auto simp add: in_respects)
-
-lemma ball_reg_right:
-  assumes a: "\<And>x. R x \<Longrightarrow> P x \<longrightarrow> Q x"
-  shows "All P \<longrightarrow> Ball R Q"
-  using a by (metis COMBC_def Collect_def Collect_mem_eq)
-
-lemma bex_reg_left:
-  assumes a: "\<And>x. R x \<Longrightarrow> Q x \<longrightarrow> P x"
-  shows "Bex R Q \<longrightarrow> Ex P"
-  using a by (metis COMBC_def Collect_def Collect_mem_eq)
-
-lemma ball_reg_left:
-  assumes a: "equivp R"
-  shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ball (Respects R) Q \<longrightarrow> All P"
-  using a by (metis equivp_reflp in_respects)
-
-lemma bex_reg_right:
-  assumes a: "equivp R"
-  shows "(\<And>x. (Q x \<longrightarrow> P x)) \<Longrightarrow> Ex Q \<longrightarrow> Bex (Respects R) P"
-  using a by (metis equivp_reflp in_respects)
-
-lemma ball_reg_eqv_range:
-  fixes P::"'a \<Rightarrow> bool"
-  and x::"'a"
-  assumes a: "equivp R2"
-  shows   "(Ball (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = All (\<lambda>f. P (f x)))"
-  apply(rule iffI)
-  apply(rule allI)
-  apply(drule_tac x="\<lambda>y. f x" in bspec)
-  apply(simp add: in_respects)
-  apply(rule impI)
-  using a equivp_reflp_symp_transp[of "R2"]
-  apply(simp add: reflp_def)
-  apply(simp)
-  apply(simp)
-  done
-
-lemma bex_reg_eqv_range:
-  assumes a: "equivp R2"
-  shows   "(Bex (Respects (R1 ===> R2)) (\<lambda>f. P (f x)) = Ex (\<lambda>f. P (f x)))"
-  apply(auto)
-  apply(rule_tac x="\<lambda>y. f x" in bexI)
-  apply(simp)
-  apply(simp add: Respects_def in_respects)
-  apply(rule impI)
-  using a equivp_reflp_symp_transp[of "R2"]
-  apply(simp add: reflp_def)
-  done
-
-(* Next four lemmas are unused *)
-lemma all_reg:
-  assumes a: "!x :: 'a. (P x --> Q x)"
-  and     b: "All P"
-  shows "All Q"
-  using a b by (metis)
-
-lemma ex_reg:
-  assumes a: "!x :: 'a. (P x --> Q x)"
-  and     b: "Ex P"
-  shows "Ex Q"
-  using a b by metis
-
-lemma ball_reg:
-  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
-  and     b: "Ball R P"
-  shows "Ball R Q"
-  using a b by (metis COMBC_def Collect_def Collect_mem_eq)
-
-lemma bex_reg:
-  assumes a: "!x :: 'a. (R x --> P x --> Q x)"
-  and     b: "Bex R P"
-  shows "Bex R Q"
-  using a b by (metis COMBC_def Collect_def Collect_mem_eq)
-
-
-lemma ball_all_comm:
-  assumes "\<And>y. (\<forall>x\<in>P. A x y) \<longrightarrow> (\<forall>x. B x y)"
-  shows "(\<forall>x\<in>P. \<forall>y. A x y) \<longrightarrow> (\<forall>x. \<forall>y. B x y)"
-  using assms by auto
-
-lemma bex_ex_comm:
-  assumes "(\<exists>y. \<exists>x. A x y) \<longrightarrow> (\<exists>y. \<exists>x\<in>P. B x y)"
-  shows "(\<exists>x. \<exists>y. A x y) \<longrightarrow> (\<exists>x\<in>P. \<exists>y. B x y)"
-  using assms by auto
-
-section {* Bounded abstraction *}
-
-definition
-  Babs :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
-where
-  "x \<in> p \<Longrightarrow> Babs p m x = m x"
-
-lemma babs_rsp:
-  assumes q: "Quotient R1 Abs1 Rep1"
-  and     a: "(R1 ===> R2) f g"
-  shows      "(R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)"
-  apply (auto simp add: Babs_def in_respects)
-  apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
-  using a apply (simp add: Babs_def)
-  apply (simp add: in_respects)
-  using Quotient_rel[OF q]
-  by metis
-
-lemma babs_prs:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "((Rep1 ---> Abs2) (Babs (Respects R1) ((Abs1 ---> Rep2) f))) = f"
-  apply (rule ext)
-  apply (simp)
-  apply (subgoal_tac "Rep1 x \<in> Respects R1")
-  apply (simp add: Babs_def Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
-  apply (simp add: in_respects Quotient_rel_rep[OF q1])
-  done
-
-lemma babs_simp:
-  assumes q: "Quotient R1 Abs Rep"
-  shows "((R1 ===> R2) (Babs (Respects R1) f) (Babs (Respects R1) g)) = ((R1 ===> R2) f g)"
-  apply(rule iffI)
-  apply(simp_all only: babs_rsp[OF q])
-  apply(auto simp add: Babs_def)
-  apply (subgoal_tac "x \<in> Respects R1 \<and> y \<in> Respects R1")
-  apply(metis Babs_def)
-  apply (simp add: in_respects)
-  using Quotient_rel[OF q]
-  by metis
-
-(* If a user proves that a particular functional relation
-   is an equivalence this may be useful in regularising *)
-lemma babs_reg_eqv:
-  shows "equivp R \<Longrightarrow> Babs (Respects R) P = P"
-  by (simp add: expand_fun_eq Babs_def in_respects equivp_reflp)
-
-
-(* 3 lemmas needed for proving repabs_inj *)
-lemma ball_rsp:
-  assumes a: "(R ===> (op =)) f g"
-  shows "Ball (Respects R) f = Ball (Respects R) g"
-  using a by (simp add: Ball_def in_respects)
-
-lemma bex_rsp:
-  assumes a: "(R ===> (op =)) f g"
-  shows "(Bex (Respects R) f = Bex (Respects R) g)"
-  using a by (simp add: Bex_def in_respects)
-
-lemma bex1_rsp:
-  assumes a: "(R ===> (op =)) f g"
-  shows "Ex1 (\<lambda>x. x \<in> Respects R \<and> f x) = Ex1 (\<lambda>x. x \<in> Respects R \<and> g x)"
-  using a
-  by (simp add: Ex1_def in_respects) auto
-
-(* 2 lemmas needed for cleaning of quantifiers *)
-lemma all_prs:
-  assumes a: "Quotient R absf repf"
-  shows "Ball (Respects R) ((absf ---> id) f) = All f"
-  using a unfolding Quotient_def Ball_def in_respects fun_map_def id_apply
-  by metis
-
-lemma ex_prs:
-  assumes a: "Quotient R absf repf"
-  shows "Bex (Respects R) ((absf ---> id) f) = Ex f"
-  using a unfolding Quotient_def Bex_def in_respects fun_map_def id_apply
-  by metis
-
-section {* Bex1_rel quantifier *}
-
-definition
-  Bex1_rel :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
-where
-  "Bex1_rel R P \<longleftrightarrow> (\<exists>x \<in> Respects R. P x) \<and> (\<forall>x \<in> Respects R. \<forall>y \<in> Respects R. ((P x \<and> P y) \<longrightarrow> (R x y)))"
-
-lemma bex1_rel_aux:
-  "\<lbrakk>\<forall>xa ya. R xa ya \<longrightarrow> x xa = y ya; Bex1_rel R x\<rbrakk> \<Longrightarrow> Bex1_rel R y"
-  unfolding Bex1_rel_def
-  apply (erule conjE)+
-  apply (erule bexE)
-  apply rule
-  apply (rule_tac x="xa" in bexI)
-  apply metis
-  apply metis
-  apply rule+
-  apply (erule_tac x="xaa" in ballE)
-  prefer 2
-  apply (metis)
-  apply (erule_tac x="ya" in ballE)
-  prefer 2
-  apply (metis)
-  apply (metis in_respects)
-  done
-
-lemma bex1_rel_aux2:
-  "\<lbrakk>\<forall>xa ya. R xa ya \<longrightarrow> x xa = y ya; Bex1_rel R y\<rbrakk> \<Longrightarrow> Bex1_rel R x"
-  unfolding Bex1_rel_def
-  apply (erule conjE)+
-  apply (erule bexE)
-  apply rule
-  apply (rule_tac x="xa" in bexI)
-  apply metis
-  apply metis
-  apply rule+
-  apply (erule_tac x="xaa" in ballE)
-  prefer 2
-  apply (metis)
-  apply (erule_tac x="ya" in ballE)
-  prefer 2
-  apply (metis)
-  apply (metis in_respects)
-  done
-
-lemma bex1_rel_rsp:
-  assumes a: "Quotient R absf repf"
-  shows "((R ===> op =) ===> op =) (Bex1_rel R) (Bex1_rel R)"
-  apply simp
-  apply clarify
-  apply rule
-  apply (simp_all add: bex1_rel_aux bex1_rel_aux2)
-  apply (erule bex1_rel_aux2)
-  apply assumption
-  done
-
-
-lemma ex1_prs:
-  assumes a: "Quotient R absf repf"
-  shows "((absf ---> id) ---> id) (Bex1_rel R) f = Ex1 f"
-apply simp
-apply (subst Bex1_rel_def)
-apply (subst Bex_def)
-apply (subst Ex1_def)
-apply simp
-apply rule
- apply (erule conjE)+
- apply (erule_tac exE)
- apply (erule conjE)
- apply (subgoal_tac "\<forall>y. R y y \<longrightarrow> f (absf y) \<longrightarrow> R x y")
-  apply (rule_tac x="absf x" in exI)
-  apply (simp)
-  apply rule+
-  using a unfolding Quotient_def
-  apply metis
- apply rule+
- apply (erule_tac x="x" in ballE)
-  apply (erule_tac x="y" in ballE)
-   apply simp
-  apply (simp add: in_respects)
- apply (simp add: in_respects)
-apply (erule_tac exE)
- apply rule
- apply (rule_tac x="repf x" in exI)
- apply (simp only: in_respects)
-  apply rule
- apply (metis Quotient_rel_rep[OF a])
-using a unfolding Quotient_def apply (simp)
-apply rule+
-using a unfolding Quotient_def in_respects
-apply metis
-done
-
-lemma bex1_bexeq_reg: "(\<exists>!x\<in>Respects R. P x) \<longrightarrow> (Bex1_rel R (\<lambda>x. P x))"
-  apply (simp add: Ex1_def Bex1_rel_def in_respects)
-  apply clarify
-  apply auto
-  apply (rule bexI)
-  apply assumption
-  apply (simp add: in_respects)
-  apply (simp add: in_respects)
-  apply auto
-  done
-
-section {* Various respects and preserve lemmas *}
-
-lemma quot_rel_rsp:
-  assumes a: "Quotient R Abs Rep"
-  shows "(R ===> R ===> op =) R R"
-  apply(rule fun_rel_id)+
-  apply(rule equals_rsp[OF a])
-  apply(assumption)+
-  done
-
-lemma o_prs:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  and     q3: "Quotient R3 Abs3 Rep3"
-  shows "(Rep1 ---> Abs3) (((Abs2 ---> Rep3) f) o ((Abs1 ---> Rep2) g)) = f o g"
-  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] Quotient_abs_rep[OF q3]
-  unfolding o_def expand_fun_eq by simp
-
-lemma o_rsp:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  and     q3: "Quotient R3 Abs3 Rep3"
-  and     a1: "(R2 ===> R3) f1 f2"
-  and     a2: "(R1 ===> R2) g1 g2"
-  shows "(R1 ===> R3) (f1 o g1) (f2 o g2)"
-  using a1 a2 unfolding o_def expand_fun_eq
-  by (auto)
-
-lemma cond_prs:
-  assumes a: "Quotient R absf repf"
-  shows "absf (if a then repf b else repf c) = (if a then b else c)"
-  using a unfolding Quotient_def by auto
-
-lemma if_prs:
-  assumes q: "Quotient R Abs Rep"
-  shows "Abs (If a (Rep b) (Rep c)) = If a b c"
-  using Quotient_abs_rep[OF q] by auto
-
-(* q not used *)
-lemma if_rsp:
-  assumes q: "Quotient R Abs Rep"
-  and     a: "a1 = a2" "R b1 b2" "R c1 c2"
-  shows "R (If a1 b1 c1) (If a2 b2 c2)"
-  using a by auto
-
-lemma let_prs:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "Abs2 (Let (Rep1 x) ((Abs1 ---> Rep2) f)) = Let x f"
-  using Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2] by auto
-
-lemma let_rsp:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     a1: "(R1 ===> R2) f g"
-  and     a2: "R1 x y"
-  shows "R2 ((Let x f)::'c) ((Let y g)::'c)"
-  using apply_rsp[OF q1 a1] a2 by auto
-
-locale quot_type =
-  fixes R :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
-  and   Abs :: "('a \<Rightarrow> bool) \<Rightarrow> 'b"
-  and   Rep :: "'b \<Rightarrow> ('a \<Rightarrow> bool)"
-  assumes equivp: "equivp R"
-  and     rep_prop: "\<And>y. \<exists>x. Rep y = R x"
-  and     rep_inverse: "\<And>x. Abs (Rep x) = x"
-  and     abs_inverse: "\<And>x. (Rep (Abs (R x))) = (R x)"
-  and     rep_inject: "\<And>x y. (Rep x = Rep y) = (x = y)"
-begin
-
-definition
-  abs::"'a \<Rightarrow> 'b"
-where
-  "abs x \<equiv> Abs (R x)"
-
-definition
-  rep::"'b \<Rightarrow> 'a"
-where
-  "rep a = Eps (Rep a)"
-
-lemma homeier_lem9:
-  shows "R (Eps (R x)) = R x"
-proof -
-  have a: "R x x" using equivp by (simp add: equivp_reflp_symp_transp reflp_def)
-  then have "R x (Eps (R x))" by (rule someI)
-  then show "R (Eps (R x)) = R x"
-    using equivp unfolding equivp_def by simp
-qed
-
-theorem homeier_thm10:
-  shows "abs (rep a) = a"
-  unfolding abs_def rep_def
-proof -
-  from rep_prop
-  obtain x where eq: "Rep a = R x" by auto
-  have "Abs (R (Eps (Rep a))) = Abs (R (Eps (R x)))" using eq by simp
-  also have "\<dots> = Abs (R x)" using homeier_lem9 by simp
-  also have "\<dots> = Abs (Rep a)" using eq by simp
-  also have "\<dots> = a" using rep_inverse by simp
-  finally
-  show "Abs (R (Eps (Rep a))) = a" by simp
-qed
-
-lemma homeier_lem7:
-  shows "(R x = R y) = (Abs (R x) = Abs (R y))" (is "?LHS = ?RHS")
-proof -
-  have "?RHS = (Rep (Abs (R x)) = Rep (Abs (R y)))" by (simp add: rep_inject)
-  also have "\<dots> = ?LHS" by (simp add: abs_inverse)
-  finally show "?LHS = ?RHS" by simp
-qed
-
-theorem homeier_thm11:
-  shows "R r r' = (abs r = abs r')"
-  unfolding abs_def
-  by (simp only: equivp[simplified equivp_def] homeier_lem7)
-
-lemma rep_refl:
-  shows "R (rep a) (rep a)"
-  unfolding rep_def
-  by (simp add: equivp[simplified equivp_def])
-
-
-lemma rep_abs_rsp:
-  shows "R f (rep (abs g)) = R f g"
-  and   "R (rep (abs g)) f = R g f"
-  by (simp_all add: homeier_thm10 homeier_thm11)
-
-lemma Quotient:
-  shows "Quotient R abs rep"
-  unfolding Quotient_def
-  apply(simp add: homeier_thm10)
-  apply(simp add: rep_refl)
-  apply(subst homeier_thm11[symmetric])
-  apply(simp add: equivp[simplified equivp_def])
-  done
-
-end
-
-section {* ML setup *}
-
-text {* Auxiliary data for the quotient package *}
-
-use "quotient_info.ML"
-
-declare [[map "fun" = (fun_map, fun_rel)]]
-
-lemmas [quot_thm] = fun_quotient
-lemmas [quot_respect] = quot_rel_rsp
-lemmas [quot_equiv] = identity_equivp
-
-
-text {* Lemmas about simplifying id's. *}
-lemmas [id_simps] =
-  id_def[symmetric]
-  fun_map_id
-  id_apply
-  id_o
-  o_id
-  eq_comp_r
-
-text {* Translation functions for the lifting process. *}
-use "quotient_term.ML"
-
-
-text {* Definitions of the quotient types. *}
-use "quotient_typ.ML"
-
-
-text {* Definitions for quotient constants. *}
-use "quotient_def.ML"
-
-
-text {*
-  An auxiliary constant for recording some information
-  about the lifted theorem in a tactic.
-*}
-definition
-  "Quot_True x \<equiv> True"
-
-lemma
-  shows QT_all: "Quot_True (All P) \<Longrightarrow> Quot_True P"
-  and   QT_ex:  "Quot_True (Ex P) \<Longrightarrow> Quot_True P"
-  and   QT_ex1: "Quot_True (Ex1 P) \<Longrightarrow> Quot_True P"
-  and   QT_lam: "Quot_True (\<lambda>x. P x) \<Longrightarrow> (\<And>x. Quot_True (P x))"
-  and   QT_ext: "(\<And>x. Quot_True (a x) \<Longrightarrow> f x = g x) \<Longrightarrow> (Quot_True a \<Longrightarrow> f = g)"
-  by (simp_all add: Quot_True_def ext)
-
-lemma QT_imp: "Quot_True a \<equiv> Quot_True b"
-  by (simp add: Quot_True_def)
-
-
-text {* Tactics for proving the lifted theorems *}
-use "quotient_tacs.ML"
-
-section {* Methods / Interface *}
-
-method_setup lifting =
-  {* Attrib.thms >> (fn thms => fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.lift_tac ctxt thms))) *}
-  {* lifts theorems to quotient types *}
-
-method_setup lifting_setup =
-  {* Attrib.thm >> (fn thms => fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.procedure_tac ctxt thms))) *}
-  {* sets up the three goals for the quotient lifting procedure *}
-
-method_setup regularize =
-  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.regularize_tac ctxt))) *}
-  {* proves the regularization goals from the quotient lifting procedure *}
-
-method_setup injection =
-  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.all_injection_tac ctxt))) *}
-  {* proves the rep/abs injection goals from the quotient lifting procedure *}
-
-method_setup cleaning =
-  {* Scan.succeed (fn ctxt => SIMPLE_METHOD (HEADGOAL (Quotient_Tacs.clean_tac ctxt))) *}
-  {* proves the cleaning goals from the quotient lifting procedure *}
-
-attribute_setup quot_lifted =
-  {* Scan.succeed Quotient_Tacs.lifted_attrib *}
-  {* lifts theorems to quotient types *}
-
-no_notation
-  rel_conj (infixr "OOO" 75) and
-  fun_map (infixr "--->" 55) and
-  fun_rel (infixr "===>" 55)
-
-end
-
--- a/Quot/Quotient_List.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,232 +0,0 @@
-(*  Title:      Quotient_List.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-*)
-theory Quotient_List
-imports Quotient Quotient_Syntax List
-begin
-
-section {* Quotient infrastructure for the list type. *}
-
-fun
-  list_rel
-where
-  "list_rel R [] [] = True"
-| "list_rel R (x#xs) [] = False"
-| "list_rel R [] (x#xs) = False"
-| "list_rel R (x#xs) (y#ys) = (R x y \<and> list_rel R xs ys)"
-
-declare [[map list = (map, list_rel)]]
-
-lemma split_list_all:
-  shows "(\<forall>x. P x) \<longleftrightarrow> P [] \<and> (\<forall>x xs. P (x#xs))"
-  apply(auto)
-  apply(case_tac x)
-  apply(simp_all)
-  done
-
-lemma map_id[id_simps]:
-  shows "map id = id"
-  apply(simp add: expand_fun_eq)
-  apply(rule allI)
-  apply(induct_tac x)
-  apply(simp_all)
-  done
-
-
-lemma list_rel_reflp:
-  shows "equivp R \<Longrightarrow> list_rel R xs xs"
-  apply(induct xs)
-  apply(simp_all add: equivp_reflp)
-  done
-
-lemma list_rel_symp:
-  assumes a: "equivp R"
-  shows "list_rel R xs ys \<Longrightarrow> list_rel R ys xs"
-  apply(induct xs ys rule: list_induct2')
-  apply(simp_all)
-  apply(rule equivp_symp[OF a])
-  apply(simp)
-  done
-
-lemma list_rel_transp:
-  assumes a: "equivp R"
-  shows "list_rel R xs1 xs2 \<Longrightarrow> list_rel R xs2 xs3 \<Longrightarrow> list_rel R xs1 xs3"
-  apply(induct xs1 xs2 arbitrary: xs3 rule: list_induct2')
-  apply(simp_all)
-  apply(case_tac xs3)
-  apply(simp_all)
-  apply(rule equivp_transp[OF a])
-  apply(auto)
-  done
-
-lemma list_equivp[quot_equiv]:
-  assumes a: "equivp R"
-  shows "equivp (list_rel R)"
-  apply(rule equivpI)
-  unfolding reflp_def symp_def transp_def
-  apply(subst split_list_all)
-  apply(simp add: equivp_reflp[OF a] list_rel_reflp[OF a])
-  apply(blast intro: list_rel_symp[OF a])
-  apply(blast intro: list_rel_transp[OF a])
-  done
-
-lemma list_rel_rel:
-  assumes q: "Quotient R Abs Rep"
-  shows "list_rel R r s = (list_rel R r r \<and> list_rel R s s \<and> (map Abs r = map Abs s))"
-  apply(induct r s rule: list_induct2')
-  apply(simp_all)
-  using Quotient_rel[OF q]
-  apply(metis)
-  done
-
-lemma list_quotient[quot_thm]:
-  assumes q: "Quotient R Abs Rep"
-  shows "Quotient (list_rel R) (map Abs) (map Rep)"
-  unfolding Quotient_def
-  apply(subst split_list_all)
-  apply(simp add: Quotient_abs_rep[OF q] abs_o_rep[OF q] map_id)
-  apply(rule conjI)
-  apply(rule allI)
-  apply(induct_tac a)
-  apply(simp)
-  apply(simp)
-  apply(simp add: Quotient_rep_reflp[OF q])
-  apply(rule allI)+
-  apply(rule list_rel_rel[OF q])
-  done
-
-
-lemma cons_prs_aux:
-  assumes q: "Quotient R Abs Rep"
-  shows "(map Abs) ((Rep h) # (map Rep t)) = h # t"
-  by (induct t) (simp_all add: Quotient_abs_rep[OF q])
-
-lemma cons_prs[quot_preserve]:
-  assumes q: "Quotient R Abs Rep"
-  shows "(Rep ---> (map Rep) ---> (map Abs)) (op #) = (op #)"
-  by (simp only: expand_fun_eq fun_map_def cons_prs_aux[OF q])
-     (simp)
-
-lemma cons_rsp[quot_respect]:
-  assumes q: "Quotient R Abs Rep"
-  shows "(R ===> list_rel R ===> list_rel R) (op #) (op #)"
-  by (auto)
-
-lemma nil_prs[quot_preserve]:
-  assumes q: "Quotient R Abs Rep"
-  shows "map Abs [] = []"
-  by simp
-
-lemma nil_rsp[quot_respect]:
-  assumes q: "Quotient R Abs Rep"
-  shows "list_rel R [] []"
-  by simp
-
-lemma map_prs_aux:
-  assumes a: "Quotient R1 abs1 rep1"
-  and     b: "Quotient R2 abs2 rep2"
-  shows "(map abs2) (map ((abs1 ---> rep2) f) (map rep1 l)) = map f l"
-  by (induct l)
-     (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
-
-
-lemma map_prs[quot_preserve]:
-  assumes a: "Quotient R1 abs1 rep1"
-  and     b: "Quotient R2 abs2 rep2"
-  shows "((abs1 ---> rep2) ---> (map rep1) ---> (map abs2)) map = map"
-  by (simp only: expand_fun_eq fun_map_def map_prs_aux[OF a b])
-     (simp)
-
-
-lemma map_rsp[quot_respect]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "((R1 ===> R2) ===> (list_rel R1) ===> list_rel R2) map map"
-  apply(simp)
-  apply(rule allI)+
-  apply(rule impI)
-  apply(rule allI)+
-  apply (induct_tac xa ya rule: list_induct2')
-  apply simp_all
-  done
-
-lemma foldr_prs_aux:
-  assumes a: "Quotient R1 abs1 rep1"
-  and     b: "Quotient R2 abs2 rep2"
-  shows "abs2 (foldr ((abs1 ---> abs2 ---> rep2) f) (map rep1 l) (rep2 e)) = foldr f l e"
-  by (induct l) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
-
-lemma foldr_prs[quot_preserve]:
-  assumes a: "Quotient R1 abs1 rep1"
-  and     b: "Quotient R2 abs2 rep2"
-  shows "((abs1 ---> abs2 ---> rep2) ---> (map rep1) ---> rep2 ---> abs2) foldr = foldr"
-  by (simp only: expand_fun_eq fun_map_def foldr_prs_aux[OF a b])
-     (simp)
-
-lemma foldl_prs_aux:
-  assumes a: "Quotient R1 abs1 rep1"
-  and     b: "Quotient R2 abs2 rep2"
-  shows "abs1 (foldl ((abs1 ---> abs2 ---> rep1) f) (rep1 e) (map rep2 l)) = foldl f e l"
-  by (induct l arbitrary:e) (simp_all add: Quotient_abs_rep[OF a] Quotient_abs_rep[OF b])
-
-
-lemma foldl_prs[quot_preserve]:
-  assumes a: "Quotient R1 abs1 rep1"
-  and     b: "Quotient R2 abs2 rep2"
-  shows "((abs1 ---> abs2 ---> rep1) ---> rep1 ---> (map rep2) ---> abs1) foldl = foldl"
-  by (simp only: expand_fun_eq fun_map_def foldl_prs_aux[OF a b])
-     (simp)
-
-lemma list_rel_empty:
-  shows "list_rel R [] b \<Longrightarrow> length b = 0"
-  by (induct b) (simp_all)
-
-lemma list_rel_len:
-  shows "list_rel R a b \<Longrightarrow> length a = length b"
-  apply (induct a arbitrary: b)
-  apply (simp add: list_rel_empty)
-  apply (case_tac b)
-  apply simp_all
-  done
-
-(* induct_tac doesn't accept 'arbitrary', so we manually 'spec' *)
-lemma foldl_rsp[quot_respect]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "((R1 ===> R2 ===> R1) ===> R1 ===> list_rel R2 ===> R1) foldl foldl"
-  apply(auto)
-  apply (subgoal_tac "R1 xa ya \<longrightarrow> list_rel R2 xb yb \<longrightarrow> R1 (foldl x xa xb) (foldl y ya yb)")
-  apply simp
-  apply (rule_tac x="xa" in spec)
-  apply (rule_tac x="ya" in spec)
-  apply (rule_tac xs="xb" and ys="yb" in list_induct2)
-  apply (rule list_rel_len)
-  apply (simp_all)
-  done
-
-lemma foldr_rsp[quot_respect]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "((R1 ===> R2 ===> R2) ===> list_rel R1 ===> R2 ===> R2) foldr foldr"
-  apply auto
-  apply(subgoal_tac "R2 xb yb \<longrightarrow> list_rel R1 xa ya \<longrightarrow> R2 (foldr x xa xb) (foldr y ya yb)")
-  apply simp
-  apply (rule_tac xs="xa" and ys="ya" in list_induct2)
-  apply (rule list_rel_len)
-  apply (simp_all)
-  done
-
-lemma list_rel_eq[id_simps]:
-  shows "(list_rel (op =)) = (op =)"
-  unfolding expand_fun_eq
-  apply(rule allI)+
-  apply(induct_tac x xa rule: list_induct2')
-  apply(simp_all)
-  done
-
-lemma list_rel_refl:
-  assumes a: "\<And>x y. R x y = (R x = R y)"
-  shows "list_rel R x x"
-  by (induct x) (auto simp add: a)
-
-end
--- a/Quot/Quotient_Option.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,80 +0,0 @@
-(*  Title:      Quotient_Option.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-*)
-theory Quotient_Option
-imports Quotient Quotient_Syntax
-begin
-
-section {* Quotient infrastructure for the option type. *}
-
-fun
-  option_rel
-where
-  "option_rel R None None = True"
-| "option_rel R (Some x) None = False"
-| "option_rel R None (Some x) = False"
-| "option_rel R (Some x) (Some y) = R x y"
-
-declare [[map option = (Option.map, option_rel)]]
-
-text {* should probably be in Option.thy *}
-lemma split_option_all:
-  shows "(\<forall>x. P x) \<longleftrightarrow> P None \<and> (\<forall>a. P (Some a))"
-  apply(auto)
-  apply(case_tac x)
-  apply(simp_all)
-  done
-
-lemma option_quotient[quot_thm]:
-  assumes q: "Quotient R Abs Rep"
-  shows "Quotient (option_rel R) (Option.map Abs) (Option.map Rep)"
-  unfolding Quotient_def
-  apply(simp add: split_option_all)
-  apply(simp add: Quotient_abs_rep[OF q] Quotient_rel_rep[OF q])
-  using q
-  unfolding Quotient_def
-  apply(blast)
-  done
-
-lemma option_equivp[quot_equiv]:
-  assumes a: "equivp R"
-  shows "equivp (option_rel R)"
-  apply(rule equivpI)
-  unfolding reflp_def symp_def transp_def
-  apply(simp_all add: split_option_all)
-  apply(blast intro: equivp_reflp[OF a])
-  apply(blast intro: equivp_symp[OF a])
-  apply(blast intro: equivp_transp[OF a])
-  done
-
-lemma option_None_rsp[quot_respect]:
-  assumes q: "Quotient R Abs Rep"
-  shows "option_rel R None None"
-  by simp
-
-lemma option_Some_rsp[quot_respect]:
-  assumes q: "Quotient R Abs Rep"
-  shows "(R ===> option_rel R) Some Some"
-  by simp
-
-lemma option_None_prs[quot_preserve]:
-  assumes q: "Quotient R Abs Rep"
-  shows "Option.map Abs None = None"
-  by simp
-
-lemma option_Some_prs[quot_preserve]:
-  assumes q: "Quotient R Abs Rep"
-  shows "(Rep ---> Option.map Abs) Some = Some"
-  apply(simp add: expand_fun_eq)
-  apply(simp add: Quotient_abs_rep[OF q])
-  done
-
-lemma option_map_id[id_simps]:
-  shows "Option.map id = id"
-  by (simp add: expand_fun_eq split_option_all)
-
-lemma option_rel_eq[id_simps]:
-  shows "option_rel (op =) = (op =)"
-  by (simp add: expand_fun_eq split_option_all)
-
-end
--- a/Quot/Quotient_Product.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,104 +0,0 @@
-(*  Title:      Quotient_Product.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-*)
-theory Quotient_Product
-imports Quotient Quotient_Syntax
-begin
-
-section {* Quotient infrastructure for the product type. *}
-
-fun
-  prod_rel
-where
-  "prod_rel R1 R2 = (\<lambda>(a, b) (c, d). R1 a c \<and> R2 b d)"
-
-declare [[map * = (prod_fun, prod_rel)]]
-
-
-lemma prod_equivp[quot_equiv]:
-  assumes a: "equivp R1"
-  assumes b: "equivp R2"
-  shows "equivp (prod_rel R1 R2)"
-  apply(rule equivpI)
-  unfolding reflp_def symp_def transp_def
-  apply(simp_all add: split_paired_all)
-  apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
-  apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
-  apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
-  done
-
-lemma prod_quotient[quot_thm]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "Quotient (prod_rel R1 R2) (prod_fun Abs1 Abs2) (prod_fun Rep1 Rep2)"
-  unfolding Quotient_def
-  apply(simp add: split_paired_all)
-  apply(simp add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
-  apply(simp add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
-  using q1 q2
-  unfolding Quotient_def
-  apply(blast)
-  done
-
-lemma Pair_rsp[quot_respect]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(R1 ===> R2 ===> prod_rel R1 R2) Pair Pair"
-  by simp
-
-lemma Pair_prs[quot_preserve]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(Rep1 ---> Rep2 ---> (prod_fun Abs1 Abs2)) Pair = Pair"
-  apply(simp add: expand_fun_eq)
-  apply(simp add: Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
-  done
-
-lemma fst_rsp[quot_respect]:
-  assumes "Quotient R1 Abs1 Rep1"
-  assumes "Quotient R2 Abs2 Rep2"
-  shows "(prod_rel R1 R2 ===> R1) fst fst"
-  by simp
-
-lemma fst_prs[quot_preserve]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(prod_fun Rep1 Rep2 ---> Abs1) fst = fst"
-  apply(simp add: expand_fun_eq)
-  apply(simp add: Quotient_abs_rep[OF q1])
-  done
-
-lemma snd_rsp[quot_respect]:
-  assumes "Quotient R1 Abs1 Rep1"
-  assumes "Quotient R2 Abs2 Rep2"
-  shows "(prod_rel R1 R2 ===> R2) snd snd"
-  by simp
-
-lemma snd_prs[quot_preserve]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(prod_fun Rep1 Rep2 ---> Abs2) snd = snd"
-  apply(simp add: expand_fun_eq)
-  apply(simp add: Quotient_abs_rep[OF q2])
-  done
-
-lemma split_rsp[quot_respect]:
-  shows "((R1 ===> R2 ===> (op =)) ===> (prod_rel R1 R2) ===> (op =)) split split"
-  by auto
-
-lemma split_prs[quot_preserve]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  and     q2: "Quotient R2 Abs2 Rep2"
-  shows "(((Abs1 ---> Abs2 ---> id) ---> prod_fun Rep1 Rep2 ---> id) split) = split"
-  by (simp add: expand_fun_eq Quotient_abs_rep[OF q1] Quotient_abs_rep[OF q2])
-
-lemma prod_fun_id[id_simps]:
-  shows "prod_fun id id = id"
-  by (simp add: prod_fun_def)
-
-lemma prod_rel_eq[id_simps]:
-  shows "prod_rel (op =) (op =) = (op =)"
-  by (simp add: expand_fun_eq)
-
-
-end
--- a/Quot/Quotient_Sum.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,96 +0,0 @@
-(*  Title:      Quotient_Sum.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-*)
-theory Quotient_Sum
-imports Quotient Quotient_Syntax
-begin
-
-section {* Quotient infrastructure for the sum type. *}
-
-fun
-  sum_rel
-where
-  "sum_rel R1 R2 (Inl a1) (Inl b1) = R1 a1 b1"
-| "sum_rel R1 R2 (Inl a1) (Inr b2) = False"
-| "sum_rel R1 R2 (Inr a2) (Inl b1) = False"
-| "sum_rel R1 R2 (Inr a2) (Inr b2) = R2 a2 b2"
-
-fun
-  sum_map
-where
-  "sum_map f1 f2 (Inl a) = Inl (f1 a)"
-| "sum_map f1 f2 (Inr a) = Inr (f2 a)"
-
-declare [[map "+" = (sum_map, sum_rel)]]
-
-
-text {* should probably be in Sum_Type.thy *}
-lemma split_sum_all:
-  shows "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (Inl x)) \<and> (\<forall>x. P (Inr x))"
-  apply(auto)
-  apply(case_tac x)
-  apply(simp_all)
-  done
-
-lemma sum_equivp[quot_equiv]:
-  assumes a: "equivp R1"
-  assumes b: "equivp R2"
-  shows "equivp (sum_rel R1 R2)"
-  apply(rule equivpI)
-  unfolding reflp_def symp_def transp_def
-  apply(simp_all add: split_sum_all)
-  apply(blast intro: equivp_reflp[OF a] equivp_reflp[OF b])
-  apply(blast intro: equivp_symp[OF a] equivp_symp[OF b])
-  apply(blast intro: equivp_transp[OF a] equivp_transp[OF b])
-  done
-
-lemma sum_quotient[quot_thm]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "Quotient (sum_rel R1 R2) (sum_map Abs1 Abs2) (sum_map Rep1 Rep2)"
-  unfolding Quotient_def
-  apply(simp add: split_sum_all)
-  apply(simp_all add: Quotient_abs_rep[OF q1] Quotient_rel_rep[OF q1])
-  apply(simp_all add: Quotient_abs_rep[OF q2] Quotient_rel_rep[OF q2])
-  using q1 q2
-  unfolding Quotient_def
-  apply(blast)+
-  done
-
-lemma sum_Inl_rsp[quot_respect]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(R1 ===> sum_rel R1 R2) Inl Inl"
-  by simp
-
-lemma sum_Inr_rsp[quot_respect]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(R2 ===> sum_rel R1 R2) Inr Inr"
-  by simp
-
-lemma sum_Inl_prs[quot_preserve]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(Rep1 ---> sum_map Abs1 Abs2) Inl = Inl"
-  apply(simp add: expand_fun_eq)
-  apply(simp add: Quotient_abs_rep[OF q1])
-  done
-
-lemma sum_Inr_prs[quot_preserve]:
-  assumes q1: "Quotient R1 Abs1 Rep1"
-  assumes q2: "Quotient R2 Abs2 Rep2"
-  shows "(Rep2 ---> sum_map Abs1 Abs2) Inr = Inr"
-  apply(simp add: expand_fun_eq)
-  apply(simp add: Quotient_abs_rep[OF q2])
-  done
-
-lemma sum_map_id[id_simps]:
-  shows "sum_map id id = id"
-  by (simp add: expand_fun_eq split_sum_all)
-
-lemma sum_rel_eq[id_simps]:
-  shows "sum_rel (op =) (op =) = (op =)"
-  by (simp add: expand_fun_eq split_sum_all)
-
-end
--- a/Quot/Quotient_Syntax.thy	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,18 +0,0 @@
-(*  Title:      Quotient_Syntax.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-*)
-
-header {* Pretty syntax for Quotient operations *}
-
-(*<*)
-theory Quotient_Syntax
-imports Quotient
-begin
-
-notation
-  rel_conj (infixr "OOO" 75) and
-  fun_map (infixr "--->" 55) and
-  fun_rel (infixr "===>" 55)
-
-end
-(*>*)
--- a/Quot/ROOT.ML	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,15 +0,0 @@
-quick_and_dirty := true;
-
-no_document use_thys
-   ["Quotient",
-    "Examples/AbsRepTest",
-    "Examples/FSet",
-    "Examples/FSet2",
-    "Examples/FSet3",
-    "Examples/IntEx",
-    "Examples/IntEx2",
-    "Examples/LFex",
-    "Examples/LamEx",
-    "Examples/LarryDatatype",
-    "Examples/LarryInt",
-    "Examples/Terms"];
--- a/Quot/quotient_def.ML	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,110 +0,0 @@
-(*  Title:      quotient_def.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-
-    Definitions for constants on quotient types.
-
-*)
-
-signature QUOTIENT_DEF =
-sig
-  val quotient_def: (binding option * mixfix) * (Attrib.binding * (term * term)) ->
-    local_theory -> (term * thm) * local_theory
-
-  val quotdef_cmd: (binding option * mixfix) * (Attrib.binding * (string * string)) ->
-    local_theory -> (term * thm) * local_theory
-
-  val quotient_lift_const: string * term -> local_theory -> (term * thm) * local_theory
-end;
-
-structure Quotient_Def: QUOTIENT_DEF =
-struct
-
-open Quotient_Info;
-open Quotient_Term;
-
-(** Interface and Syntax Setup **)
-
-(* The ML-interface for a quotient definition takes
-   as argument:
-
-    - an optional binding and mixfix annotation
-    - attributes
-    - the new constant as term
-    - the rhs of the definition as term
-
-   It returns the defined constant and its definition
-   theorem; stores the data in the qconsts data slot.
-
-   Restriction: At the moment the right-hand side of the
-   definition must be a constant. Similarly the left-hand 
-   side must be a constant.
-*)
-fun error_msg bind str = 
-let 
-  val name = Binding.name_of bind
-  val pos = Position.str_of (Binding.pos_of bind)
-in
-  error ("Head of quotient_definition " ^ 
-    (quote str) ^ " differs from declaration " ^ name ^ pos)
-end
-
-fun quotient_def ((optbind, mx), (attr, (lhs, rhs))) lthy =
-let
-  val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined."
-  val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction"
-  
-  fun sanity_test NONE _ = true
-    | sanity_test (SOME bind) str =
-        if Name.of_binding bind = str then true
-        else error_msg bind str
-
-  val _ = sanity_test optbind lhs_str
-
-  val qconst_bname = Binding.name lhs_str
-  val absrep_trm = absrep_fun AbsF lthy (fastype_of rhs, lhs_ty) $ rhs
-  val prop = Logic.mk_equals (lhs, Syntax.check_term lthy absrep_trm)
-  val (_, prop') = LocalDefs.cert_def lthy prop
-  val (_, newrhs) = Primitive_Defs.abs_def prop'
-
-  val ((trm, (_ , thm)), lthy') = Local_Theory.define ((qconst_bname, mx), (attr, newrhs)) lthy
-
-  (* data storage *)
-  fun qcinfo phi = transform_qconsts phi {qconst = trm, rconst = rhs, def = thm}
-  fun trans_name phi = (fst o dest_Const o #qconst) (qcinfo phi)
-  val lthy'' = Local_Theory.declaration true
-                 (fn phi => qconsts_update_gen (trans_name phi) (qcinfo phi)) lthy'
-in
-  ((trm, thm), lthy'')
-end
-
-fun quotdef_cmd (decl, (attr, (lhs_str, rhs_str))) lthy =
-let
-  val lhs = Syntax.read_term lthy lhs_str
-  val rhs = Syntax.read_term lthy rhs_str
-  val lthy' = Variable.declare_term lhs lthy
-  val lthy'' = Variable.declare_term rhs lthy'
-in
-  quotient_def (decl, (attr, (lhs, rhs))) lthy''
-end
-
-fun quotient_lift_const (b, t) ctxt =
-  quotient_def ((NONE, NoSyn), (Attrib.empty_binding,
-    (Quotient_Term.quotient_lift_const (b, t) ctxt, t))) ctxt
-
-local
-  structure P = OuterParse;
-in
-
-val quotdef_decl = (P.binding >> SOME) -- P.opt_mixfix' --| P.$$$ "where"
-
-val quotdef_parser =
-  Scan.optional quotdef_decl (NONE, NoSyn) -- 
-    P.!!! (SpecParse.opt_thm_name ":" -- (P.term --| P.$$$ "is" -- P.term))
-end
-
-val _ =
-  OuterSyntax.local_theory "quotient_definition"
-    "definition for constants over the quotient type"
-      OuterKeyword.thy_decl (quotdef_parser >> (snd oo quotdef_cmd))
-
-end; (* structure *)
--- a/Quot/quotient_info.ML	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,293 +0,0 @@
-(*  Title:      quotient_info.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-
-    Data slots for the quotient package.
-
-*)
-
-
-signature QUOTIENT_INFO =
-sig
-  exception NotFound
-
-  type maps_info = {mapfun: string, relmap: string}
-  val maps_defined: theory -> string -> bool
-  val maps_lookup: theory -> string -> maps_info     (* raises NotFound *)
-  val maps_update_thy: string -> maps_info -> theory -> theory
-  val maps_update: string -> maps_info -> Proof.context -> Proof.context
-  val print_mapsinfo: Proof.context -> unit
-
-  type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
-  val transform_quotdata: morphism -> quotdata_info -> quotdata_info
-  val quotdata_lookup_raw: theory -> string -> quotdata_info option
-  val quotdata_lookup: theory -> string -> quotdata_info     (* raises NotFound *)
-  val quotdata_update_thy: string -> quotdata_info -> theory -> theory
-  val quotdata_update_gen: string -> quotdata_info -> Context.generic -> Context.generic
-  val quotdata_dest: Proof.context -> quotdata_info list
-  val print_quotinfo: Proof.context -> unit
-
-  type qconsts_info = {qconst: term, rconst: term, def: thm}
-  val transform_qconsts: morphism -> qconsts_info -> qconsts_info
-  val qconsts_lookup: theory -> term -> qconsts_info     (* raises NotFound *)
-  val qconsts_update_thy: string -> qconsts_info -> theory -> theory
-  val qconsts_update_gen: string -> qconsts_info -> Context.generic -> Context.generic
-  val qconsts_dest: Proof.context -> qconsts_info list
-  val print_qconstinfo: Proof.context -> unit
-
-  val equiv_rules_get: Proof.context -> thm list
-  val equiv_rules_add: attribute
-  val rsp_rules_get: Proof.context -> thm list
-  val rsp_rules_add: attribute
-  val prs_rules_get: Proof.context -> thm list
-  val prs_rules_add: attribute
-  val id_simps_get: Proof.context -> thm list
-  val quotient_rules_get: Proof.context -> thm list
-  val quotient_rules_add: attribute
-end;
-
-
-structure Quotient_Info: QUOTIENT_INFO =
-struct
-
-exception NotFound
-
-
-(** data containers **)
-
-(* info about map- and rel-functions for a type *)
-type maps_info = {mapfun: string, relmap: string}
-
-structure MapsData = Theory_Data
-  (type T = maps_info Symtab.table
-   val empty = Symtab.empty
-   val extend = I
-   val merge = Symtab.merge (K true))
-
-fun maps_defined thy s =
-  Symtab.defined (MapsData.get thy) s
-
-fun maps_lookup thy s =
-  case (Symtab.lookup (MapsData.get thy) s) of
-    SOME map_fun => map_fun
-  | NONE => raise NotFound
-
-fun maps_update_thy k minfo = MapsData.map (Symtab.update (k, minfo))
-fun maps_update k minfo = ProofContext.theory (maps_update_thy k minfo)
-
-fun maps_attribute_aux s minfo = Thm.declaration_attribute
-  (fn _ => Context.mapping (maps_update_thy s minfo) (maps_update s minfo))
-
-(* attribute to be used in declare statements *)
-fun maps_attribute (ctxt, (tystr, (mapstr, relstr))) =
-let
-  val thy = ProofContext.theory_of ctxt
-  val tyname = Sign.intern_type thy tystr
-  val mapname = Sign.intern_const thy mapstr
-  val relname = Sign.intern_const thy relstr
-
-  fun sanity_check s = (Const (s, dummyT) |> Syntax.check_term ctxt; ())
-  val _ = List.app sanity_check [mapname, relname]
-in
-  maps_attribute_aux tyname {mapfun = mapname, relmap = relname}
-end
-
-val maps_attr_parser =
-  Args.context -- Scan.lift
-    ((Args.name --| OuterParse.$$$ "=") --
-      (OuterParse.$$$ "(" |-- Args.name --| OuterParse.$$$ "," --
-        Args.name --| OuterParse.$$$ ")"))
-
-val _ = Context.>> (Context.map_theory
-  (Attrib.setup @{binding "map"} (maps_attr_parser >> maps_attribute)
-    "declaration of map information"))
-
-fun print_mapsinfo ctxt =
-let
-  fun prt_map (ty_name, {mapfun, relmap}) =
-    Pretty.block (Library.separate (Pretty.brk 2)
-      (map Pretty.str
-        ["type:", ty_name,
-        "map:", mapfun,
-        "relation map:", relmap]))
-in
-  MapsData.get (ProofContext.theory_of ctxt)
-  |> Symtab.dest
-  |> map (prt_map)
-  |> Pretty.big_list "maps for type constructors:"
-  |> Pretty.writeln
-end
-
-
-(* info about quotient types *)
-type quotdata_info = {qtyp: typ, rtyp: typ, equiv_rel: term, equiv_thm: thm}
-
-structure QuotData = Theory_Data
-  (type T = quotdata_info Symtab.table
-   val empty = Symtab.empty
-   val extend = I
-   val merge = Symtab.merge (K true))
-
-fun transform_quotdata phi {qtyp, rtyp, equiv_rel, equiv_thm} =
-  {qtyp = Morphism.typ phi qtyp,
-   rtyp = Morphism.typ phi rtyp,
-   equiv_rel = Morphism.term phi equiv_rel,
-   equiv_thm = Morphism.thm phi equiv_thm}
-
-fun quotdata_lookup_raw thy str = Symtab.lookup (QuotData.get thy) str
-
-fun quotdata_lookup thy str =
-  case Symtab.lookup (QuotData.get thy) str of
-    SOME qinfo => qinfo
-  | NONE => raise NotFound
-
-fun quotdata_update_thy str qinfo = QuotData.map (Symtab.update (str, qinfo))
-fun quotdata_update_gen str qinfo = Context.mapping (quotdata_update_thy str qinfo) I
-
-fun quotdata_dest lthy =
-  map snd (Symtab.dest (QuotData.get (ProofContext.theory_of lthy)))
-
-fun print_quotinfo ctxt =
-let
-  fun prt_quot {qtyp, rtyp, equiv_rel, equiv_thm} =
-    Pretty.block (Library.separate (Pretty.brk 2)
-     [Pretty.str "quotient type:",
-      Syntax.pretty_typ ctxt qtyp,
-      Pretty.str "raw type:",
-      Syntax.pretty_typ ctxt rtyp,
-      Pretty.str "relation:",
-      Syntax.pretty_term ctxt equiv_rel,
-      Pretty.str "equiv. thm:",
-      Syntax.pretty_term ctxt (prop_of equiv_thm)])
-in
-  QuotData.get (ProofContext.theory_of ctxt)
-  |> Symtab.dest
-  |> map (prt_quot o snd)
-  |> Pretty.big_list "quotients:"
-  |> Pretty.writeln
-end
-
-
-(* info about quotient constants *)
-type qconsts_info = {qconst: term, rconst: term, def: thm}
-
-fun qconsts_info_eq (x : qconsts_info, y : qconsts_info) = #qconst x = #qconst y
-
-(* We need to be able to lookup instances of lifted constants,
-   for example given "nat fset" we need to find "'a fset";
-   but overloaded constants share the same name *)
-structure QConstsData = Theory_Data
-  (type T = (qconsts_info list) Symtab.table
-   val empty = Symtab.empty
-   val extend = I
-   val merge = Symtab.merge_list qconsts_info_eq)
-
-fun transform_qconsts phi {qconst, rconst, def} =
-  {qconst = Morphism.term phi qconst,
-   rconst = Morphism.term phi rconst,
-   def = Morphism.thm phi def}
-
-fun qconsts_update_thy name qcinfo = QConstsData.map (Symtab.cons_list (name, qcinfo))
-fun qconsts_update_gen name qcinfo = Context.mapping (qconsts_update_thy name qcinfo) I
-
-fun qconsts_dest lthy =
-  flat (map snd (Symtab.dest (QConstsData.get (ProofContext.theory_of lthy))))
-
-fun qconsts_lookup thy t =
-  let
-    val (name, qty) = dest_Const t
-    fun matches x =
-      let
-        val (name', qty') = dest_Const (#qconst x);
-      in
-        name = name' andalso Sign.typ_instance thy (qty, qty')
-      end
-  in
-    case Symtab.lookup (QConstsData.get thy) name of
-      NONE => raise NotFound
-    | SOME l =>
-      (case (find_first matches l) of
-        SOME x => x
-      | NONE => raise NotFound)
-  end
-
-fun print_qconstinfo ctxt =
-let
-  fun prt_qconst {qconst, rconst, def} =
-    Pretty.block (separate (Pretty.brk 1)
-     [Syntax.pretty_term ctxt qconst,
-      Pretty.str ":=",
-      Syntax.pretty_term ctxt rconst,
-      Pretty.str "as",
-      Syntax.pretty_term ctxt (prop_of def)])
-in
-  QConstsData.get (ProofContext.theory_of ctxt)
-  |> Symtab.dest
-  |> map snd
-  |> flat
-  |> map prt_qconst
-  |> Pretty.big_list "quotient constants:"
-  |> Pretty.writeln
-end
-
-(* equivalence relation theorems *)
-structure EquivRules = Named_Thms
-  (val name = "quot_equiv"
-   val description = "Equivalence relation theorems.")
-
-val equiv_rules_get = EquivRules.get
-val equiv_rules_add = EquivRules.add
-
-(* respectfulness theorems *)
-structure RspRules = Named_Thms
-  (val name = "quot_respect"
-   val description = "Respectfulness theorems.")
-
-val rsp_rules_get = RspRules.get
-val rsp_rules_add = RspRules.add
-
-(* preservation theorems *)
-structure PrsRules = Named_Thms
-  (val name = "quot_preserve"
-   val description = "Preservation theorems.")
-
-val prs_rules_get = PrsRules.get
-val prs_rules_add = PrsRules.add
-
-(* id simplification theorems *)
-structure IdSimps = Named_Thms
-  (val name = "id_simps"
-   val description = "Identity simp rules for maps.")
-
-val id_simps_get = IdSimps.get
-
-(* quotient theorems *)
-structure QuotientRules = Named_Thms
-  (val name = "quot_thm"
-   val description = "Quotient theorems.")
-
-val quotient_rules_get = QuotientRules.get
-val quotient_rules_add = QuotientRules.add
-
-(* setup of the theorem lists *)
-
-val _ = Context.>> (Context.map_theory
-  (EquivRules.setup #>
-   RspRules.setup #>
-   PrsRules.setup #>
-   IdSimps.setup #>
-   QuotientRules.setup))
-
-(* setup of the printing commands *)
-
-fun improper_command (pp_fn, cmd_name, descr_str) =
-  OuterSyntax.improper_command cmd_name descr_str
-    OuterKeyword.diag (Scan.succeed (Toplevel.keep (pp_fn o Toplevel.context_of)))
-
-val _ = map improper_command
-  [(print_mapsinfo, "print_maps", "prints out all map functions"),
-   (print_quotinfo, "print_quotients", "prints out all quotients"),
-   (print_qconstinfo, "print_quotconsts", "prints out all quotient constants")]
-
-
-end; (* structure *)
-
--- a/Quot/quotient_tacs.ML	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,665 +0,0 @@
-(*  Title:      quotient_tacs.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-
-    Tactics for solving goal arising from lifting
-    theorems to quotient types.
-*)
-
-signature QUOTIENT_TACS =
-sig
-  val regularize_tac: Proof.context -> int -> tactic
-  val injection_tac: Proof.context -> int -> tactic
-  val all_injection_tac: Proof.context -> int -> tactic
-  val clean_tac: Proof.context -> int -> tactic
-  val procedure_tac: Proof.context -> thm -> int -> tactic
-  val lift_tac: Proof.context -> thm list -> int -> tactic
-  val quotient_tac: Proof.context -> int -> tactic
-  val quot_true_tac: Proof.context -> (term -> term) -> int -> tactic
-  val lifted_attrib: attribute
-end;
-
-structure Quotient_Tacs: QUOTIENT_TACS =
-struct
-
-open Quotient_Info;
-open Quotient_Term;
-
-
-(** various helper fuctions **)
-
-(* Since HOL_basic_ss is too "big" for us, we *)
-(* need to set up our own minimal simpset.    *)
-fun mk_minimal_ss ctxt =
-  Simplifier.context ctxt empty_ss
-    setsubgoaler asm_simp_tac
-    setmksimps (mksimps [])
-
-(* composition of two theorems, used in maps *)
-fun OF1 thm1 thm2 = thm2 RS thm1
-
-(* prints a warning, if the subgoal is not solved *)
-fun WARN (tac, msg) i st =
- case Seq.pull (SOLVED' tac i st) of
-     NONE    => (warning msg; Seq.single st)
-   | seqcell => Seq.make (fn () => seqcell)
-
-fun RANGE_WARN tacs = RANGE (map WARN tacs)
-
-fun atomize_thm thm =
-let
-  val thm' = Thm.freezeT (forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? *)
-  val thm'' = ObjectLogic.atomize (cprop_of thm')
-in
-  @{thm equal_elim_rule1} OF [thm'', thm']
-end
-
-
-
-(*** Regularize Tactic ***)
-
-(** solvers for equivp and quotient assumptions **)
-
-fun equiv_tac ctxt =
-  REPEAT_ALL_NEW (resolve_tac (equiv_rules_get ctxt))
-
-fun equiv_solver_tac ss = equiv_tac (Simplifier.the_context ss)
-val equiv_solver = Simplifier.mk_solver' "Equivalence goal solver" equiv_solver_tac
-
-fun quotient_tac ctxt =
-  (REPEAT_ALL_NEW (FIRST'
-    [rtac @{thm identity_quotient},
-     resolve_tac (quotient_rules_get ctxt)]))
-
-fun quotient_solver_tac ss = quotient_tac (Simplifier.the_context ss)
-val quotient_solver =
-  Simplifier.mk_solver' "Quotient goal solver" quotient_solver_tac
-
-fun solve_quotient_assm ctxt thm =
-  case Seq.pull (quotient_tac ctxt 1 thm) of
-    SOME (t, _) => t
-  | _ => error "Solve_quotient_assm failed. Possibly a quotient theorem is missing."
-
-
-fun prep_trm thy (x, (T, t)) =
-  (cterm_of thy (Var (x, T)), cterm_of thy t)
-
-fun prep_ty thy (x, (S, ty)) =
-  (ctyp_of thy (TVar (x, S)), ctyp_of thy ty)
-
-fun get_match_inst thy pat trm =
-let
-  val univ = Unify.matchers thy [(pat, trm)]
-  val SOME (env, _) = Seq.pull univ             (* raises BIND, if no unifier *)
-  val tenv = Vartab.dest (Envir.term_env env)
-  val tyenv = Vartab.dest (Envir.type_env env)
-in
-  (map (prep_ty thy) tyenv, map (prep_trm thy) tenv)
-end
-
-(* Calculates the instantiations for the lemmas:
-
-      ball_reg_eqv_range and bex_reg_eqv_range
-
-   Since the left-hand-side contains a non-pattern '?P (f ?x)'
-   we rely on unification/instantiation to check whether the
-   theorem applies and return NONE if it doesn't.
-*)
-fun calculate_inst ctxt ball_bex_thm redex R1 R2 =
-let
-  val thy = ProofContext.theory_of ctxt
-  fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm))
-  val ty_inst = map (SOME o ctyp_of thy) [domain_type (fastype_of R2)]
-  val trm_inst = map (SOME o cterm_of thy) [R2, R1]
-in
-  case try (Drule.instantiate' ty_inst trm_inst) ball_bex_thm of
-    NONE => NONE
-  | SOME thm' =>
-      (case try (get_match_inst thy (get_lhs thm')) redex of
-        NONE => NONE
-      | SOME inst2 => try (Drule.instantiate inst2) thm')
-end
-
-fun ball_bex_range_simproc ss redex =
-let
-  val ctxt = Simplifier.the_context ss
-in
-  case redex of
-    (Const (@{const_name "Ball"}, _) $ (Const (@{const_name "Respects"}, _) $
-      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
-        calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2
-
-  | (Const (@{const_name "Bex"}, _) $ (Const (@{const_name "Respects"}, _) $
-      (Const (@{const_name "fun_rel"}, _) $ R1 $ R2)) $ _) =>
-        calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2
-
-  | _ => NONE
-end
-
-(* Regularize works as follows:
-
-  0. preliminary simplification step according to
-     ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range
-
-  1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left)
-
-  2. monos
-
-  3. commutation rules for ball and bex (ball_all_comm bex_ex_comm)
-
-  4. then rel-equalities, which need to be instantiated with 'eq_imp_rel'
-     to avoid loops
-
-  5. then simplification like 0
-
-  finally jump back to 1
-*)
-
-fun regularize_tac ctxt =
-let
-  val thy = ProofContext.theory_of ctxt
-  val ball_pat = @{term "Ball (Respects (R1 ===> R2)) P"}
-  val bex_pat  = @{term "Bex (Respects (R1 ===> R2)) P"}
-  val simproc = Simplifier.simproc_i thy "" [ball_pat, bex_pat] (K (ball_bex_range_simproc))
-  val simpset = (mk_minimal_ss ctxt)
-                       addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp}
-                       addsimprocs [simproc]
-                       addSolver equiv_solver addSolver quotient_solver
-  val eq_imp_rel = @{lemma "equivp R ==> a = b --> R a b" by (simp add: equivp_reflp)}
-  val eq_eqvs = map (OF1 eq_imp_rel) (equiv_rules_get ctxt)
-in
-  simp_tac simpset THEN'
-  REPEAT_ALL_NEW (CHANGED o FIRST'
-    [resolve_tac @{thms ball_reg_right bex_reg_left bex1_bexeq_reg},
-     resolve_tac (Inductive.get_monos ctxt),
-     resolve_tac @{thms ball_all_comm bex_ex_comm},
-     resolve_tac eq_eqvs,
-     simp_tac simpset])
-end
-
-
-
-(*** Injection Tactic ***)
-
-(* Looks for Quot_True assumptions, and in case its parameter
-   is an application, it returns the function and the argument.
-*)
-fun find_qt_asm asms =
-let
-  fun find_fun trm =
-    case trm of
-      (Const(@{const_name Trueprop}, _) $ (Const (@{const_name Quot_True}, _) $ _)) => true
-    | _ => false
-in
- case find_first find_fun asms of
-   SOME (_ $ (_ $ (f $ a))) => SOME (f, a)
- | _ => NONE
-end
-
-fun quot_true_simple_conv ctxt fnctn ctrm =
-  case (term_of ctrm) of
-    (Const (@{const_name Quot_True}, _) $ x) =>
-    let
-      val fx = fnctn x;
-      val thy = ProofContext.theory_of ctxt;
-      val cx = cterm_of thy x;
-      val cfx = cterm_of thy fx;
-      val cxt = ctyp_of thy (fastype_of x);
-      val cfxt = ctyp_of thy (fastype_of fx);
-      val thm = Drule.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp}
-    in
-      Conv.rewr_conv thm ctrm
-    end
-
-fun quot_true_conv ctxt fnctn ctrm =
-  case (term_of ctrm) of
-    (Const (@{const_name Quot_True}, _) $ _) =>
-      quot_true_simple_conv ctxt fnctn ctrm
-  | _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm
-  | Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm
-  | _ => Conv.all_conv ctrm
-
-fun quot_true_tac ctxt fnctn =
-   CONVERSION
-    ((Conv.params_conv ~1 (fn ctxt =>
-       (Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt)
-
-fun dest_comb (f $ a) = (f, a)
-fun dest_bcomb ((_ $ l) $ r) = (l, r)
-
-fun unlam t =
-  case t of
-    (Abs a) => snd (Term.dest_abs a)
-  | _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0)))
-
-fun dest_fun_type (Type("fun", [T, S])) = (T, S)
-  | dest_fun_type _ = error "dest_fun_type"
-
-val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl
-
-(* We apply apply_rsp only in case if the type needs lifting.
-   This is the case if the type of the data in the Quot_True
-   assumption is different from the corresponding type in the goal.
-*)
-val apply_rsp_tac =
-  Subgoal.FOCUS (fn {concl, asms, context,...} =>
-  let
-    val bare_concl = HOLogic.dest_Trueprop (term_of concl)
-    val qt_asm = find_qt_asm (map term_of asms)
-  in
-    case (bare_concl, qt_asm) of
-      (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) =>
-         if fastype_of qt_fun = fastype_of f
-         then no_tac
-         else
-           let
-             val ty_x = fastype_of x
-             val ty_b = fastype_of qt_arg
-             val ty_f = range_type (fastype_of f)
-             val thy = ProofContext.theory_of context
-             val ty_inst = map (SOME o (ctyp_of thy)) [ty_x, ty_b, ty_f]
-             val t_inst = map (SOME o (cterm_of thy)) [R2, f, g, x, y];
-             val inst_thm = Drule.instantiate' ty_inst
-               ([NONE, NONE, NONE] @ t_inst) @{thm apply_rsp}
-           in
-             (rtac inst_thm THEN' quotient_tac context) 1
-           end
-    | _ => no_tac
-  end)
-
-(* Instantiates and applies 'equals_rsp'. Since the theorem is
-   complex we rely on instantiation to tell us if it applies
-*)
-fun equals_rsp_tac R ctxt =
-let
-  val thy = ProofContext.theory_of ctxt
-in
-  case try (cterm_of thy) R of (* There can be loose bounds in R *)
-    SOME ctm =>
-      let
-        val ty = domain_type (fastype_of R)
-      in
-        case try (Drule.instantiate' [SOME (ctyp_of thy ty)]
-          [SOME (cterm_of thy R)]) @{thm equals_rsp} of
-          SOME thm => rtac thm THEN' quotient_tac ctxt
-        | NONE => K no_tac
-      end
-  | _ => K no_tac
-end
-
-fun rep_abs_rsp_tac ctxt =
-  SUBGOAL (fn (goal, i) =>
-    case (try bare_concl goal) of
-      SOME (rel $ _ $ (rep $ (abs $ _))) =>
-        let
-          val thy = ProofContext.theory_of ctxt;
-          val (ty_a, ty_b) = dest_fun_type (fastype_of abs);
-          val ty_inst = map (SOME o (ctyp_of thy)) [ty_a, ty_b];
-        in
-          case try (map (SOME o (cterm_of thy))) [rel, abs, rep] of
-            SOME t_inst =>
-              (case try (Drule.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of
-                SOME inst_thm => (rtac inst_thm THEN' quotient_tac ctxt) i
-              | NONE => no_tac)
-          | NONE => no_tac
-        end
-    | _ => no_tac)
-
-
-
-(* Injection means to prove that the regularised theorem implies
-   the abs/rep injected one.
-
-   The deterministic part:
-    - remove lambdas from both sides
-    - prove Ball/Bex/Babs equalities using ball_rsp, bex_rsp, babs_rsp
-    - prove Ball/Bex relations unfolding fun_rel_id
-    - reflexivity of equality
-    - prove equality of relations using equals_rsp
-    - use user-supplied RSP theorems
-    - solve 'relation of relations' goals using quot_rel_rsp
-    - remove rep_abs from the right side
-      (Lambdas under respects may have left us some assumptions)
-
-   Then in order:
-    - split applications of lifted type (apply_rsp)
-    - split applications of non-lifted type (cong_tac)
-    - apply extentionality
-    - assumption
-    - reflexivity of the relation
-*)
-fun injection_match_tac ctxt = SUBGOAL (fn (goal, i) =>
-(case (bare_concl goal) of
-    (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *)
-  (Const (@{const_name fun_rel}, _) $ _ $ _) $ (Abs _) $ (Abs _)
-      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
-
-    (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *)
-| (Const (@{const_name "op ="},_) $
-    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
-    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
-      => rtac @{thm ball_rsp} THEN' dtac @{thm QT_all}
-
-    (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *)
-| (Const (@{const_name fun_rel}, _) $ _ $ _) $
-    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
-    (Const(@{const_name Ball},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
-      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
-
-    (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *)
-| Const (@{const_name "op ="},_) $
-    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
-    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
-      => rtac @{thm bex_rsp} THEN' dtac @{thm QT_ex}
-
-    (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *)
-| (Const (@{const_name fun_rel}, _) $ _ $ _) $
-    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
-    (Const(@{const_name Bex},_) $ (Const (@{const_name Respects}, _) $ _) $ _)
-      => rtac @{thm fun_rel_id} THEN' quot_true_tac ctxt unlam
-
-| (Const (@{const_name fun_rel}, _) $ _ $ _) $
-    (Const(@{const_name Bex1_rel},_) $ _) $ (Const(@{const_name Bex1_rel},_) $ _)
-      => rtac @{thm bex1_rel_rsp} THEN' quotient_tac ctxt
-
-| (_ $
-    (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _) $
-    (Const(@{const_name Babs},_) $ (Const (@{const_name Respects}, _) $ _) $ _))
-      => rtac @{thm babs_rsp} THEN' RANGE [quotient_tac ctxt]
-
-| Const (@{const_name "op ="},_) $ (R $ _ $ _) $ (_ $ _ $ _) =>
-   (rtac @{thm refl} ORELSE'
-    (equals_rsp_tac R ctxt THEN' RANGE [
-       quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)]))
-
-    (* reflexivity of operators arising from Cong_tac *)
-| Const (@{const_name "op ="},_) $ _ $ _ => rtac @{thm refl}
-
-   (* respectfulness of constants; in particular of a simple relation *)
-| _ $ (Const _) $ (Const _)  (* fun_rel, list_rel, etc but not equality *)
-    => resolve_tac (rsp_rules_get ctxt) THEN_ALL_NEW quotient_tac ctxt
-
-    (* R (...) (Rep (Abs ...)) ----> R (...) (...) *)
-    (* observe fun_map *)
-| _ $ _ $ _
-    => (rtac @{thm quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt)
-       ORELSE' rep_abs_rsp_tac ctxt
-
-| _ => K no_tac
-) i)
-
-fun injection_step_tac ctxt rel_refl =
- FIRST' [
-    injection_match_tac ctxt,
-
-    (* R (t $ ...) (t' $ ...) ----> apply_rsp   provided type of t needs lifting *)
-    apply_rsp_tac ctxt THEN'
-                 RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)],
-
-    (* (op =) (t $ ...) (t' $ ...) ----> Cong   provided type of t does not need lifting *)
-    (* merge with previous tactic *)
-    Cong_Tac.cong_tac @{thm cong} THEN'
-                 RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)],
-
-    (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *)
-    rtac @{thm ext} THEN' quot_true_tac ctxt unlam,
-
-    (* resolving with R x y assumptions *)
-    atac,
-
-    (* reflexivity of the basic relations *)
-    (* R ... ... *)
-    resolve_tac rel_refl]
-
-fun injection_tac ctxt =
-let
-  val rel_refl = map (OF1 @{thm equivp_reflp}) (equiv_rules_get ctxt)
-in
-  injection_step_tac ctxt rel_refl
-end
-
-fun all_injection_tac ctxt =
-  REPEAT_ALL_NEW (injection_tac ctxt)
-
-
-
-(*** Cleaning of the Theorem ***)
-
-(* expands all fun_maps, except in front of the (bound) variables listed in xs *)
-fun fun_map_simple_conv xs ctrm =
-  case (term_of ctrm) of
-    ((Const (@{const_name "fun_map"}, _) $ _ $ _) $ h $ _) =>
-        if member (op=) xs h
-        then Conv.all_conv ctrm
-        else Conv.rewr_conv @{thm fun_map_def[THEN eq_reflection]} ctrm
-  | _ => Conv.all_conv ctrm
-
-fun fun_map_conv xs ctxt ctrm =
-  case (term_of ctrm) of
-      _ $ _ => (Conv.comb_conv (fun_map_conv xs ctxt) then_conv
-                fun_map_simple_conv xs) ctrm
-    | Abs _ => Conv.abs_conv (fn (x, ctxt) => fun_map_conv ((term_of x)::xs) ctxt) ctxt ctrm
-    | _ => Conv.all_conv ctrm
-
-fun fun_map_tac ctxt = CONVERSION (fun_map_conv [] ctxt)
-
-(* custom matching functions *)
-fun mk_abs u i t =
-  if incr_boundvars i u aconv t then Bound i else
-  case t of
-    t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2
-  | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t')
-  | Bound j => if i = j then error "make_inst" else t
-  | _ => t
-
-fun make_inst lhs t =
-let
-  val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs;
-  val _ $ (Abs (_, _, (_ $ g))) = t;
-in
-  (f, Abs ("x", T, mk_abs u 0 g))
-end
-
-fun make_inst_id lhs t =
-let
-  val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs;
-  val _ $ (Abs (_, _, g)) = t;
-in
-  (f, Abs ("x", T, mk_abs u 0 g))
-end
-
-(* Simplifies a redex using the 'lambda_prs' theorem.
-   First instantiates the types and known subterms.
-   Then solves the quotient assumptions to get Rep2 and Abs1
-   Finally instantiates the function f using make_inst
-   If Rep2 is an identity then the pattern is simpler and
-   make_inst_id is used
-*)
-fun lambda_prs_simple_conv ctxt ctrm =
-  case (term_of ctrm) of
-    (Const (@{const_name fun_map}, _) $ r1 $ a2) $ (Abs _) =>
-      let
-        val thy = ProofContext.theory_of ctxt
-        val (ty_b, ty_a) = dest_fun_type (fastype_of r1)
-        val (ty_c, ty_d) = dest_fun_type (fastype_of a2)
-        val tyinst = map (SOME o (ctyp_of thy)) [ty_a, ty_b, ty_c, ty_d]
-        val tinst = [NONE, NONE, SOME (cterm_of thy r1), NONE, SOME (cterm_of thy a2)]
-        val thm1 = Drule.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]}
-        val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1)
-        val thm3 = MetaSimplifier.rewrite_rule @{thms id_apply[THEN eq_reflection]} thm2
-        val (insp, inst) =
-          if ty_c = ty_d
-          then make_inst_id (term_of (Thm.lhs_of thm3)) (term_of ctrm)
-          else make_inst (term_of (Thm.lhs_of thm3)) (term_of ctrm)
-        val thm4 = Drule.instantiate ([], [(cterm_of thy insp, cterm_of thy inst)]) thm3
-      in
-        Conv.rewr_conv thm4 ctrm
-      end
-  | _ => Conv.all_conv ctrm
-
-fun lambda_prs_conv ctxt = More_Conv.top_conv lambda_prs_simple_conv ctxt
-fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt)
-
-
-(* Cleaning consists of:
-
-  1. unfolding of ---> in front of everything, except
-     bound variables (this prevents lambda_prs from
-     becoming stuck)
-
-  2. simplification with lambda_prs
-
-  3. simplification with:
-
-      - Quotient_abs_rep Quotient_rel_rep
-        babs_prs all_prs ex_prs ex1_prs
-
-      - id_simps and preservation lemmas and
-
-      - symmetric versions of the definitions
-        (that is definitions of quotient constants
-         are folded)
-
-  4. test for refl
-*)
-fun clean_tac lthy =
-let
-  val defs = map (symmetric o #def) (qconsts_dest lthy)
-  val prs = prs_rules_get lthy
-  val ids = id_simps_get lthy
-  val thms = @{thms Quotient_abs_rep Quotient_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs
-
-  val ss = (mk_minimal_ss lthy) addsimps thms addSolver quotient_solver
-in
-  EVERY' [fun_map_tac lthy,
-          lambda_prs_tac lthy,
-          simp_tac ss,
-          TRY o rtac refl]
-end
-
-
-
-(** Tactic for Generalising Free Variables in a Goal **)
-
-fun inst_spec ctrm =
-   Drule.instantiate' [SOME (ctyp_of_term ctrm)] [NONE, SOME ctrm] @{thm spec}
-
-fun inst_spec_tac ctrms =
-  EVERY' (map (dtac o inst_spec) ctrms)
-
-fun all_list xs trm =
-  fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm
-
-fun apply_under_Trueprop f =
-  HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop
-
-fun gen_frees_tac ctxt =
-  SUBGOAL (fn (concl, i) =>
-    let
-      val thy = ProofContext.theory_of ctxt
-      val vrs = Term.add_frees concl []
-      val cvrs = map (cterm_of thy o Free) vrs
-      val concl' = apply_under_Trueprop (all_list vrs) concl
-      val goal = Logic.mk_implies (concl', concl)
-      val rule = Goal.prove ctxt [] [] goal
-        (K (EVERY1 [inst_spec_tac (rev cvrs), atac]))
-    in
-      rtac rule i
-    end)
-
-
-(** The General Shape of the Lifting Procedure **)
-
-(* - A is the original raw theorem
-   - B is the regularized theorem
-   - C is the rep/abs injected version of B
-   - D is the lifted theorem
-
-   - 1st prem is the regularization step
-   - 2nd prem is the rep/abs injection step
-   - 3rd prem is the cleaning part
-
-   the Quot_True premise in 2nd records the lifted theorem
-*)
-val lifting_procedure_thm =
-  @{lemma  "[|A;
-              A --> B;
-              Quot_True D ==> B = C;
-              C = D|] ==> D"
-      by (simp add: Quot_True_def)}
-
-fun lift_match_error ctxt msg rtrm qtrm =
-let
-  val rtrm_str = Syntax.string_of_term ctxt rtrm
-  val qtrm_str = Syntax.string_of_term ctxt qtrm
-  val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str,
-    "", "does not match with original theorem", rtrm_str]
-in
-  error msg
-end
-
-fun procedure_inst ctxt rtrm qtrm =
-let
-  val thy = ProofContext.theory_of ctxt
-  val rtrm' = HOLogic.dest_Trueprop rtrm
-  val qtrm' = HOLogic.dest_Trueprop qtrm
-  val reg_goal = regularize_trm_chk ctxt (rtrm', qtrm')
-    handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
-  val inj_goal = inj_repabs_trm_chk ctxt (reg_goal, qtrm')
-    handle (LIFT_MATCH msg) => lift_match_error ctxt msg rtrm qtrm
-in
-  Drule.instantiate' []
-    [SOME (cterm_of thy rtrm'),
-     SOME (cterm_of thy reg_goal),
-     NONE,
-     SOME (cterm_of thy inj_goal)] lifting_procedure_thm
-end
-
-(* the tactic leaves three subgoals to be proved *)
-fun procedure_tac ctxt rthm =
-  ObjectLogic.full_atomize_tac
-  THEN' gen_frees_tac ctxt
-  THEN' SUBGOAL (fn (goal, i) =>
-    let
-      val rthm' = atomize_thm rthm
-      val rule = procedure_inst ctxt (prop_of rthm') goal
-    in
-      (rtac rule THEN' rtac rthm') i
-    end)
-
-
-(* Automatic Proofs *)
-
-val msg1 = "The regularize proof failed."
-val msg2 = cat_lines ["The injection proof failed.",
-                      "This is probably due to missing respects lemmas.",
-                      "Try invoking the injection method manually to see",
-                      "which lemmas are missing."]
-val msg3 = "The cleaning proof failed."
-
-fun lift_tac ctxt rthms =
-let
-  fun mk_tac rthm =
-    procedure_tac ctxt rthm
-    THEN' RANGE_WARN
-      [(regularize_tac ctxt, msg1),
-       (all_injection_tac ctxt, msg2),
-       (clean_tac ctxt, msg3)]
-in
-  simp_tac (mk_minimal_ss ctxt) (* unfolding multiple &&& *)
-  THEN' RANGE (map mk_tac rthms)
-end
-
-(* An Attribute which automatically constructs the qthm *)
-fun lifted_attrib_aux context thm =
-let
-  val ctxt = Context.proof_of context
-  val ((_, [thm']), ctxt') = Variable.import false [thm] ctxt
-  val goal = (quotient_lift_all ctxt' o prop_of) thm'
-in
-  Goal.prove ctxt' [] [] goal (K (lift_tac ctxt' [thm] 1))
-  |> singleton (ProofContext.export ctxt' ctxt)
-end;
-
-val lifted_attrib = Thm.rule_attribute lifted_attrib_aux
-
-end; (* structure *)
--- a/Quot/quotient_term.ML	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,786 +0,0 @@
-(*  Title:      quotient_term.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-
-    Constructs terms corresponding to goals from
-    lifting theorems to quotient types.
-*)
-
-signature QUOTIENT_TERM =
-sig
-  exception LIFT_MATCH of string
-
-  datatype flag = AbsF | RepF
-
-  val absrep_fun: flag -> Proof.context -> typ * typ -> term
-  val absrep_fun_chk: flag -> Proof.context -> typ * typ -> term
-
-  (* Allows Nitpick to represent quotient types as single elements from raw type *)
-  val absrep_const_chk: flag -> Proof.context -> string -> term
-
-  val equiv_relation: Proof.context -> typ * typ -> term
-  val equiv_relation_chk: Proof.context -> typ * typ -> term
-
-  val regularize_trm: Proof.context -> term * term -> term
-  val regularize_trm_chk: Proof.context -> term * term -> term
-
-  val inj_repabs_trm: Proof.context -> term * term -> term
-  val inj_repabs_trm_chk: Proof.context -> term * term -> term
-
-  val quotient_lift_const: string * term -> local_theory -> term
-  val quotient_lift_all: Proof.context -> term -> term
-end;
-
-structure Quotient_Term: QUOTIENT_TERM =
-struct
-
-open Quotient_Info;
-
-exception LIFT_MATCH of string
-
-
-
-(*** Aggregate Rep/Abs Function ***)
-
-
-(* The flag RepF is for types in negative position; AbsF is for types
-   in positive position. Because of this, function types need to be
-   treated specially, since there the polarity changes.
-*)
-
-datatype flag = AbsF | RepF
-
-fun negF AbsF = RepF
-  | negF RepF = AbsF
-
-fun is_identity (Const (@{const_name "id"}, _)) = true
-  | is_identity _ = false
-
-fun mk_identity ty = Const (@{const_name "id"}, ty --> ty)
-
-fun mk_fun_compose flag (trm1, trm2) =
-  case flag of
-    AbsF => Const (@{const_name "comp"}, dummyT) $ trm1 $ trm2
-  | RepF => Const (@{const_name "comp"}, dummyT) $ trm2 $ trm1
-
-fun get_mapfun ctxt s =
-let
-  val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("No map function for type " ^ quote s ^ " found.")
-  val mapfun = #mapfun (maps_lookup thy s) handle Quotient_Info.NotFound => raise exn
-in
-  Const (mapfun, dummyT)
-end
-
-(* makes a Free out of a TVar *)
-fun mk_Free (TVar ((x, i), _)) = Free (unprefix "'" x ^ string_of_int i, dummyT)
-
-(* produces an aggregate map function for the
-   rty-part of a quotient definition; abstracts
-   over all variables listed in vs (these variables
-   correspond to the type variables in rty)
-
-   for example for: (?'a list * ?'b)
-   it produces:     %a b. prod_map (map a) b
-*)
-fun mk_mapfun ctxt vs rty =
-let
-  val vs' = map (mk_Free) vs
-
-  fun mk_mapfun_aux rty =
-    case rty of
-      TVar _ => mk_Free rty
-    | Type (_, []) => mk_identity rty
-    | Type (s, tys) => list_comb (get_mapfun ctxt s, map mk_mapfun_aux tys)
-    | _ => raise LIFT_MATCH "mk_mapfun (default)"
-in
-  fold_rev Term.lambda vs' (mk_mapfun_aux rty)
-end
-
-(* looks up the (varified) rty and qty for
-   a quotient definition
-*)
-fun get_rty_qty ctxt s =
-let
-  val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("No quotient type " ^ quote s ^ " found.")
-  val qdata = (quotdata_lookup thy s) handle Quotient_Info.NotFound => raise exn
-in
-  (#rtyp qdata, #qtyp qdata)
-end
-
-(* takes two type-environments and looks
-   up in both of them the variable v, which
-   must be listed in the environment
-*)
-fun double_lookup rtyenv qtyenv v =
-let
-  val v' = fst (dest_TVar v)
-in
-  (snd (the (Vartab.lookup rtyenv v')), snd (the (Vartab.lookup qtyenv v')))
-end
-
-(* matches a type pattern with a type *)
-fun match ctxt err ty_pat ty =
-let
-  val thy = ProofContext.theory_of ctxt
-in
-  Sign.typ_match thy (ty_pat, ty) Vartab.empty
-  handle MATCH_TYPE => err ctxt ty_pat ty
-end
-
-(* produces the rep or abs constant for a qty *)
-fun absrep_const flag ctxt qty_str =
-let
-  val thy = ProofContext.theory_of ctxt
-  val qty_name = Long_Name.base_name qty_str
-in
-  case flag of
-    AbsF => Const (Sign.full_bname thy ("abs_" ^ qty_name), dummyT)
-  | RepF => Const (Sign.full_bname thy ("rep_" ^ qty_name), dummyT)
-end
-
-(* Lets Nitpick represent elements of quotient types as elements of the raw type *)
-fun absrep_const_chk flag ctxt qty_str =
-  Syntax.check_term ctxt (absrep_const flag ctxt qty_str)
-
-fun absrep_match_err ctxt ty_pat ty =
-let
-  val ty_pat_str = Syntax.string_of_typ ctxt ty_pat
-  val ty_str = Syntax.string_of_typ ctxt ty
-in
-  raise LIFT_MATCH (space_implode " "
-    ["absrep_fun (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"])
-end
-
-
-(** generation of an aggregate absrep function **)
-
-(* - In case of equal types we just return the identity.
-
-   - In case of TFrees we also return the identity.
-
-   - In case of function types we recurse taking
-     the polarity change into account.
-
-   - If the type constructors are equal, we recurse for the
-     arguments and build the appropriate map function.
-
-   - If the type constructors are unequal, there must be an
-     instance of quotient types:
-
-       - we first look up the corresponding rty_pat and qty_pat
-         from the quotient definition; the arguments of qty_pat
-         must be some distinct TVars
-       - we then match the rty_pat with rty and qty_pat with qty;
-         if matching fails the types do not correspond -> error
-       - the matching produces two environments; we look up the
-         assignments for the qty_pat variables and recurse on the
-         assignments
-       - we prefix the aggregate map function for the rty_pat,
-         which is an abstraction over all type variables
-       - finally we compose the result with the appropriate
-         absrep function in case at least one argument produced
-         a non-identity function /
-         otherwise we just return the appropriate absrep
-         function
-
-     The composition is necessary for types like
-
-        ('a list) list / ('a foo) foo
-
-     The matching is necessary for types like
-
-        ('a * 'a) list / 'a bar
-
-     The test is necessary in order to eliminate superfluous
-     identity maps.
-*)
-
-fun absrep_fun flag ctxt (rty, qty) =
-  if rty = qty
-  then mk_identity rty
-  else
-    case (rty, qty) of
-      (Type ("fun", [ty1, ty2]), Type ("fun", [ty1', ty2'])) =>
-        let
-          val arg1 = absrep_fun (negF flag) ctxt (ty1, ty1')
-          val arg2 = absrep_fun flag ctxt (ty2, ty2')
-        in
-          list_comb (get_mapfun ctxt "fun", [arg1, arg2])
-        end
-    | (Type (s, tys), Type (s', tys')) =>
-        if s = s'
-        then
-           let
-             val args = map (absrep_fun flag ctxt) (tys ~~ tys')
-           in
-             list_comb (get_mapfun ctxt s, args)
-           end
-        else
-           let
-             val (rty_pat, qty_pat as Type (_, vs)) = get_rty_qty ctxt s'
-             val rtyenv = match ctxt absrep_match_err rty_pat rty
-             val qtyenv = match ctxt absrep_match_err qty_pat qty
-             val args_aux = map (double_lookup rtyenv qtyenv) vs
-             val args = map (absrep_fun flag ctxt) args_aux
-             val map_fun = mk_mapfun ctxt vs rty_pat
-             val result = list_comb (map_fun, args)
-           in
-             if forall is_identity args
-             then absrep_const flag ctxt s'
-             else mk_fun_compose flag (absrep_const flag ctxt s', result)
-           end
-    | (TFree x, TFree x') =>
-        if x = x'
-        then mk_identity rty
-        else raise (LIFT_MATCH "absrep_fun (frees)")
-    | (TVar _, TVar _) => raise (LIFT_MATCH "absrep_fun (vars)")
-    | _ => raise (LIFT_MATCH "absrep_fun (default)")
-
-fun absrep_fun_chk flag ctxt (rty, qty) =
-  absrep_fun flag ctxt (rty, qty)
-  |> Syntax.check_term ctxt
-
-
-
-
-(*** Aggregate Equivalence Relation ***)
-
-
-(* works very similar to the absrep generation,
-   except there is no need for polarities
-*)
-
-(* instantiates TVars so that the term is of type ty *)
-fun force_typ ctxt trm ty =
-let
-  val thy = ProofContext.theory_of ctxt
-  val trm_ty = fastype_of trm
-  val ty_inst = Sign.typ_match thy (trm_ty, ty) Vartab.empty
-in
-  map_types (Envir.subst_type ty_inst) trm
-end
-
-fun is_eq (Const (@{const_name "op ="}, _)) = true
-  | is_eq _ = false
-
-fun mk_rel_compose (trm1, trm2) =
-  Const (@{const_name "rel_conj"}, dummyT) $ trm1 $ trm2
-
-fun get_relmap ctxt s =
-let
-  val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("get_relmap (no relation map function found for type " ^ s ^ ")")
-  val relmap = #relmap (maps_lookup thy s) handle Quotient_Info.NotFound => raise exn
-in
-  Const (relmap, dummyT)
-end
-
-fun mk_relmap ctxt vs rty =
-let
-  val vs' = map (mk_Free) vs
-
-  fun mk_relmap_aux rty =
-    case rty of
-      TVar _ => mk_Free rty
-    | Type (_, []) => HOLogic.eq_const rty
-    | Type (s, tys) => list_comb (get_relmap ctxt s, map mk_relmap_aux tys)
-    | _ => raise LIFT_MATCH ("mk_relmap (default)")
-in
-  fold_rev Term.lambda vs' (mk_relmap_aux rty)
-end
-
-fun get_equiv_rel ctxt s =
-let
-  val thy = ProofContext.theory_of ctxt
-  val exn = LIFT_MATCH ("get_quotdata (no quotient found for type " ^ s ^ ")")
-in
-  #equiv_rel (quotdata_lookup thy s) handle Quotient_Info.NotFound => raise exn
-end
-
-fun equiv_match_err ctxt ty_pat ty =
-let
-  val ty_pat_str = Syntax.string_of_typ ctxt ty_pat
-  val ty_str = Syntax.string_of_typ ctxt ty
-in
-  raise LIFT_MATCH (space_implode " "
-    ["equiv_relation (Types ", quote ty_pat_str, "and", quote ty_str, " do not match.)"])
-end
-
-(* builds the aggregate equivalence relation
-   that will be the argument of Respects
-*)
-fun equiv_relation ctxt (rty, qty) =
-  if rty = qty
-  then HOLogic.eq_const rty
-  else
-    case (rty, qty) of
-      (Type (s, tys), Type (s', tys')) =>
-       if s = s'
-       then
-         let
-           val args = map (equiv_relation ctxt) (tys ~~ tys')
-         in
-           list_comb (get_relmap ctxt s, args)
-         end
-       else
-         let
-           val (rty_pat, qty_pat as Type (_, vs)) = get_rty_qty ctxt s'
-           val rtyenv = match ctxt equiv_match_err rty_pat rty
-           val qtyenv = match ctxt equiv_match_err qty_pat qty
-           val args_aux = map (double_lookup rtyenv qtyenv) vs
-           val args = map (equiv_relation ctxt) args_aux
-           val rel_map = mk_relmap ctxt vs rty_pat
-           val result = list_comb (rel_map, args)
-           val eqv_rel = get_equiv_rel ctxt s'
-           val eqv_rel' = force_typ ctxt eqv_rel ([rty, rty] ---> @{typ bool})
-         in
-           if forall is_eq args
-           then eqv_rel'
-           else mk_rel_compose (result, eqv_rel')
-         end
-      | _ => HOLogic.eq_const rty
-
-fun equiv_relation_chk ctxt (rty, qty) =
-  equiv_relation ctxt (rty, qty)
-  |> Syntax.check_term ctxt
-
-
-
-(*** Regularization ***)
-
-(* Regularizing an rtrm means:
-
- - Quantifiers over types that need lifting are replaced
-   by bounded quantifiers, for example:
-
-      All P  ----> All (Respects R) P
-
-   where the aggregate relation R is given by the rty and qty;
-
- - Abstractions over types that need lifting are replaced
-   by bounded abstractions, for example:
-
-      %x. P  ----> Ball (Respects R) %x. P
-
- - Equalities over types that need lifting are replaced by
-   corresponding equivalence relations, for example:
-
-      A = B  ----> R A B
-
-   or
-
-      A = B  ----> (R ===> R) A B
-
-   for more complicated types of A and B
-
-
- The regularize_trm accepts raw theorems in which equalities
- and quantifiers match exactly the ones in the lifted theorem
- but also accepts partially regularized terms.
-
- This means that the raw theorems can have:
-   Ball (Respects R),  Bex (Respects R), Bex1_rel (Respects R), Babs, R
- in the places where:
-   All, Ex, Ex1, %, (op =)
- is required the lifted theorem.
-
-*)
-
-val mk_babs = Const (@{const_name Babs}, dummyT)
-val mk_ball = Const (@{const_name Ball}, dummyT)
-val mk_bex  = Const (@{const_name Bex}, dummyT)
-val mk_bex1_rel = Const (@{const_name Bex1_rel}, dummyT)
-val mk_resp = Const (@{const_name Respects}, dummyT)
-
-(* - applies f to the subterm of an abstraction,
-     otherwise to the given term,
-   - used by regularize, therefore abstracted
-     variables do not have to be treated specially
-*)
-fun apply_subt f (trm1, trm2) =
-  case (trm1, trm2) of
-    (Abs (x, T, t), Abs (_ , _, t')) => Abs (x, T, f (t, t'))
-  | _ => f (trm1, trm2)
-
-fun term_mismatch str ctxt t1 t2 =
-let
-  val t1_str = Syntax.string_of_term ctxt t1
-  val t2_str = Syntax.string_of_term ctxt t2
-  val t1_ty_str = Syntax.string_of_typ ctxt (fastype_of t1)
-  val t2_ty_str = Syntax.string_of_typ ctxt (fastype_of t2)
-in
-  raise LIFT_MATCH (cat_lines [str, t1_str ^ "::" ^ t1_ty_str, t2_str ^ "::" ^ t2_ty_str])
-end
-
-(* the major type of All and Ex quantifiers *)
-fun qnt_typ ty = domain_type (domain_type ty)
-
-(* Checks that two types match, for example:
-     rty -> rty   matches   qty -> qty *)
-fun matches_typ thy rT qT =
-  if rT = qT then true else
-  case (rT, qT) of
-    (Type (rs, rtys), Type (qs, qtys)) =>
-      if rs = qs then
-        if length rtys <> length qtys then false else
-        forall (fn x => x = true) (map2 (matches_typ thy) rtys qtys)
-      else
-        (case Quotient_Info.quotdata_lookup_raw thy qs of
-          SOME quotinfo => Sign.typ_instance thy (rT, #rtyp quotinfo)
-        | NONE => false)
-  | _ => false
-
-
-(* produces a regularized version of rtrm
-
-   - the result might contain dummyTs
-
-   - for regularisation we do not need any
-     special treatment of bound variables
-*)
-fun regularize_trm ctxt (rtrm, qtrm) =
-  case (rtrm, qtrm) of
-    (Abs (x, ty, t), Abs (_, ty', t')) =>
-       let
-         val subtrm = Abs(x, ty, regularize_trm ctxt (t, t'))
-       in
-         if ty = ty' then subtrm
-         else mk_babs $ (mk_resp $ equiv_relation ctxt (ty, ty')) $ subtrm
-       end
-  | (Const (@{const_name "Babs"}, T) $ resrel $ (t as (Abs (_, ty, _))), t' as (Abs (_, ty', _))) =>
-       let
-         val subtrm = regularize_trm ctxt (t, t')
-         val needres = mk_resp $ equiv_relation_chk ctxt (ty, ty')
-       in
-         if resrel <> needres
-         then term_mismatch "regularize (Babs)" ctxt resrel needres
-         else mk_babs $ resrel $ subtrm
-       end
-
-  | (Const (@{const_name "All"}, ty) $ t, Const (@{const_name "All"}, ty') $ t') =>
-       let
-         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
-       in
-         if ty = ty' then Const (@{const_name "All"}, ty) $ subtrm
-         else mk_ball $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
-       end
-
-  | (Const (@{const_name "Ex"}, ty) $ t, Const (@{const_name "Ex"}, ty') $ t') =>
-       let
-         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
-       in
-         if ty = ty' then Const (@{const_name "Ex"}, ty) $ subtrm
-         else mk_bex $ (mk_resp $ equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
-       end
-
-  | (Const (@{const_name "Ex1"}, ty) $ (Abs (_, _,
-      (Const (@{const_name "op &"}, _) $ (Const (@{const_name "op :"}, _) $ _ $
-        (Const (@{const_name "Respects"}, _) $ resrel)) $ (t $ _)))),
-     Const (@{const_name "Ex1"}, ty') $ t') =>
-       let
-         val t_ = incr_boundvars (~1) t
-         val subtrm = apply_subt (regularize_trm ctxt) (t_, t')
-         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
-       in
-         if resrel <> needrel
-         then term_mismatch "regularize (Bex1)" ctxt resrel needrel
-         else mk_bex1_rel $ resrel $ subtrm
-       end
-
-  | (Const (@{const_name "Ex1"}, ty) $ t, Const (@{const_name "Ex1"}, ty') $ t') =>
-       let
-         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
-       in
-         if ty = ty' then Const (@{const_name "Ex1"}, ty) $ subtrm
-         else mk_bex1_rel $ (equiv_relation ctxt (qnt_typ ty, qnt_typ ty')) $ subtrm
-       end
-
-  | (Const (@{const_name "Ball"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t,
-     Const (@{const_name "All"}, ty') $ t') =>
-       let
-         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
-         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
-       in
-         if resrel <> needrel
-         then term_mismatch "regularize (Ball)" ctxt resrel needrel
-         else mk_ball $ (mk_resp $ resrel) $ subtrm
-       end
-
-  | (Const (@{const_name "Bex"}, ty) $ (Const (@{const_name "Respects"}, _) $ resrel) $ t,
-     Const (@{const_name "Ex"}, ty') $ t') =>
-       let
-         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
-         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
-       in
-         if resrel <> needrel
-         then term_mismatch "regularize (Bex)" ctxt resrel needrel
-         else mk_bex $ (mk_resp $ resrel) $ subtrm
-       end
-
-  | (Const (@{const_name "Bex1_rel"}, ty) $ resrel $ t, Const (@{const_name "Ex1"}, ty') $ t') =>
-       let
-         val subtrm = apply_subt (regularize_trm ctxt) (t, t')
-         val needrel = equiv_relation_chk ctxt (qnt_typ ty, qnt_typ ty')
-       in
-         if resrel <> needrel
-         then term_mismatch "regularize (Bex1_res)" ctxt resrel needrel
-         else mk_bex1_rel $ resrel $ subtrm
-       end
-
-  | (* equalities need to be replaced by appropriate equivalence relations *)
-    (Const (@{const_name "op ="}, ty), Const (@{const_name "op ="}, ty')) =>
-         if ty = ty' then rtrm
-         else equiv_relation ctxt (domain_type ty, domain_type ty')
-
-  | (* in this case we just check whether the given equivalence relation is correct *)
-    (rel, Const (@{const_name "op ="}, ty')) =>
-       let
-         val rel_ty = fastype_of rel
-         val rel' = equiv_relation_chk ctxt (domain_type rel_ty, domain_type ty')
-       in
-         if rel' aconv rel then rtrm
-         else term_mismatch "regularise (relation mismatch)" ctxt rel rel'
-       end
-
-  | (_, Const _) =>
-       let
-         val thy = ProofContext.theory_of ctxt
-         fun same_const (Const (s, T)) (Const (s', T')) = (s = s') andalso matches_typ thy T T'
-           | same_const _ _ = false
-       in
-         if same_const rtrm qtrm then rtrm
-         else
-           let
-             val rtrm' = #rconst (qconsts_lookup thy qtrm)
-               handle Quotient_Info.NotFound => term_mismatch "regularize(constant notfound)" ctxt rtrm qtrm
-           in
-             if Pattern.matches thy (rtrm', rtrm)
-             then rtrm else term_mismatch "regularize(constant mismatch)" ctxt rtrm qtrm
-           end
-       end
-
-  | (((t1 as Const (@{const_name "split"}, _)) $ Abs (v1, ty, Abs(v1', ty', s1))),
-     ((t2 as Const (@{const_name "split"}, _)) $ Abs (v2, _ , Abs(v2', _  , s2)))) =>
-       regularize_trm ctxt (t1, t2) $ Abs (v1, ty, Abs (v1', ty', regularize_trm ctxt (s1, s2)))
-
-  | (((t1 as Const (@{const_name "split"}, _)) $ Abs (v1, ty, s1)),
-     ((t2 as Const (@{const_name "split"}, _)) $ Abs (v2, _ , s2))) =>
-       regularize_trm ctxt (t1, t2) $ Abs (v1, ty, regularize_trm ctxt (s1, s2))
-
-  | (t1 $ t2, t1' $ t2') =>
-       regularize_trm ctxt (t1, t1') $ regularize_trm ctxt (t2, t2')
-
-  | (Bound i, Bound i') =>
-       if i = i' then rtrm
-       else raise (LIFT_MATCH "regularize (bounds mismatch)")
-
-  | _ =>
-       let
-         val rtrm_str = Syntax.string_of_term ctxt rtrm
-         val qtrm_str = Syntax.string_of_term ctxt qtrm
-       in
-         raise (LIFT_MATCH ("regularize failed (default: " ^ rtrm_str ^ "," ^ qtrm_str ^ ")"))
-       end
-
-fun regularize_trm_chk ctxt (rtrm, qtrm) =
-  regularize_trm ctxt (rtrm, qtrm)
-  |> Syntax.check_term ctxt
-
-
-
-(*** Rep/Abs Injection ***)
-
-(*
-Injection of Rep/Abs means:
-
-  For abstractions:
-
-  * If the type of the abstraction needs lifting, then we add Rep/Abs
-    around the abstraction; otherwise we leave it unchanged.
-
-  For applications:
-
-  * If the application involves a bounded quantifier, we recurse on
-    the second argument. If the application is a bounded abstraction,
-    we always put an Rep/Abs around it (since bounded abstractions
-    are assumed to always need lifting). Otherwise we recurse on both
-    arguments.
-
-  For constants:
-
-  * If the constant is (op =), we leave it always unchanged.
-    Otherwise the type of the constant needs lifting, we put
-    and Rep/Abs around it.
-
-  For free variables:
-
-  * We put a Rep/Abs around it if the type needs lifting.
-
-  Vars case cannot occur.
-*)
-
-fun mk_repabs ctxt (T, T') trm =
-  absrep_fun RepF ctxt (T, T') $ (absrep_fun AbsF ctxt (T, T') $ trm)
-
-fun inj_repabs_err ctxt msg rtrm qtrm =
-let
-  val rtrm_str = Syntax.string_of_term ctxt rtrm
-  val qtrm_str = Syntax.string_of_term ctxt qtrm
-in
-  raise LIFT_MATCH (space_implode " " [msg, quote rtrm_str, "and", quote qtrm_str])
-end
-
-
-(* bound variables need to be treated properly,
-   as the type of subterms needs to be calculated   *)
-fun inj_repabs_trm ctxt (rtrm, qtrm) =
- case (rtrm, qtrm) of
-    (Const (@{const_name "Ball"}, T) $ r $ t, Const (@{const_name "All"}, _) $ t') =>
-       Const (@{const_name "Ball"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
-
-  | (Const (@{const_name "Bex"}, T) $ r $ t, Const (@{const_name "Ex"}, _) $ t') =>
-       Const (@{const_name "Bex"}, T) $ r $ (inj_repabs_trm ctxt (t, t'))
-
-  | (Const (@{const_name "Babs"}, T) $ r $ t, t' as (Abs _)) =>
-      let
-        val rty = fastype_of rtrm
-        val qty = fastype_of qtrm
-      in
-        mk_repabs ctxt (rty, qty) (Const (@{const_name "Babs"}, T) $ r $ (inj_repabs_trm ctxt (t, t')))
-      end
-
-  | (Abs (x, T, t), Abs (x', T', t')) =>
-      let
-        val rty = fastype_of rtrm
-        val qty = fastype_of qtrm
-        val (y, s) = Term.dest_abs (x, T, t)
-        val (_, s') = Term.dest_abs (x', T', t')
-        val yvar = Free (y, T)
-        val result = Term.lambda_name (y, yvar) (inj_repabs_trm ctxt (s, s'))
-      in
-        if rty = qty then result
-        else mk_repabs ctxt (rty, qty) result
-      end
-
-  | (t $ s, t' $ s') =>
-       (inj_repabs_trm ctxt (t, t')) $ (inj_repabs_trm ctxt (s, s'))
-
-  | (Free (_, T), Free (_, T')) =>
-        if T = T' then rtrm
-        else mk_repabs ctxt (T, T') rtrm
-
-  | (_, Const (@{const_name "op ="}, _)) => rtrm
-
-  | (_, Const (_, T')) =>
-      let
-        val rty = fastype_of rtrm
-      in
-        if rty = T' then rtrm
-        else mk_repabs ctxt (rty, T') rtrm
-      end
-
-  | _ => inj_repabs_err ctxt "injection (default):" rtrm qtrm
-
-fun inj_repabs_trm_chk ctxt (rtrm, qtrm) =
-  inj_repabs_trm ctxt (rtrm, qtrm)
-  |> Syntax.check_term ctxt
-
-
-
-(*** Wrapper for automatically transforming an rthm into a qthm ***)
-
-(* subst_tys takes a list of (rty, qty) substitution pairs
-   and replaces all occurences of rty in the given type
-   by appropriate qty, with substitution *)
-fun subst_ty thy ty (rty, qty) r =
-  if r <> NONE then r else
-  case try (Sign.typ_match thy (rty, ty)) Vartab.empty of
-    SOME inst => SOME (Envir.subst_type inst qty)
-  | NONE => NONE
-fun subst_tys thy substs ty =
-  case fold (subst_ty thy ty) substs NONE of
-    SOME ty => ty
-  | NONE =>
-      (case ty of
-        Type (s, tys) => Type (s, map (subst_tys thy substs) tys)
-      | x => x)
-
-(* subst_trms takes a list of (rtrm, qtrm) substitution pairs
-   and if the given term matches any of the raw terms it
-   returns the appropriate qtrm instantiated. If none of
-   them matched it returns NONE. *)
-fun subst_trm thy t (rtrm, qtrm) s =
-  if s <> NONE then s else
-    case try (Pattern.match thy (rtrm, t)) (Vartab.empty, Vartab.empty) of
-      SOME inst => SOME (Envir.subst_term inst qtrm)
-    | NONE => NONE;
-fun subst_trms thy substs t = fold (subst_trm thy t) substs NONE
-
-(* prepares type and term substitution pairs to be used by above
-   functions that let replace all raw constructs by appropriate
-   lifted counterparts. *)
-fun get_ty_trm_substs ctxt =
-let
-  val thy = ProofContext.theory_of ctxt
-  val quot_infos  = Quotient_Info.quotdata_dest ctxt
-  val const_infos = Quotient_Info.qconsts_dest ctxt
-  val ty_substs = map (fn ri => (#rtyp ri, #qtyp ri)) quot_infos
-  val const_substs = map (fn ci => (#rconst ci, #qconst ci)) const_infos
-  fun rel_eq rel = HOLogic.eq_const (subst_tys thy ty_substs (domain_type (fastype_of rel)))
-  val rel_substs = map (fn ri => (#equiv_rel ri, rel_eq (#equiv_rel ri))) quot_infos
-in
-  (ty_substs, const_substs @ rel_substs)
-end
-
-fun quotient_lift_const (b, t) ctxt =
-let
-  val thy = ProofContext.theory_of ctxt
-  val (ty_substs, _) = get_ty_trm_substs ctxt;
-  val (_, ty) = dest_Const t;
-  val nty = subst_tys thy ty_substs ty;
-in
-  Free(b, nty)
-end
-
-(*
-Takes a term and
-
-* replaces raw constants by the quotient constants
-
-* replaces equivalence relations by equalities
-
-* replaces raw types by the quotient types
-
-*)
-
-fun quotient_lift_all ctxt t =
-let
-  val thy = ProofContext.theory_of ctxt
-  val (ty_substs, substs) = get_ty_trm_substs ctxt
-  fun lift_aux t =
-    case subst_trms thy substs t of
-      SOME x => x
-    | NONE =>
-      (case t of
-        a $ b => lift_aux a $ lift_aux b
-      | Abs(a, ty, s) =>
-          let
-            val (y, s') = Term.dest_abs (a, ty, s)
-            val nty = subst_tys thy ty_substs ty
-          in
-            Abs(y, nty, abstract_over (Free (y, nty), lift_aux s'))
-          end
-      | Free(n, ty) => Free(n, subst_tys thy ty_substs ty)
-      | Var(n, ty) => Var(n, subst_tys thy ty_substs ty)
-      | Bound i => Bound i
-      | Const(s, ty) => Const(s, subst_tys thy ty_substs ty))
-in
-  lift_aux t
-end
-
-
-end; (* structure *)
-
-
-
--- a/Quot/quotient_typ.ML	Thu Feb 25 07:48:57 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,309 +0,0 @@
-(*  Title:      quotient_typ.thy
-    Author:     Cezary Kaliszyk and Christian Urban
-
-    Definition of a quotient type.
-
-*)
-
-signature QUOTIENT_TYPE =
-sig
-  val quotient_type: ((string list * binding * mixfix) * (typ * term)) list
-    -> Proof.context -> Proof.state
-
-  val quotient_type_cmd: ((((string list * binding) * mixfix) * string) * string) list
-    -> Proof.context -> Proof.state
-end;
-
-structure Quotient_Type: QUOTIENT_TYPE =
-struct
-
-open Quotient_Info;
-
-(* wrappers for define, note, Attrib.internal and theorem_i *)
-fun define (name, mx, rhs) lthy =
-let
-  val ((rhs, (_ , thm)), lthy') =
-     Local_Theory.define ((name, mx), (Attrib.empty_binding, rhs)) lthy
-in
-  ((rhs, thm), lthy')
-end
-
-fun note (name, thm, attrs) lthy =
-let
-  val ((_,[thm']), lthy') = Local_Theory.note ((name, attrs), [thm]) lthy
-in
-  (thm', lthy')
-end
-
-fun intern_attr at = Attrib.internal (K at)
-
-fun theorem after_qed goals ctxt =
-let
-  val goals' = map (rpair []) goals
-  fun after_qed' thms = after_qed (the_single thms)
-in
-  Proof.theorem_i NONE after_qed' [goals'] ctxt
-end
-
-
-
-(*** definition of quotient types ***)
-
-val mem_def1 = @{lemma "y : S ==> S y" by (simp add: mem_def)}
-val mem_def2 = @{lemma "S y ==> y : S" by (simp add: mem_def)}
-
-(* constructs the term lambda (c::rty => bool). EX (x::rty). c = rel x *)
-fun typedef_term rel rty lthy =
-let
-  val [x, c] =
-    [("x", rty), ("c", HOLogic.mk_setT rty)]
-    |> Variable.variant_frees lthy [rel]
-    |> map Free
-in
-  lambda c (HOLogic.exists_const rty $
-     lambda x (HOLogic.mk_eq (c, (rel $ x))))
-end
-
-
-(* makes the new type definitions and proves non-emptyness *)
-fun typedef_make (vs, qty_name, mx, rel, rty) lthy =
-let
-  val typedef_tac =
-     EVERY1 (map rtac [@{thm exI}, mem_def2, @{thm exI}, @{thm refl}])
-in
-  Local_Theory.theory_result
-    (Typedef.add_typedef false NONE
-       (qty_name, vs, mx)
-          (typedef_term rel rty lthy)
-             NONE typedef_tac) lthy
-end
-
-
-(* tactic to prove the quot_type theorem for the new type *)
-fun typedef_quot_type_tac equiv_thm (typedef_info: Typedef.info) =
-let
-  val rep_thm = #Rep typedef_info RS mem_def1
-  val rep_inv = #Rep_inverse typedef_info
-  val abs_inv = mem_def2 RS #Abs_inverse typedef_info
-  val rep_inj = #Rep_inject typedef_info
-in
-  (rtac @{thm quot_type.intro} THEN' RANGE [
-    rtac equiv_thm,
-    rtac rep_thm,
-    rtac rep_inv,
-    EVERY' (map rtac [abs_inv, @{thm exI}, @{thm refl}]),
-    rtac rep_inj]) 1
-end
-
-
-(* proves the quot_type theorem for the new type *)
-fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy =
-let
-  val quot_type_const = Const (@{const_name "quot_type"}, dummyT)
-  val goal =
-    HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep)
-    |> Syntax.check_term lthy
-in
-  Goal.prove lthy [] [] goal
-    (K (typedef_quot_type_tac equiv_thm typedef_info))
-end
-
-(* proves the quotient theorem for the new type *)
-fun typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_type_thm) lthy =
-let
-  val quotient_const = Const (@{const_name "Quotient"}, dummyT)
-  val goal =
-    HOLogic.mk_Trueprop (quotient_const $ rel $ abs $ rep)
-    |> Syntax.check_term lthy
-
-  val typedef_quotient_thm_tac =
-    EVERY1 [
-      K (rewrite_goals_tac [abs_def, rep_def]),
-      rtac @{thm quot_type.Quotient},
-      rtac quot_type_thm]
-in
-  Goal.prove lthy [] [] goal
-    (K typedef_quotient_thm_tac)
-end
-
-
-(* main function for constructing a quotient type *)
-fun mk_quotient_type (((vs, qty_name, mx), (rty, rel)), equiv_thm) lthy =
-let
-  (* generates the typedef *)
-  val ((qty_full_name, typedef_info), lthy1) = typedef_make (vs, qty_name, mx, rel, rty) lthy
-
-  (* abs and rep functions from the typedef *)
-  val Abs_ty = #abs_type typedef_info
-  val Rep_ty = #rep_type typedef_info
-  val Abs_name = #Abs_name typedef_info
-  val Rep_name = #Rep_name typedef_info
-  val Abs_const = Const (Abs_name, Rep_ty --> Abs_ty)
-  val Rep_const = Const (Rep_name, Abs_ty --> Rep_ty)
-
-  (* more useful abs and rep definitions *)
-  val abs_const = Const (@{const_name "quot_type.abs"}, dummyT )
-  val rep_const = Const (@{const_name "quot_type.rep"}, dummyT )
-  val abs_trm = Syntax.check_term lthy1 (abs_const $ rel $ Abs_const)
-  val rep_trm = Syntax.check_term lthy1 (rep_const $ Rep_const)
-  val abs_name = Binding.prefix_name "abs_" qty_name
-  val rep_name = Binding.prefix_name "rep_" qty_name
-
-  val ((abs, abs_def), lthy2) = define (abs_name, NoSyn, abs_trm) lthy1
-  val ((rep, rep_def), lthy3) = define (rep_name, NoSyn, rep_trm) lthy2
-
-  (* quot_type theorem *)
-  val quot_thm = typedef_quot_type_thm (rel, Abs_const, Rep_const, equiv_thm, typedef_info) lthy3
-
-  (* quotient theorem *)
-  val quotient_thm = typedef_quotient_thm (rel, abs, rep, abs_def, rep_def, quot_thm) lthy3
-  val quotient_thm_name = Binding.prefix_name "Quotient_" qty_name
-
-  (* name equivalence theorem *)
-  val equiv_thm_name = Binding.suffix_name "_equivp" qty_name
-
-  (* storing the quot-info *)
-  fun qinfo phi = transform_quotdata phi
-    {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm}
-  val lthy4 = Local_Theory.declaration true
-    (fn phi => quotdata_update_gen qty_full_name (qinfo phi)) lthy3
-in
-  lthy4
-  |> note (quotient_thm_name, quotient_thm, [intern_attr quotient_rules_add])
-  ||>> note (equiv_thm_name, equiv_thm, [intern_attr equiv_rules_add])
-end
-
-
-(* sanity checks for the quotient type specifications *)
-fun sanity_check ((vs, qty_name, _), (rty, rel)) =
-let
-  val rty_tfreesT = map fst (Term.add_tfreesT rty [])
-  val rel_tfrees = map fst (Term.add_tfrees rel [])
-  val rel_frees = map fst (Term.add_frees rel [])
-  val rel_vars = Term.add_vars rel []
-  val rel_tvars = Term.add_tvars rel []
-  val qty_str = Binding.str_of qty_name ^ ": "
-
-  val illegal_rel_vars =
-    if null rel_vars andalso null rel_tvars then []
-    else [qty_str ^ "illegal schematic variable(s) in the relation."]
-
-  val dup_vs =
-    (case duplicates (op =) vs of
-       [] => []
-     | dups => [qty_str ^ "duplicate type variable(s) on the lhs: " ^ commas_quote dups])
-
-  val extra_rty_tfrees =
-    (case subtract (op =) vs rty_tfreesT of
-       [] => []
-     | extras => [qty_str ^ "extra type variable(s) on the lhs: " ^ commas_quote extras])
-
-  val extra_rel_tfrees =
-    (case subtract (op =) vs rel_tfrees of
-       [] => []
-     | extras => [qty_str ^ "extra type variable(s) in the relation: " ^ commas_quote extras])
-
-  val illegal_rel_frees =
-    (case rel_frees of
-      [] => []
-    | xs => [qty_str ^ "illegal variable(s) in the relation: " ^ commas_quote xs])
-
-  val errs = illegal_rel_vars @ dup_vs @ extra_rty_tfrees @ extra_rel_tfrees @ illegal_rel_frees
-in
-  if null errs then () else error (cat_lines errs)
-end
-
-(* check for existence of map functions *)
-fun map_check ctxt (_, (rty, _)) =
-let
-  val thy = ProofContext.theory_of ctxt
-
-  fun map_check_aux rty warns =
-    case rty of
-      Type (_, []) => warns
-    | Type (s, _) => if maps_defined thy s then warns else s::warns
-    | _ => warns
-
-  val warns = map_check_aux rty []
-in
-  if null warns then ()
-  else warning ("No map function defined for " ^ commas warns ^
-    ". This will cause problems later on.")
-end
-
-
-
-(*** interface and syntax setup ***)
-
-
-(* the ML-interface takes a list of 5-tuples consisting of:
-
- - the name of the quotient type
- - its free type variables (first argument)
- - its mixfix annotation
- - the type to be quotient
- - the relation according to which the type is quotient
-
- it opens a proof-state in which one has to show that the
- relations are equivalence relations
-*)
-
-fun quotient_type quot_list lthy =
-let
-  (* sanity check *)
-  val _ = List.app sanity_check quot_list
-  val _ = List.app (map_check lthy) quot_list
-
-  fun mk_goal (rty, rel) =
-  let
-    val equivp_ty = ([rty, rty] ---> @{typ bool}) --> @{typ bool}
-  in
-    HOLogic.mk_Trueprop (Const (@{const_name equivp}, equivp_ty) $ rel)
-  end
-
-  val goals = map (mk_goal o snd) quot_list
-
-  fun after_qed thms lthy =
-    fold_map mk_quotient_type (quot_list ~~ thms) lthy |> snd
-in
-  theorem after_qed goals lthy
-end
-
-fun quotient_type_cmd specs lthy =
-let
-  fun parse_spec ((((vs, qty_name), mx), rty_str), rel_str) lthy =
-  let
-    (* new parsing with proper declaration *)
-    val rty = Syntax.read_typ lthy rty_str
-    val lthy1 = Variable.declare_typ rty lthy
-    val rel = 
-      Syntax.parse_term lthy1 rel_str
-      |> Syntax.type_constraint (rty --> rty --> @{typ bool}) 
-      |> Syntax.check_term lthy1 
-    val lthy2 = Variable.declare_term rel lthy1
-  in
-    (((vs, qty_name, mx), (rty, rel)), lthy2)
-  end
-
-  val (spec', lthy') = fold_map parse_spec specs lthy
-in
-  quotient_type spec' lthy'
-end
-
-local
-  structure P = OuterParse;
-in
-
-val quotspec_parser =
-  P.and_list1 ((P.type_args -- P.binding) -- P.opt_infix -- 
-    (P.$$$ "=" |-- P.typ) -- (P.$$$ "/" |-- P.term))
-end
-
-val _ = OuterKeyword.keyword "/"
-
-val _ =
-  OuterSyntax.local_theory_to_proof "quotient_type"
-    "quotient type definitions (require equivalence proofs)"
-       OuterKeyword.thy_goal (quotspec_parser >> quotient_type_cmd)
-
-end; (* structure *)