moved Nominal to "toplevel"
authorChristian Urban <urbanc@in.tum.de>
Thu, 25 Feb 2010 07:48:33 +0100
changeset 1258 7d8949da7d99
parent 1252 4b0563bc4b03
child 1259 db158e995bfc
moved Nominal to "toplevel"
Nominal/Abs.thy
Nominal/Fv.thy
Nominal/LFex.thy
Nominal/LamEx.thy
Nominal/LamEx2.thy
Nominal/Nominal2_Atoms.thy
Nominal/Nominal2_Base.thy
Nominal/Nominal2_Eqvt.thy
Nominal/Nominal2_Supp.thy
Nominal/Perm.thy
Nominal/Rsp.thy
Nominal/Terms.thy
Nominal/Test.thy
Nominal/nominal_atoms.ML
Nominal/nominal_permeq.ML
Nominal/nominal_thmdecls.ML
Quot/Nominal/Abs.thy
Quot/Nominal/Fv.thy
Quot/Nominal/LFex.thy
Quot/Nominal/LamEx.thy
Quot/Nominal/LamEx2.thy
Quot/Nominal/Nominal2_Atoms.thy
Quot/Nominal/Nominal2_Base.thy
Quot/Nominal/Nominal2_Eqvt.thy
Quot/Nominal/Nominal2_Supp.thy
Quot/Nominal/Perm.thy
Quot/Nominal/Rsp.thy
Quot/Nominal/Terms.thy
Quot/Nominal/Test.thy
Quot/Nominal/nominal_atoms.ML
Quot/Nominal/nominal_permeq.ML
Quot/Nominal/nominal_thmdecls.ML
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Abs.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,506 @@
+theory Abs
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "../Quotient" "../Quotient_Product"
+begin
+
+(* the next three lemmas that should be in Nominal \<dots>\<dots>must be cleaned *)
+lemma ball_image: 
+  shows "(\<forall>x \<in> p \<bullet> S. P x) = (\<forall>x \<in> S. P (p \<bullet> x))"
+apply(auto)
+apply(drule_tac x="p \<bullet> x" in bspec)
+apply(simp add: mem_permute_iff)
+apply(simp)
+apply(drule_tac x="(- p) \<bullet> x" in bspec)
+apply(rule_tac p1="p" in mem_permute_iff[THEN iffD1])
+apply(simp)
+apply(simp)
+done
+
+lemma fresh_star_plus:
+  fixes p q::perm
+  shows "\<lbrakk>a \<sharp>* p;  a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)"
+  unfolding fresh_star_def
+  by (simp add: fresh_plus_perm)
+
+lemma fresh_star_permute_iff:
+  shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x"
+apply(simp add: fresh_star_def)
+apply(simp add: ball_image)
+apply(simp add: fresh_permute_iff)
+done
+
+fun
+  alpha_gen 
+where
+  alpha_gen[simp del]:
+  "alpha_gen (bs, x) R f pi (cs, y) \<longleftrightarrow> f x - bs = f y - cs \<and> (f x - bs) \<sharp>* pi \<and> R (pi \<bullet> x) y"
+
+notation
+  alpha_gen ("_ \<approx>gen _ _ _ _" [100, 100, 100, 100, 100] 100)
+
+lemma [mono]: "R1 \<le> R2 \<Longrightarrow> alpha_gen x R1 \<le> alpha_gen x R2"
+  by (cases x) (auto simp add: le_fun_def le_bool_def alpha_gen.simps)
+
+lemma alpha_gen_refl:
+  assumes a: "R x x"
+  shows "(bs, x) \<approx>gen R f 0 (bs, x)"
+  using a by (simp add: alpha_gen fresh_star_def fresh_zero_perm)
+
+lemma alpha_gen_sym:
+  assumes a: "(bs, x) \<approx>gen R f p (cs, y)"
+  and     b: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x"
+  shows "(cs, y) \<approx>gen R f (- p) (bs, x)"
+  using a b by (simp add: alpha_gen fresh_star_def fresh_def supp_minus_perm)
+
+lemma alpha_gen_trans:
+  assumes a: "(bs, x) \<approx>gen R f p1 (cs, y)"
+  and     b: "(cs, y) \<approx>gen R f p2 (ds, z)"
+  and     c: "\<lbrakk>R (p1 \<bullet> x) y; R (p2 \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((p2 + p1) \<bullet> x) z"
+  shows "(bs, x) \<approx>gen R f (p2 + p1) (ds, z)"
+  using a b c using supp_plus_perm
+  apply(simp add: alpha_gen fresh_star_def fresh_def)
+  apply(blast)
+  done
+
+lemma alpha_gen_eqvt:
+  assumes a: "(bs, x) \<approx>gen R f q (cs, y)"
+  and     b: "R (q \<bullet> x) y \<Longrightarrow> R (p \<bullet> (q \<bullet> x)) (p \<bullet> y)"
+  and     c: "p \<bullet> (f x) = f (p \<bullet> x)"
+  and     d: "p \<bullet> (f y) = f (p \<bullet> y)"
+  shows "(p \<bullet> bs, p \<bullet> x) \<approx>gen R f (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
+  using a b
+  apply(simp add: alpha_gen c[symmetric] d[symmetric] Diff_eqvt[symmetric])
+  apply(simp add: permute_eqvt[symmetric])
+  apply(simp add: fresh_star_permute_iff)
+  apply(clarsimp)
+  done
+
+lemma alpha_gen_compose_sym:
+  assumes b: "\<exists>pi. (aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)"
+  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
+  shows "\<exists>pi. (ab, s) \<approx>gen R f pi (aa, t)"
+  using b apply -
+  apply(erule exE)
+  apply(rule_tac x="- pi" in exI)
+  apply(simp add: alpha_gen.simps)
+  apply(erule conjE)+
+  apply(rule conjI)
+  apply(simp add: fresh_star_def fresh_minus_perm)
+  apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))")
+  apply simp
+  apply(rule a)
+  apply assumption
+  done
+
+lemma alpha_gen_compose_trans:
+  assumes b: "\<exists>pi\<Colon>perm. (aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)"
+  and c: "\<exists>pi\<Colon>perm. (ab, ta) \<approx>gen R f pi (ac, sa)"
+  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
+  shows "\<exists>pi\<Colon>perm. (aa, t) \<approx>gen R f pi (ac, sa)"
+  using b c apply -
+  apply(simp add: alpha_gen.simps)
+  apply(erule conjE)+
+  apply(erule exE)+
+  apply(erule conjE)+
+  apply(rule_tac x="pia + pi" in exI)
+  apply(simp add: fresh_star_plus)
+  apply(drule_tac x="- pia \<bullet> sa" in spec)
+  apply(drule mp)
+  apply(rotate_tac 4)
+  apply(drule_tac pi="- pia" in a)
+  apply(simp)
+  apply(rotate_tac 6)
+  apply(drule_tac pi="pia" in a)
+  apply(simp)
+  done
+
+lemma alpha_gen_atom_eqvt:
+  assumes a: "\<And>x. pi \<bullet> (f x) = f (pi \<bullet> x)"
+  and     b: "\<exists>pia. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R (pi \<bullet> x1) (pi \<bullet> x2)) f pia ({atom b}, s)"
+  shows  "\<exists>pia. ({atom (pi \<bullet> a)}, pi \<bullet> t) \<approx>gen R f pia ({atom (pi \<bullet> b)}, pi \<bullet> s)"
+  using b 
+  apply -
+  apply(erule exE)
+  apply(rule_tac x="pi \<bullet> pia" in exI)
+  apply(simp add: alpha_gen.simps)
+  apply(erule conjE)+
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
+  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt)
+  apply(subst permute_eqvt[symmetric])
+  apply(simp)
+  done
+
+fun
+  alpha_abs 
+where
+  "alpha_abs (bs, x) (cs, y) = (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))"
+
+notation
+  alpha_abs ("_ \<approx>abs _")
+
+lemma alpha_abs_swap:
+  assumes a1: "a \<notin> (supp x) - bs"
+  and     a2: "b \<notin> (supp x) - bs"
+  shows "(bs, x) \<approx>abs ((a \<rightleftharpoons> b) \<bullet> bs, (a \<rightleftharpoons> b) \<bullet> x)"
+  apply(simp)
+  apply(rule_tac x="(a \<rightleftharpoons> b)" in exI)
+  apply(simp add: alpha_gen)
+  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
+  apply(simp add: swap_set_not_in[OF a1 a2])
+  apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
+  using a1 a2
+  apply(simp add: fresh_star_def fresh_def)
+  apply(blast)
+  apply(simp add: supp_swap)
+  done
+
+fun
+  supp_abs_fun
+where
+  "supp_abs_fun (bs, x) = (supp x) - bs"
+
+lemma supp_abs_fun_lemma:
+  assumes a: "x \<approx>abs y" 
+  shows "supp_abs_fun x = supp_abs_fun y"
+  using a
+  apply(induct rule: alpha_abs.induct)
+  apply(simp add: alpha_gen)
+  done
+  
+quotient_type 'a abs = "(atom set \<times> 'a::pt)" / "alpha_abs"
+  apply(rule equivpI)
+  unfolding reflp_def symp_def transp_def
+  apply(simp_all)
+  (* refl *)
+  apply(clarify)
+  apply(rule exI)
+  apply(rule alpha_gen_refl)
+  apply(simp)
+  (* symm *)
+  apply(clarify)
+  apply(rule exI)
+  apply(rule alpha_gen_sym)
+  apply(assumption)
+  apply(clarsimp)
+  (* trans *)
+  apply(clarify)
+  apply(rule exI)
+  apply(rule alpha_gen_trans)
+  apply(assumption)
+  apply(assumption)
+  apply(simp)
+  done
+
+quotient_definition
+  "Abs::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs"
+is
+  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)"
+
+lemma [quot_respect]:
+  shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair"
+  apply(clarsimp)
+  apply(rule exI)
+  apply(rule alpha_gen_refl)
+  apply(simp)
+  done
+
+lemma [quot_respect]:
+  shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute"
+  apply(clarsimp)
+  apply(rule exI)
+  apply(rule alpha_gen_eqvt)
+  apply(assumption)
+  apply(simp_all add: supp_eqvt)
+  done
+
+lemma [quot_respect]:
+  shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun"
+  apply(simp add: supp_abs_fun_lemma)
+  done
+
+lemma abs_induct:
+  "\<lbrakk>\<And>as (x::'a::pt). P (Abs as x)\<rbrakk> \<Longrightarrow> P t"
+  apply(lifting prod.induct[where 'a="atom set" and 'b="'a"])
+  done
+
+(* TEST case *)
+lemmas abs_induct2 = prod.induct[where 'a="atom set" and 'b="'a::pt", quot_lifted]
+thm abs_induct abs_induct2
+
+instantiation abs :: (pt) pt
+begin
+
+quotient_definition
+  "permute_abs::perm \<Rightarrow> ('a::pt abs) \<Rightarrow> 'a abs"
+is
+  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
+
+lemma permute_ABS [simp]:
+  fixes x::"'a::pt"  (* ??? has to be 'a \<dots> 'b does not work *)
+  shows "(p \<bullet> (Abs as x)) = Abs (p \<bullet> as) (p \<bullet> x)"
+  by (lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"])
+
+instance
+  apply(default)
+  apply(induct_tac [!] x rule: abs_induct)
+  apply(simp_all)
+  done
+
+end
+
+quotient_definition
+  "supp_Abs_fun :: ('a::pt) abs \<Rightarrow> atom \<Rightarrow> bool"
+is
+  "supp_abs_fun"
+
+lemma supp_Abs_fun_simp:
+  shows "supp_Abs_fun (Abs bs x) = (supp x) - bs"
+  by (lifting supp_abs_fun.simps(1))
+
+lemma supp_Abs_fun_eqvt [eqvt]:
+  shows "(p \<bullet> supp_Abs_fun x) = supp_Abs_fun (p \<bullet> x)"
+  apply(induct_tac x rule: abs_induct)
+  apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt)
+  done
+
+lemma supp_Abs_fun_fresh:
+  shows "a \<sharp> Abs bs x \<Longrightarrow> a \<sharp> supp_Abs_fun (Abs bs x)"
+  apply(rule fresh_fun_eqvt_app)
+  apply(simp add: eqvts_raw)
+  apply(simp)
+  done
+
+lemma Abs_swap:
+  assumes a1: "a \<notin> (supp x) - bs"
+  and     a2: "b \<notin> (supp x) - bs"
+  shows "(Abs bs x) = (Abs ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x))"
+  using a1 a2 by (lifting alpha_abs_swap)
+
+lemma Abs_supports:
+  shows "((supp x) - as) supports (Abs as x)"
+  unfolding supports_def
+  apply(clarify)
+  apply(simp (no_asm))
+  apply(subst Abs_swap[symmetric])
+  apply(simp_all)
+  done
+
+lemma supp_Abs_subset1:
+  fixes x::"'a::fs"
+  shows "(supp x) - as \<subseteq> supp (Abs as x)"
+  apply(simp add: supp_conv_fresh)
+  apply(auto)
+  apply(drule_tac supp_Abs_fun_fresh)
+  apply(simp only: supp_Abs_fun_simp)
+  apply(simp add: fresh_def)
+  apply(simp add: supp_finite_atom_set finite_supp)
+  done
+
+lemma supp_Abs_subset2:
+  fixes x::"'a::fs"
+  shows "supp (Abs as x) \<subseteq> (supp x) - as"
+  apply(rule supp_is_subset)
+  apply(rule Abs_supports)
+  apply(simp add: finite_supp)
+  done
+
+lemma supp_Abs:
+  fixes x::"'a::fs"
+  shows "supp (Abs as x) = (supp x) - as"
+  apply(rule_tac subset_antisym)
+  apply(rule supp_Abs_subset2)
+  apply(rule supp_Abs_subset1)
+  done
+
+instance abs :: (fs) fs
+  apply(default)
+  apply(induct_tac x rule: abs_induct)
+  apply(simp add: supp_Abs)
+  apply(simp add: finite_supp)
+  done
+
+lemma Abs_fresh_iff:
+  fixes x::"'a::fs"
+  shows "a \<sharp> Abs bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
+  apply(simp add: fresh_def)
+  apply(simp add: supp_Abs)
+  apply(auto)
+  done
+
+lemma Abs_eq_iff:
+  shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
+  by (lifting alpha_abs.simps(1))
+
+
+
+(* 
+  below is a construction site for showing that in the
+  single-binder case, the old and new alpha equivalence 
+  coincide
+*)
+
+fun
+  alpha1
+where
+  "alpha1 (a, x) (b, y) \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<rightleftharpoons> b) \<bullet> y \<and> a \<sharp> y)"
+
+notation 
+  alpha1 ("_ \<approx>abs1 _")
+
+thm swap_set_not_in
+
+lemma qq:
+  fixes S::"atom set"
+  assumes a: "supp p \<inter> S = {}"
+  shows "p \<bullet> S = S"
+using a
+apply(simp add: supp_perm permute_set_eq)
+apply(auto)
+apply(simp only: disjoint_iff_not_equal)
+apply(simp)
+apply (metis permute_atom_def_raw)
+apply(rule_tac x="(- p) \<bullet> x" in exI)
+apply(simp)
+apply(simp only: disjoint_iff_not_equal)
+apply(simp)
+apply(metis permute_minus_cancel)
+done
+
+lemma alpha_abs_swap:
+  assumes a1: "(supp x - bs) \<sharp>* p"
+  and     a2: "(supp x - bs) \<sharp>* p"
+  shows "(bs, x) \<approx>abs (p \<bullet> bs, p \<bullet> x)"
+  apply(simp)
+  apply(rule_tac x="p" in exI)
+  apply(simp add: alpha_gen)
+  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
+  apply(rule conjI)
+  apply(rule sym)
+  apply(rule qq)
+  using a1 a2
+  apply(auto)[1]
+  oops
+
+
+
+lemma
+  assumes a: "(a, x) \<approx>abs1 (b, y)" "sort_of a = sort_of b"
+  shows "({a}, x) \<approx>abs ({b}, y)"
+using a
+apply(simp)
+apply(erule disjE)
+apply(simp)
+apply(rule exI)
+apply(rule alpha_gen_refl)
+apply(simp)
+apply(rule_tac x="(a \<rightleftharpoons> b)" in  exI)
+apply(simp add: alpha_gen)
+apply(simp add: fresh_def)
+apply(rule conjI)
+apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in  permute_eq_iff[THEN iffD1])
+apply(rule trans)
+apply(simp add: Diff_eqvt supp_eqvt)
+apply(subst swap_set_not_in)
+back
+apply(simp)
+apply(simp)
+apply(simp add: permute_set_eq)
+apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in fresh_star_permute_iff[THEN iffD1])
+apply(simp add: permute_self)
+apply(simp add: Diff_eqvt supp_eqvt)
+apply(simp add: permute_set_eq)
+apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
+apply(simp add: fresh_star_def fresh_def)
+apply(blast)
+apply(simp add: supp_swap)
+done
+
+thm supp_perm
+
+lemma perm_induct_test:
+  fixes P :: "perm => bool"
+  assumes zero: "P 0"
+  assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
+  assumes plus: "\<And>p1 p2. \<lbrakk>supp (p1 + p2) = (supp p1 \<union> supp p2); P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)"
+  shows "P p"
+sorry
+
+lemma tt1:
+  assumes a: "finite (supp p)"
+  shows "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
+using a
+unfolding fresh_star_def fresh_def
+apply(induct F\<equiv>"supp p" arbitrary: p rule: finite.induct)
+apply(simp add: supp_perm)
+defer
+apply(case_tac "a \<in> A")
+apply(simp add: insert_absorb)
+apply(subgoal_tac "A = supp p - {a}")
+prefer 2
+apply(blast)
+apply(case_tac "p \<bullet> a = a")
+apply(simp add: supp_perm)
+apply(drule_tac x="p + (((- p) \<bullet> a) \<rightleftharpoons> a)" in meta_spec)
+apply(simp)
+apply(drule meta_mp)
+apply(rule subset_antisym)
+apply(rule subsetI)
+apply(simp)
+apply(simp add: supp_perm)
+apply(case_tac "xa = p \<bullet> a")
+apply(simp)
+apply(case_tac "p \<bullet> a = (- p) \<bullet> a")
+apply(simp)
+defer
+apply(simp)
+oops
+
+lemma tt:
+  "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
+apply(induct p rule: perm_induct_test)
+apply(simp)
+apply(rule swap_fresh_fresh)
+apply(case_tac "a \<in> supp x")
+apply(simp add: fresh_star_def)
+apply(drule_tac x="a" in bspec)
+apply(simp)
+apply(simp add: fresh_def)
+apply(simp add: supp_swap)
+apply(simp add: fresh_def)
+apply(case_tac "b \<in> supp x")
+apply(simp add: fresh_star_def)
+apply(drule_tac x="b" in bspec)
+apply(simp)
+apply(simp add: fresh_def)
+apply(simp add: supp_swap)
+apply(simp add: fresh_def)
+apply(simp)
+apply(drule meta_mp)
+apply(simp add: fresh_star_def fresh_def)
+apply(drule meta_mp)
+apply(simp add: fresh_star_def fresh_def)
+apply(simp)
+done
+
+lemma yy:
+  assumes "S1 - {x} = S2 - {x}" "x \<in> S1" "x \<in> S2"
+  shows "S1 = S2"
+using assms
+apply (metis insert_Diff_single insert_absorb)
+done
+
+
+lemma
+  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of b"
+  shows "(a, x) \<approx>abs1 (b, y)"
+using a
+apply(case_tac "a = b")
+apply(simp)
+oops
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Fv.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,399 @@
+theory Fv
+imports "Nominal2_Atoms" "Abs"
+begin
+
+(* Bindings are given as a list which has a length being equal
+   to the length of the number of constructors.
+
+   Each element is a list whose length is equal to the number
+   of arguents.
+
+   Every element specifies bindings of this argument given as
+   a tuple: function, bound argument.
+
+  Eg:
+nominal_datatype
+
+   C1
+ | C2 x y z bind x in z
+ | C3 x y z bind f x in z bind g y in z
+
+yields:
+[
+ [],
+ [[], [], [(NONE, 0)]],
+ [[], [], [(SOME (Const f), 0), (Some (Const g), 1)]]]
+
+A SOME binding has to have a function returning an atom set,
+and a NONE binding has to be on an argument that is an atom
+or an atom set.
+
+How the procedure works:
+  For each of the defined datatypes,
+  For each of the constructors,
+  It creates a union of free variables for each argument.
+
+  For an argument the free variables are the variables minus
+  bound variables.
+
+  The variables are:
+    For an atom, a singleton set with the atom itself.
+    For an atom set, the atom set itself.
+    For a recursive argument, the appropriate fv function applied to it.
+    (* TODO: This one is not implemented *)
+    For other arguments it should be an appropriate fv function stored
+      in the database.
+  The bound variables are a union of results of all bindings that
+  involve the given argument. For a paricular binding the result is:
+    For a function applied to an argument this function with the argument.
+    For an atom, a singleton set with the atom itself.
+    For an atom set, the atom set itself.
+    For a recursive argument, the appropriate fv function applied to it.
+    (* TODO: This one is not implemented *)
+    For other arguments it should be an appropriate fv function stored
+      in the database.
+*)
+
+ML {*
+  open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *);
+  (* TODO: It is the same as one in 'nominal_atoms' *)
+  fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom});
+  val noatoms = @{term "{} :: atom set"};
+  fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x];
+  fun mk_union sets =
+    fold (fn a => fn b =>
+      if a = noatoms then b else
+      if b = noatoms then a else
+      HOLogic.mk_binop @{const_name union} (a, b)) (rev sets) noatoms;
+  fun mk_diff a b =
+    if b = noatoms then a else
+    if b = a then noatoms else
+    HOLogic.mk_binop @{const_name minus} (a, b);
+  fun mk_atoms t =
+    let
+      val ty = fastype_of t;
+      val atom_ty = HOLogic.dest_setT ty --> @{typ atom};
+      val img_ty = atom_ty --> ty --> @{typ "atom set"};
+    in
+      (Const (@{const_name image}, img_ty) $ Const (@{const_name atom}, atom_ty) $ t)
+    end;
+  (* Copy from Term *)
+  fun is_funtype (Type ("fun", [_, _])) = true
+    | is_funtype _ = false;
+  (* Similar to one in USyntax *)
+  fun mk_pair (fst, snd) =
+    let val ty1 = fastype_of fst
+      val ty2 = fastype_of snd
+      val c = HOLogic.pair_const ty1 ty2
+    in c $ fst $ snd
+    end;
+
+*}
+
+(* TODO: Notice datatypes without bindings and replace alpha with equality *)
+ML {*
+(* Currently needs just one full_tname to access Datatype *)
+fun define_fv_alpha full_tname bindsall lthy =
+let
+  val thy = ProofContext.theory_of lthy;
+  val {descr, ...} = Datatype.the_info thy full_tname;
+  val sorts = []; (* TODO *)
+  fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
+  val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
+    "fv_" ^ name_of_typ (nth_dtyp i)) descr);
+  val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr;
+  val fv_frees = map Free (fv_names ~~ fv_types);
+  val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
+    "alpha_" ^ name_of_typ (nth_dtyp i)) descr);
+  val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr;
+  val alpha_frees = map Free (alpha_names ~~ alpha_types);
+  fun fv_alpha_constr i (cname, dts) bindcs =
+    let
+      val Ts = map (typ_of_dtyp descr sorts) dts;
+      val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
+      val args = map Free (names ~~ Ts);
+      val names2 = Name.variant_list ("pi" :: names) (Datatype_Prop.make_tnames Ts);
+      val args2 = map Free (names2 ~~ Ts);
+      val c = Const (cname, Ts ---> (nth_dtyp i));
+      val fv_c = nth fv_frees i;
+      val alpha = nth alpha_frees i;
+      fun fv_bind args (NONE, i) =
+            if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else
+            (* TODO we assume that all can be 'atomized' *)
+            if (is_funtype o fastype_of) (nth args i) then mk_atoms (nth args i) else
+            mk_single_atom (nth args i)
+        | fv_bind args (SOME f, i) = f $ (nth args i);
+      fun fv_arg ((dt, x), bindxs) =
+        let
+          val arg =
+            if is_rec_type dt then nth fv_frees (body_index dt) $ x else
+            (* TODO: we just assume everything can be 'atomized' *)
+            if (is_funtype o fastype_of) x then mk_atoms x else
+            HOLogic.mk_set @{typ atom} [mk_atom (fastype_of x) $ x]
+          val sub = mk_union (map (fv_bind args) bindxs)
+        in
+          mk_diff arg sub
+        end;
+      val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq
+        (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ bindcs))))
+      val alpha_rhs =
+        HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2)));
+      fun alpha_arg ((dt, bindxs), (arg, arg2)) =
+        if bindxs = [] then (
+          if is_rec_type dt then (nth alpha_frees (body_index dt) $ arg $ arg2)
+          else (HOLogic.mk_eq (arg, arg2)))
+        else
+          if is_rec_type dt then let
+            (* THE HARD CASE *)
+            val lhs_binds = mk_union (map (fv_bind args) bindxs);
+            val lhs = mk_pair (lhs_binds, arg);
+            val rhs_binds = mk_union (map (fv_bind args2) bindxs);
+            val rhs = mk_pair (rhs_binds, arg2);
+            val alpha = nth alpha_frees (body_index dt);
+            val fv = nth fv_frees (body_index dt);
+            val alpha_gen_pre = Const (@{const_name alpha_gen}, dummyT) $ lhs $ alpha $ fv $ (Free ("pi", @{typ perm})) $ rhs;
+            val alpha_gen_t = Syntax.check_term lthy alpha_gen_pre
+          in
+            HOLogic.mk_exists ("pi", @{typ perm}, alpha_gen_t)
+          (* TODO Add some test that is makes sense *)
+          end else @{term "True"}
+      val alpha_lhss = map (HOLogic.mk_Trueprop o alpha_arg) (dts ~~ bindcs ~~ (args ~~ args2))
+      val alpha_eq = Logic.list_implies (alpha_lhss, alpha_rhs)
+    in
+      (fv_eq, alpha_eq)
+    end;
+  fun fv_alpha_eq (i, (_, _, constrs)) binds = map2 (fv_alpha_constr i) constrs binds;
+  val (fv_eqs, alpha_eqs) = split_list (flat (map2 fv_alpha_eq descr bindsall))
+  val add_binds = map (fn x => (Attrib.empty_binding, x))
+  val (fvs, lthy') = (Primrec.add_primrec
+    (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) (add_binds fv_eqs) lthy)
+  val (alphas, lthy'') = (Inductive.add_inductive_i
+     {quiet_mode = false, verbose = true, alt_name = Binding.empty,
+      coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false}
+     (map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) alpha_names alpha_types) []
+     (add_binds alpha_eqs) [] lthy')
+in
+  ((fvs, alphas), lthy'')
+end
+*}
+
+(* tests
+atom_decl name
+
+datatype ty =
+  Var "name set"
+
+ML {* Syntax.check_term @{context} (mk_atoms @{term "a :: name set"}) *}
+
+local_setup {* define_fv_alpha "Fv.ty" [[[[]]]] *}
+print_theorems
+
+
+datatype rtrm1 =
+  rVr1 "name"
+| rAp1 "rtrm1" "rtrm1"
+| rLm1 "name" "rtrm1"        --"name is bound in trm1"
+| rLt1 "bp" "rtrm1" "rtrm1"   --"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) = {atom x}"
+| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)"
+
+setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Fv.rtrm1", "Fv.bp"] *}
+
+local_setup {* define_fv_alpha "Fv.rtrm1"
+  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(SOME @{term bv1}, 0)], [], [(SOME @{term bv1}, 0)]]],
+   [[], [[]], [[], []]]] *}
+print_theorems
+*)
+
+
+ML {*
+fun alpha_inj_tac dist_inj intrs elims =
+  SOLVED' (asm_full_simp_tac (HOL_ss addsimps intrs)) ORELSE'
+  (rtac @{thm iffI} THEN' RANGE [
+     (eresolve_tac elims THEN_ALL_NEW
+       asm_full_simp_tac (HOL_ss addsimps dist_inj)
+     ),
+     asm_full_simp_tac (HOL_ss addsimps intrs)])
+*}
+
+ML {*
+fun build_alpha_inj_gl thm =
+  let
+    val prop = prop_of thm;
+    val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop);
+    val hyps = map HOLogic.dest_Trueprop (Logic.strip_imp_prems prop);
+    fun list_conj l = foldr1 HOLogic.mk_conj l;
+  in
+    if hyps = [] then concl
+    else HOLogic.mk_eq (concl, list_conj hyps)
+  end;
+*}
+
+ML {*
+fun build_alpha_inj intrs dist_inj elims ctxt =
+let
+  val ((_, thms_imp), ctxt') = Variable.import false intrs ctxt;
+  val gls = map (HOLogic.mk_Trueprop o build_alpha_inj_gl) thms_imp;
+  fun tac _ = alpha_inj_tac dist_inj intrs elims 1;
+  val thms = map (fn gl => Goal.prove ctxt' [] [] gl tac) gls;
+in
+  Variable.export ctxt' ctxt thms
+end
+*}
+
+ML {*
+fun build_alpha_refl_gl alphas (x, y, z) =
+let
+  fun build_alpha alpha =
+    let
+      val ty = domain_type (fastype_of alpha);
+      val var = Free(x, ty);
+      val var2 = Free(y, ty);
+      val var3 = Free(z, ty);
+      val symp = HOLogic.mk_imp (alpha $ var $ var2, alpha $ var2 $ var);
+      val transp = HOLogic.mk_imp (alpha $ var $ var2,
+        HOLogic.mk_all (z, ty,
+          HOLogic.mk_imp (alpha $ var2 $ var3, alpha $ var $ var3)))
+    in
+      ((alpha $ var $ var), (symp, transp))
+    end;
+  val (refl_eqs, eqs) = split_list (map build_alpha alphas)
+  val (sym_eqs, trans_eqs) = split_list eqs
+  fun conj l = @{term Trueprop} $ foldr1 HOLogic.mk_conj l
+in
+  (conj refl_eqs, (conj sym_eqs, conj trans_eqs))
+end
+*}
+
+ML {*
+fun reflp_tac induct inj =
+  rtac induct THEN_ALL_NEW
+  asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
+  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
+  (rtac @{thm exI[of _ "0 :: perm"]} THEN'
+   asm_full_simp_tac (HOL_ss addsimps
+     @{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv}))
+*}
+
+ML {*
+fun symp_tac induct inj eqvt =
+  ((rtac @{thm impI} THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
+  asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
+  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
+  (etac @{thm alpha_gen_compose_sym} THEN' eresolve_tac eqvt)
+*}
+
+ML {*
+fun imp_elim_tac case_rules =
+  Subgoal.FOCUS (fn {concl, context, ...} =>
+    case term_of concl of
+      _ $ (_ $ asm $ _) =>
+        let
+          fun filter_fn case_rule = (
+            case Logic.strip_assums_hyp (prop_of case_rule) of
+              ((_ $ asmc) :: _) =>
+                let
+                  val thy = ProofContext.theory_of context
+                in
+                  Pattern.matches thy (asmc, asm)
+                end
+            | _ => false)
+          val matching_rules = filter filter_fn case_rules
+        in
+         (rtac impI THEN' rotate_tac (~1) THEN' eresolve_tac matching_rules) 1
+        end
+    | _ => no_tac
+  )
+*}
+
+ML {*
+fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt =
+  ((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
+  (TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW
+  (
+    asm_full_simp_tac (HOL_ss addsimps alpha_inj @ term_inj @ distinct) THEN'
+    TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
+    (etac @{thm alpha_gen_compose_trans} THEN' RANGE [atac, eresolve_tac eqvt])
+  )
+*}
+
+lemma transp_aux:
+  "(\<And>xa ya. R xa ya \<longrightarrow> (\<forall>z. R ya z \<longrightarrow> R xa z)) \<Longrightarrow> transp R"
+  unfolding transp_def
+  by blast
+
+ML {*
+fun equivp_tac reflps symps transps =
+  simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def})
+  THEN' rtac conjI THEN' rtac allI THEN'
+  resolve_tac reflps THEN'
+  rtac conjI THEN' rtac allI THEN' rtac allI THEN'
+  resolve_tac symps THEN'
+  rtac @{thm transp_aux} THEN' resolve_tac transps
+*}
+
+ML {*
+fun build_equivps alphas term_induct alpha_induct term_inj alpha_inj distinct cases eqvt ctxt =
+let
+  val ([x, y, z], ctxt') = Variable.variant_fixes ["x","y","z"] ctxt;
+  val (reflg, (symg, transg)) = build_alpha_refl_gl alphas (x, y, z)
+  fun reflp_tac' _ = reflp_tac term_induct alpha_inj 1;
+  fun symp_tac' _ = symp_tac alpha_induct alpha_inj eqvt 1;
+  fun transp_tac' _ = transp_tac ctxt alpha_induct alpha_inj term_inj distinct cases eqvt 1;
+  val reflt = Goal.prove ctxt' [] [] reflg reflp_tac';
+  val symt = Goal.prove ctxt' [] [] symg symp_tac';
+  val transt = Goal.prove ctxt' [] [] transg transp_tac';
+  val [refltg, symtg, transtg] = Variable.export ctxt' ctxt [reflt, symt, transt]
+  val reflts = HOLogic.conj_elims refltg
+  val symts = HOLogic.conj_elims symtg
+  val transts = HOLogic.conj_elims transtg
+  fun equivp alpha =
+    let
+      val equivp = Const (@{const_name equivp}, fastype_of alpha --> @{typ bool})
+      val goal = @{term Trueprop} $ (equivp $ alpha)
+      fun tac _ = equivp_tac reflts symts transts 1
+    in
+      Goal.prove ctxt [] [] goal tac
+    end
+in
+  map equivp alphas
+end
+*}
+
+(*
+Tests:
+prove alpha1_reflp_aux: {* fst (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *}
+by (tactic {* reflp_tac @{thm rtrm1_bp.induct} @{thms alpha1_inj} 1 *})
+
+prove alpha1_symp_aux: {* (fst o snd) (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *}
+by (tactic {* symp_tac @{thm alpha_rtrm1_alpha_bp.induct} @{thms alpha1_inj} @{thms alpha1_eqvt} 1 *})
+
+prove alpha1_transp_aux: {* (snd o snd) (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *}
+by (tactic {* transp_tac @{context} @{thm alpha_rtrm1_alpha_bp.induct} @{thms alpha1_inj} @{thms rtrm1.inject bp.inject} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} 1 *})
+
+lemma alpha1_equivp:
+  "equivp alpha_rtrm1"
+  "equivp alpha_bp"
+apply (tactic {*
+  (simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def})
+  THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN'
+  resolve_tac (HOLogic.conj_elims @{thm alpha1_reflp_aux})
+  THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN'
+  resolve_tac (HOLogic.conj_elims @{thm alpha1_symp_aux}) THEN' rtac @{thm transp_aux}
+  THEN' resolve_tac (HOLogic.conj_elims @{thm alpha1_transp_aux})
+)
+1 *})
+done*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/LFex.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,236 @@
+theory LFex
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp"
+begin
+
+atom_decl name
+atom_decl ident
+
+datatype rkind =
+    Type
+  | KPi "rty" "name" "rkind"
+and rty =
+    TConst "ident"
+  | TApp "rty" "rtrm"
+  | TPi "rty" "name" "rty"
+and rtrm =
+    Const "ident"
+  | Var "name"
+  | App "rtrm" "rtrm"
+  | Lam "rty" "name" "rtrm"
+
+
+setup {* snd o define_raw_perms ["rkind", "rty", "rtrm"] ["LFex.rkind", "LFex.rty", "LFex.rtrm"] *}
+
+local_setup {*
+  snd o define_fv_alpha "LFex.rkind"
+  [[ [], [[], [(NONE, 1)], [(NONE, 1)]] ],
+   [ [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]] ],
+   [ [[]], [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]]]] *}
+notation
+    alpha_rkind  ("_ \<approx>ki _" [100, 100] 100)
+and alpha_rty    ("_ \<approx>ty _" [100, 100] 100)
+and alpha_rtrm   ("_ \<approx>tr _" [100, 100] 100)
+thm fv_rkind_fv_rty_fv_rtrm.simps alpha_rkind_alpha_rty_alpha_rtrm.intros
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_rkind_alpha_rty_alpha_rtrm_inj}, []), (build_alpha_inj @{thms alpha_rkind_alpha_rty_alpha_rtrm.intros} @{thms rkind.distinct rty.distinct rtrm.distinct rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} ctxt)) ctxt)) *}
+thm alpha_rkind_alpha_rty_alpha_rtrm_inj
+
+lemma rfv_eqvt[eqvt]:
+  "((pi\<bullet>fv_rkind t1) = fv_rkind (pi\<bullet>t1))"
+  "((pi\<bullet>fv_rty t2) = fv_rty (pi\<bullet>t2))"
+  "((pi\<bullet>fv_rtrm t3) = fv_rtrm (pi\<bullet>t3))"
+apply(induct t1 and t2 and t3 rule: rkind_rty_rtrm.inducts)
+apply(simp_all add: union_eqvt Diff_eqvt)
+apply(simp_all add: permute_set_eq atom_eqvt)
+done
+
+lemma alpha_eqvt:
+  "t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)"
+  "t2 \<approx>ty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>ty (pi \<bullet> s2)"
+  "t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)"
+apply(induct rule: alpha_rkind_alpha_rty_alpha_rtrm.inducts)
+apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm.intros)
+apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm_inj)
+apply (rule alpha_gen_atom_eqvt)
+apply (simp add: rfv_eqvt)
+apply assumption
+apply (rule alpha_gen_atom_eqvt)
+apply (simp add: rfv_eqvt)
+apply assumption
+apply (rule alpha_gen_atom_eqvt)
+apply (simp add: rfv_eqvt)
+apply assumption
+done
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_equivps}, []),
+  (build_equivps [@{term alpha_rkind}, @{term alpha_rty}, @{term alpha_rtrm}]
+     @{thm rkind_rty_rtrm.induct} @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} 
+     @{thms rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj}
+     @{thms rkind.distinct rty.distinct rtrm.distinct}
+     @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases}
+     @{thms alpha_eqvt} ctxt)) ctxt)) *}
+thm alpha_equivps
+
+local_setup  {* define_quotient_type
+  [(([], @{binding kind}, NoSyn), (@{typ rkind}, @{term alpha_rkind})),
+   (([], @{binding ty},   NoSyn), (@{typ rty},   @{term alpha_rty}  )),
+   (([], @{binding trm},  NoSyn), (@{typ rtrm},  @{term alpha_rtrm} ))]
+  (ALLGOALS (resolve_tac @{thms alpha_equivps}))
+*}
+
+local_setup {*
+(fn ctxt => ctxt
+ |> snd o (Quotient_Def.quotient_lift_const ("TYP", @{term Type}))
+ |> snd o (Quotient_Def.quotient_lift_const ("KPI", @{term KPi}))
+ |> snd o (Quotient_Def.quotient_lift_const ("TCONST", @{term TConst}))
+ |> snd o (Quotient_Def.quotient_lift_const ("TAPP", @{term TApp}))
+ |> snd o (Quotient_Def.quotient_lift_const ("TPI", @{term TPi}))
+ |> snd o (Quotient_Def.quotient_lift_const ("CONS", @{term Const}))
+ |> snd o (Quotient_Def.quotient_lift_const ("VAR", @{term Var}))
+ |> snd o (Quotient_Def.quotient_lift_const ("APP", @{term App}))
+ |> snd o (Quotient_Def.quotient_lift_const ("LAM", @{term Lam}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_kind", @{term fv_rkind}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_ty", @{term fv_rty}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_trm", @{term fv_rtrm}))) *}
+print_theorems
+
+local_setup {* prove_const_rsp @{binding rfv_rsp} [@{term fv_rkind}, @{term fv_rty}, @{term fv_rtrm}]
+  (fn _ => fvbv_rsp_tac @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} @{thms fv_rkind_fv_rty_fv_rtrm.simps} 1) *}
+local_setup {* prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"}]
+  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *}
+local_setup {* prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rty \<Rightarrow> rty"}]
+  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *}
+local_setup {* prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"}]
+  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *}
+ML {* fun const_rsp_tac _ = constr_rsp_tac @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj}
+  @{thms rfv_rsp} @{thms alpha_equivps} 1 *}
+local_setup {* prove_const_rsp Binding.empty [@{term TConst}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term TApp}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term Var}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term App}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term Const}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term KPi}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term TPi}] const_rsp_tac *}
+local_setup {* prove_const_rsp Binding.empty [@{term Lam}] const_rsp_tac *}
+
+lemmas kind_ty_trm_induct = rkind_rty_rtrm.induct[quot_lifted]
+
+thm rkind_rty_rtrm.inducts
+lemmas kind_ty_trm_inducts = rkind_rty_rtrm.inducts[quot_lifted]
+
+instantiation kind and ty and trm :: pt
+begin
+
+quotient_definition
+  "permute_kind :: perm \<Rightarrow> kind \<Rightarrow> kind"
+is
+  "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"
+
+quotient_definition
+  "permute_ty :: perm \<Rightarrow> ty \<Rightarrow> ty"
+is
+  "permute :: perm \<Rightarrow> rty \<Rightarrow> rty"
+
+quotient_definition
+  "permute_trm :: perm \<Rightarrow> trm \<Rightarrow> trm"
+is
+  "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"
+
+instance by default (simp_all add:
+  permute_rkind_permute_rty_permute_rtrm_zero[quot_lifted]
+  permute_rkind_permute_rty_permute_rtrm_append[quot_lifted])
+
+end
+
+(*
+Lifts, but slow and not needed?.
+lemmas alpha_kind_alpha_ty_alpha_trm_induct = alpha_rkind_alpha_rty_alpha_rtrm.induct[unfolded alpha_gen, quot_lifted, folded alpha_gen]
+*)
+
+lemmas permute_ktt[simp] = permute_rkind_permute_rty_permute_rtrm.simps[quot_lifted]
+
+lemmas kind_ty_trm_inj = alpha_rkind_alpha_rty_alpha_rtrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
+
+lemmas fv_kind_ty_trm = fv_rkind_fv_rty_fv_rtrm.simps[quot_lifted]
+
+lemmas fv_eqvt = rfv_eqvt[quot_lifted]
+
+lemma supports:
+  "{} supports TYP"
+  "(supp (atom i)) supports (TCONST i)"
+  "(supp A \<union> supp M) supports (TAPP A M)"
+  "(supp (atom i)) supports (CONS i)"
+  "(supp (atom x)) supports (VAR x)"
+  "(supp M \<union> supp N) supports (APP M N)"
+  "(supp ty \<union> supp (atom na) \<union> supp ki) supports (KPI ty na ki)"
+  "(supp ty \<union> supp (atom na) \<union> supp ty2) supports (TPI ty na ty2)"
+  "(supp ty \<union> supp (atom na) \<union> supp trm) supports (LAM ty na trm)"
+apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh)
+apply(rule_tac [!] allI)+
+apply(rule_tac [!] impI)
+apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
+apply(simp_all add: fresh_atom)
+done
+
+lemma kind_ty_trm_fs:
+  "finite (supp (x\<Colon>kind))"
+  "finite (supp (y\<Colon>ty))"
+  "finite (supp (z\<Colon>trm))"
+apply(induct x and y and z rule: kind_ty_trm_inducts)
+apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
+apply(simp_all add: supp_atom)
+done
+
+instance kind and ty and trm :: fs
+apply(default)
+apply(simp_all only: kind_ty_trm_fs)
+done
+
+lemma supp_eqs:
+  "supp TYP = {}"
+  "supp rkind = fv_kind rkind \<Longrightarrow> supp (KPI rty name rkind) = supp rty \<union> supp (Abs {atom name} rkind)"
+  "supp (TCONST i) = {atom i}"
+  "supp (TAPP A M) = supp A \<union> supp M"
+  "supp rty2 = fv_ty rty2 \<Longrightarrow> supp (TPI rty1 name rty2) = supp rty1 \<union> supp (Abs {atom name} rty2)"
+  "supp (CONS i) = {atom i}"
+  "supp (VAR x) = {atom x}"
+  "supp (APP M N) = supp M \<union> supp N"
+  "supp rtrm = fv_trm rtrm \<Longrightarrow> supp (LAM rty name rtrm) = supp rty \<union> supp (Abs {atom name} rtrm)"
+  apply(simp_all (no_asm) add: supp_def)
+  apply(simp_all only: kind_ty_trm_inj Abs_eq_iff alpha_gen)
+  apply(simp_all only: insert_eqvt empty_eqvt atom_eqvt supp_eqvt[symmetric] fv_eqvt[symmetric])
+  apply(simp_all add: Collect_imp_eq Collect_neg_eq[symmetric] Set.Un_commute)
+  apply(simp_all add: supp_at_base[simplified supp_def])
+  done
+
+lemma supp_fv:
+  "supp t1 = fv_kind t1"
+  "supp t2 = fv_ty t2"
+  "supp t3 = fv_trm t3"
+  apply(induct t1 and t2 and t3 rule: kind_ty_trm_inducts)
+  apply(simp_all (no_asm) only: supp_eqs fv_kind_ty_trm)
+  apply(simp_all)
+  apply(subst supp_eqs)
+  apply(simp_all add: supp_Abs)
+  apply(subst supp_eqs)
+  apply(simp_all add: supp_Abs)
+  apply(subst supp_eqs)
+  apply(simp_all add: supp_Abs)
+  done
+
+lemma supp_rkind_rty_rtrm:
+  "supp TYP = {}"
+  "supp (KPI A x K) = supp A \<union> (supp K - {atom x})"
+  "supp (TCONST i) = {atom i}"
+  "supp (TAPP A M) = supp A \<union> supp M"
+  "supp (TPI A x B) = supp A \<union> (supp B - {atom x})"
+  "supp (CONS i) = {atom i}"
+  "supp (VAR x) = {atom x}"
+  "supp (APP M N) = supp M \<union> supp N"
+  "supp (LAM A x M) = supp A \<union> (supp M - {atom x})"
+  by (simp_all only: supp_fv fv_kind_ty_trm)
+
+end
+
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/LamEx.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,624 @@
+theory LamEx
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
+begin
+
+atom_decl name
+
+datatype rlam =
+  rVar "name"
+| rApp "rlam" "rlam"
+| rLam "name" "rlam"
+
+fun
+  rfv :: "rlam \<Rightarrow> atom set"
+where
+  rfv_var: "rfv (rVar a) = {atom a}"
+| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
+| rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}"
+
+instantiation rlam :: pt
+begin
+
+primrec
+  permute_rlam
+where
+  "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)"
+| "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)"
+| "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)"
+
+instance
+apply default
+apply(induct_tac [!] x)
+apply(simp_all)
+done
+
+end
+
+instantiation rlam :: fs
+begin
+
+lemma neg_conj:
+  "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)"
+  by simp
+
+lemma infinite_Un:
+  "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
+  by simp
+
+instance
+apply default
+apply(induct_tac x)
+(* var case *)
+apply(simp add: supp_def)
+apply(fold supp_def)[1]
+apply(simp add: supp_at_base)
+(* app case *)
+apply(simp only: supp_def)
+apply(simp only: permute_rlam.simps) 
+apply(simp only: rlam.inject) 
+apply(simp only: neg_conj)
+apply(simp only: Collect_disj_eq)
+apply(simp only: infinite_Un)
+apply(simp only: Collect_disj_eq)
+apply(simp)
+(* lam case *)
+apply(simp only: supp_def)
+apply(simp only: permute_rlam.simps) 
+apply(simp only: rlam.inject) 
+apply(simp only: neg_conj)
+apply(simp only: Collect_disj_eq)
+apply(simp only: infinite_Un)
+apply(simp only: Collect_disj_eq)
+apply(simp)
+apply(fold supp_def)[1]
+apply(simp add: supp_at_base)
+done
+
+end
+
+
+(* for the eqvt proof of the alpha-equivalence *)
+declare permute_rlam.simps[eqvt]
+
+lemma rfv_eqvt[eqvt]:
+  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
+apply(induct t)
+apply(simp_all)
+apply(simp add: permute_set_eq atom_eqvt)
+apply(simp add: union_eqvt)
+apply(simp add: Diff_eqvt)
+apply(simp add: permute_set_eq atom_eqvt)
+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. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s)
+       \<Longrightarrow> rLam a t \<approx> rLam b s"
+
+lemma a3_inverse:
+  assumes "rLam a t \<approx> rLam b s"
+  shows "\<exists>pi. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s)"
+using assms
+apply(erule_tac alpha.cases)
+apply(auto)
+done
+
+text {* should be automatic with new version of eqvt-machinery *}
+lemma alpha_eqvt:
+  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 ?p1="- pi" in permute_eq_iff[THEN iffD1])
+apply(simp only: Diff_eqvt rfv_eqvt insert_eqvt atom_eqvt empty_eqvt)
+apply(simp)
+apply(rule conjI)
+apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
+apply(simp add: Diff_eqvt rfv_eqvt atom_eqvt insert_eqvt empty_eqvt)
+apply(subst permute_eqvt[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="0" in exI)
+apply(simp_all add: fresh_star_def fresh_zero_perm)
+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="- pi" in exI)
+apply(simp)
+apply(simp add: fresh_star_def fresh_minus_perm)
+apply(erule conjE)+
+apply(rotate_tac 3)
+apply(drule_tac pi="- pi" in alpha_eqvt)
+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_plus)
+apply(drule_tac x="- pia \<bullet> sa" in spec)
+apply(drule mp)
+apply(rotate_tac 7)
+apply(drule_tac pi="- pia" in alpha_eqvt)
+apply(simp)
+apply(rotate_tac 9)
+apply(drule_tac pi="pia" in alpha_eqvt)
+apply(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_all)
+  done
+
+inductive
+    alpha2 :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx>2 _" [100, 100] 100)
+where
+  a21: "a = b \<Longrightarrow> (rVar a) \<approx>2 (rVar b)"
+| a22: "\<lbrakk>t1 \<approx>2 t2; s1 \<approx>2 s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx>2 rApp t2 s2"
+| a23: "(a = b \<and> t \<approx>2 s) \<or> (a \<noteq> b \<and> ((a \<leftrightarrow> b) \<bullet> t) \<approx>2 s \<and> atom b \<notin> rfv t)\<Longrightarrow> rLam a t \<approx>2 rLam b s"
+
+lemma fv_vars:
+  fixes a::name
+  assumes a1: "\<forall>x \<in> rfv t - {atom a}. pi \<bullet> x = x"
+  shows "(pi \<bullet> t) \<approx>2 ((a \<leftrightarrow> (pi \<bullet> a)) \<bullet> t)"
+using a1
+apply(induct t)
+apply(auto)
+apply(rule a21)
+apply(case_tac "name = a")
+apply(simp)
+apply(simp)
+defer
+apply(rule a22)
+apply(simp)
+apply(simp)
+apply(rule a23)
+apply(case_tac "a = name")
+apply(simp)
+oops
+
+
+lemma 
+  assumes a1: "t \<approx>2 s"
+  shows "t \<approx> s"
+using a1
+apply(induct)
+apply(rule alpha.intros)
+apply(simp)
+apply(rule alpha.intros)
+apply(simp)
+apply(simp)
+apply(rule alpha.intros)
+apply(erule disjE)
+apply(rule_tac x="0" in exI)
+apply(simp add: fresh_star_def fresh_zero_perm)
+apply(erule conjE)+
+apply(drule alpha_rfv)
+apply(simp)
+apply(rule_tac x="(a \<leftrightarrow> b)" in exI)
+apply(simp)
+apply(erule conjE)+
+apply(rule conjI)
+apply(drule alpha_rfv)
+apply(drule sym)
+apply(simp)
+apply(simp add: rfv_eqvt[symmetric])
+defer
+apply(subgoal_tac "atom a \<sharp> (rfv t - {atom a})")
+apply(subgoal_tac "atom b \<sharp> (rfv t - {atom a})")
+
+defer
+sorry
+
+lemma 
+  assumes a1: "t \<approx> s"
+  shows "t \<approx>2 s"
+using a1
+apply(induct)
+apply(rule alpha2.intros)
+apply(simp)
+apply(rule alpha2.intros)
+apply(simp)
+apply(simp)
+apply(clarify)
+apply(rule alpha2.intros)
+apply(frule alpha_rfv)
+apply(rotate_tac 4)
+apply(drule sym)
+apply(simp)
+apply(drule sym)
+apply(simp)
+oops
+
+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> atom set"
+is
+  "rfv"
+
+lemma perm_rsp[quot_respect]:
+  "(op = ===> alpha ===> alpha) permute permute"
+  apply(auto)
+  apply(rule alpha_eqvt)
+  apply(simp)
+  done
+
+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="0" in exI)
+  unfolding fresh_star_def 
+  apply(simp add: fresh_star_def fresh_zero_perm)
+  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"
+  apply (lifting rlam.induct)
+  done
+
+instantiation lam :: pt
+begin
+
+quotient_definition
+  "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam"
+is
+  "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam"
+
+lemma permute_lam [simp]:
+  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 permute_rlam.simps)
+done
+
+instance
+apply default
+apply(induct_tac [!] x rule: lam_induct)
+apply(simp_all)
+done
+
+end
+
+lemma fv_lam [simp]: 
+  shows "fv (Var a) = {atom a}"
+  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
+  and   "fv (Lam a t) = fv t - {atom a}"
+apply(lifting rfv_var rfv_app rfv_lam)
+done
+
+lemma fv_eqvt:
+  shows "(p \<bullet> fv t) = fv (p \<bullet> t)"
+apply(lifting rfv_eqvt)
+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. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s)\<rbrakk> 
+   \<Longrightarrow> Lam a t = Lam b s"
+  apply  (lifting a3)
+  done
+
+lemma a3_inv:
+  assumes "Lam a t = Lam b s"
+  shows "\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s)"
+using assms
+apply(lifting a3_inverse)
+done
+
+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. fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a}) \<sharp>* pi \<and> (pi \<bullet> t) = s\<rbrakk> 
+   \<Longrightarrow> P\<rbrakk>
+    \<Longrightarrow> P"
+  by (lifting alpha.cases)
+
+(* not sure whether needed *)
+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. fv t - {atom a} = fv s - {atom b} \<and>
+         (fv t - {atom a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s)\<rbrakk> 
+     \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
+    \<Longrightarrow> qxb qx qxa"
+  by (lifting alpha.induct)
+
+(* should they lift automatically *)
+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(regularize)
+prefer 2
+apply(regularize)
+prefer 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 Lam_pseudo_inject:
+  shows "(Lam a t = Lam b s) = 
+      (\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s))"
+apply(rule iffI)
+apply(rule a3_inv)
+apply(assumption)
+apply(rule a3)
+apply(assumption)
+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 only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+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)"
+  apply (simp add: supp_def)
+  done
+
+lemma var_supp:
+  shows "(supp (Var a)) = {a:::name}"
+  using var_supp1 by (simp add: supp_at_base)
+
+lemma app_supp:
+  shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)"
+apply(simp only: supp_def lam_inject)
+apply(simp add: Collect_imp_eq Collect_neg_eq)
+done
+
+(* supp for lam *)
+lemma lam_supp1:
+  shows "(supp (atom x, t)) supports (Lam x t) "
+apply(simp add: supports_def)
+apply(fold fresh_def)
+apply(simp add: fresh_Pair swap_fresh_fresh)
+apply(clarify)
+apply(subst swap_at_base_simps(3))
+apply(simp_all add: fresh_atom)
+done
+
+lemma lam_fsupp1:
+  assumes a: "finite (supp t)"
+  shows "finite (supp (Lam x t))"
+apply(rule supports_finite)
+apply(rule lam_supp1)
+apply(simp add: a supp_Pair supp_atom)
+done
+
+instance lam :: fs
+apply(default)
+apply(induct_tac x rule: lam_induct)
+apply(simp add: var_supp)
+apply(simp add: app_supp)
+apply(simp add: lam_fsupp1)
+done
+
+lemma supp_fv:
+  shows "supp t = fv t"
+apply(induct t rule: lam_induct)
+apply(simp add: var_supp)
+apply(simp add: app_supp)
+apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)")
+apply(simp add: supp_Abs)
+apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
+apply(simp add: Lam_pseudo_inject)
+apply(simp add: Abs_eq_iff alpha_gen)
+apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric])
+done
+
+lemma lam_supp2:
+  shows "supp (Lam x t) = supp (Abs {atom x} t)"
+apply(simp add: supp_def permute_set_eq atom_eqvt)
+apply(simp add: Lam_pseudo_inject)
+apply(simp add: Abs_eq_iff supp_fv alpha_gen)
+done
+
+lemma lam_supp:
+  shows "supp (Lam x t) = ((supp t) - {atom x})"
+apply(simp add: lam_supp2)
+apply(simp add: supp_Abs)
+done
+
+lemma fresh_lam:
+  "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)"
+apply(simp add: fresh_def)
+apply(simp add: lam_supp)
+apply(auto)
+done
+
+lemma lam_induct_strong:
+  fixes a::"'a::fs"
+  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; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
+  shows "P a lam"
+proof -
+  have "\<And>pi 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 a. P a (pi \<bullet> lam1)" by fact
+    have b2: "\<And>pi 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 a. P a (pi \<bullet> lam)" by fact
+    obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
+      apply(rule obtain_atom)
+      apply(auto)
+      sorry
+    from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" 
+      apply -
+      apply(rule a3)
+      apply(blast)
+      apply(simp add: fresh_Pair)
+      done
+    have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
+      apply(rule swap_fresh_fresh)
+      using fr
+      apply(simp add: fresh_lam fresh_Pair)
+      apply(simp add: fresh_lam fresh_Pair)
+      done
+    show "P a (pi \<bullet> Lam name lam)" 
+      apply (simp)
+      apply(subst eq[symmetric])
+      using p
+      apply(simp only: permute_lam)
+      apply(simp add: flip_def)
+      done
+  qed
+  then have "P a (0 \<bullet> lam)" by blast
+  then show "P a lam" by simp 
+qed
+
+
+lemma var_fresh:
+  fixes a::"name"
+  shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)"
+  apply(simp add: fresh_def)
+  apply(simp add: var_supp1)
+  done
+
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/LamEx2.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,567 @@
+theory LamEx
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
+begin
+
+atom_decl name
+
+datatype rlam =
+  rVar "name"
+| rApp "rlam" "rlam"
+| rLam "name" "rlam"
+
+fun
+  rfv :: "rlam \<Rightarrow> atom set"
+where
+  rfv_var: "rfv (rVar a) = {atom a}"
+| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
+| rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}"
+
+instantiation rlam :: pt
+begin
+
+primrec
+  permute_rlam
+where
+  "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)"
+| "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)"
+| "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)"
+
+instance
+apply default
+apply(induct_tac [!] x)
+apply(simp_all)
+done
+
+end
+
+instantiation rlam :: fs
+begin
+
+lemma neg_conj:
+  "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)"
+  by simp
+
+lemma infinite_Un:
+  "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
+  by simp
+
+instance
+apply default
+apply(induct_tac x)
+(* var case *)
+apply(simp add: supp_def)
+apply(fold supp_def)[1]
+apply(simp add: supp_at_base)
+(* app case *)
+apply(simp only: supp_def)
+apply(simp only: permute_rlam.simps) 
+apply(simp only: rlam.inject) 
+apply(simp only: neg_conj)
+apply(simp only: Collect_disj_eq)
+apply(simp only: infinite_Un)
+apply(simp only: Collect_disj_eq)
+apply(simp)
+(* lam case *)
+apply(simp only: supp_def)
+apply(simp only: permute_rlam.simps) 
+apply(simp only: rlam.inject) 
+apply(simp only: neg_conj)
+apply(simp only: Collect_disj_eq)
+apply(simp only: infinite_Un)
+apply(simp only: Collect_disj_eq)
+apply(simp)
+apply(fold supp_def)[1]
+apply(simp add: supp_at_base)
+done
+
+end
+
+
+(* for the eqvt proof of the alpha-equivalence *)
+declare permute_rlam.simps[eqvt]
+
+lemma rfv_eqvt[eqvt]:
+  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
+apply(induct t)
+apply(simp_all)
+apply(simp add: permute_set_eq atom_eqvt)
+apply(simp add: union_eqvt)
+apply(simp add: Diff_eqvt)
+apply(simp add: permute_set_eq atom_eqvt)
+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. (({atom a}, t) \<approx>gen alpha rfv pi ({atom b}, s)) \<Longrightarrow> rLam a t \<approx> rLam b s"
+print_theorems
+thm alpha.induct
+
+lemma a3_inverse:
+  assumes "rLam a t \<approx> rLam b s"
+  shows "\<exists>pi. (({atom a}, t) \<approx>gen alpha rfv pi ({atom b}, s))"
+using assms
+apply(erule_tac alpha.cases)
+apply(auto)
+done
+
+text {* should be automatic with new version of eqvt-machinery *}
+lemma alpha_eqvt:
+  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(rule alpha_gen_atom_eqvt)
+apply(rule rfv_eqvt)
+apply assumption
+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="0" in exI)
+apply(rule alpha_gen_refl)
+apply(assumption)
+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 alpha_gen_compose_sym)
+  apply(erule alpha_eqvt)
+  done
+
+lemma alpha_trans:
+  shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
+apply(induct arbitrary: t3 rule: alpha.induct)
+apply(simp add: a1)
+apply(rotate_tac 4)
+apply(erule alpha.cases)
+apply(simp_all add: a2)
+apply(erule alpha.cases)
+apply(simp_all)
+apply(rule a3)
+apply(erule alpha_gen_compose_trans)
+apply(assumption)
+apply(erule alpha_eqvt)
+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_all add: alpha_gen.simps)
+  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> atom set"
+is
+  "rfv"
+
+lemma perm_rsp[quot_respect]:
+  "(op = ===> alpha ===> alpha) permute permute"
+  apply(auto)
+  apply(rule alpha_eqvt)
+  apply(simp)
+  done
+
+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="0" in exI)
+  unfolding fresh_star_def 
+  apply(simp add: fresh_star_def fresh_zero_perm alpha_gen.simps)
+  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"
+  apply (lifting rlam.induct)
+  done
+
+instantiation lam :: pt
+begin
+
+quotient_definition
+  "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam"
+is
+  "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam"
+
+lemma permute_lam [simp]:
+  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 permute_rlam.simps)
+done
+
+instance
+apply default
+apply(induct_tac [!] x rule: lam_induct)
+apply(simp_all)
+done
+
+end
+
+lemma fv_lam [simp]: 
+  shows "fv (Var a) = {atom a}"
+  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
+  and   "fv (Lam a t) = fv t - {atom a}"
+apply(lifting rfv_var rfv_app rfv_lam)
+done
+
+lemma fv_eqvt:
+  shows "(p \<bullet> fv t) = fv (p \<bullet> t)"
+apply(lifting rfv_eqvt)
+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 alpha_gen_rsp_pre:
+  assumes a5: "\<And>t s. R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s)"
+  and     a1: "R s1 t1"
+  and     a2: "R s2 t2"
+  and     a3: "\<And>a b c d. R a b \<Longrightarrow> R c d \<Longrightarrow> R1 a c = R2 b d"
+  and     a4: "\<And>x y. R x y \<Longrightarrow> fv1 x = fv2 y"
+  shows   "(a, s1) \<approx>gen R1 fv1 pi (b, s2) = (a, t1) \<approx>gen R2 fv2 pi (b, t2)"
+apply (simp add: alpha_gen.simps)
+apply (simp only: a4[symmetric, OF a1] a4[symmetric, OF a2])
+apply auto
+apply (subst a3[symmetric])
+apply (rule a5)
+apply (rule a1)
+apply (rule a2)
+apply (assumption)
+apply (subst a3)
+apply (rule a5)
+apply (rule a1)
+apply (rule a2)
+apply (assumption)
+done
+
+lemma [quot_respect]: "(prod_rel op = alpha ===>
+           (alpha ===> alpha ===> op =) ===> (alpha ===> op =) ===> op = ===> prod_rel op = alpha ===> op =)
+           alpha_gen alpha_gen"
+apply simp
+apply clarify
+apply (rule alpha_gen_rsp_pre[of "alpha",OF alpha_eqvt])
+apply auto
+done
+
+(* pi_abs would be also sufficient to prove the next lemma *)
+lemma replam_eqvt: "pi \<bullet> (rep_lam x) = rep_lam (pi \<bullet> x)"
+apply (unfold rep_lam_def)
+sorry
+
+lemma [quot_preserve]: "(prod_fun id rep_lam --->
+           (abs_lam ---> abs_lam ---> id) ---> (abs_lam ---> id) ---> id ---> (prod_fun id rep_lam) ---> id)
+           alpha_gen = alpha_gen"
+apply (simp add: expand_fun_eq alpha_gen.simps Quotient_abs_rep[OF Quotient_lam])
+apply (simp add: replam_eqvt)
+apply (simp only: Quotient_abs_rep[OF Quotient_lam])
+apply auto
+done
+
+
+lemma alpha_prs [quot_preserve]: "(rep_lam ---> rep_lam ---> id) alpha = (op =)"
+apply (simp add: expand_fun_eq)
+apply (simp add: Quotient_rel_rep[OF Quotient_lam])
+done
+
+lemma a3:
+  "\<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s) \<Longrightarrow> Lam a t = Lam b s"
+  apply (unfold alpha_gen)
+  apply (lifting a3[unfolded alpha_gen])
+  done
+
+
+lemma a3_inv:
+  "Lam a t = Lam b s \<Longrightarrow> \<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s)"
+  apply (unfold alpha_gen)
+  apply (lifting a3_inverse[unfolded alpha_gen])
+  done
+
+lemma alpha_cases: 
+  "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
+    \<And>t1 t2 s1 s2. \<lbrakk>a1 = App t1 s1; a2 = App t2 s2; t1 = t2; s1 = s2\<rbrakk> \<Longrightarrow> P;
+    \<And>a t b s. \<lbrakk>a1 = Lam a t; a2 = Lam b s; \<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s)\<rbrakk>
+   \<Longrightarrow> P\<rbrakk>
+    \<Longrightarrow> P"
+unfolding alpha_gen
+apply (lifting alpha.cases[unfolded alpha_gen])
+done
+
+(* not sure whether needed *)
+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>a t b s. \<exists>pi. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. x1 = x2 \<and> qxb x1 x2) fv pi ({atom b}, s) \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
+    \<Longrightarrow> qxb qx qxa"
+unfolding alpha_gen by (lifting alpha.induct[unfolded alpha_gen])
+
+(* should they lift automatically *)
+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(regularize)
+prefer 2
+apply(regularize)
+prefer 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
+
+thm a3_inv
+lemma Lam_pseudo_inject:
+  shows "(Lam a t = Lam b s) = (\<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s))"
+apply(rule iffI)
+apply(rule a3_inv)
+apply(assumption)
+apply(rule a3)
+apply(assumption)
+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 only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+apply (erule alpha.cases)
+apply (simp_all only: rlam.distinct)
+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)"
+  apply (simp add: supp_def)
+  done
+
+lemma var_supp:
+  shows "(supp (Var a)) = {a:::name}"
+  using var_supp1 by (simp add: supp_at_base)
+
+lemma app_supp:
+  shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)"
+apply(simp only: supp_def lam_inject)
+apply(simp add: Collect_imp_eq Collect_neg_eq)
+done
+
+(* supp for lam *)
+lemma lam_supp1:
+  shows "(supp (atom x, t)) supports (Lam x t) "
+apply(simp add: supports_def)
+apply(fold fresh_def)
+apply(simp add: fresh_Pair swap_fresh_fresh)
+apply(clarify)
+apply(subst swap_at_base_simps(3))
+apply(simp_all add: fresh_atom)
+done
+
+lemma lam_fsupp1:
+  assumes a: "finite (supp t)"
+  shows "finite (supp (Lam x t))"
+apply(rule supports_finite)
+apply(rule lam_supp1)
+apply(simp add: a supp_Pair supp_atom)
+done
+
+instance lam :: fs
+apply(default)
+apply(induct_tac x rule: lam_induct)
+apply(simp add: var_supp)
+apply(simp add: app_supp)
+apply(simp add: lam_fsupp1)
+done
+
+lemma supp_fv:
+  shows "supp t = fv t"
+apply(induct t rule: lam_induct)
+apply(simp add: var_supp)
+apply(simp add: app_supp)
+apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)")
+apply(simp add: supp_Abs)
+apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
+apply(simp add: Lam_pseudo_inject)
+apply(simp add: Abs_eq_iff)
+apply(simp add: alpha_gen.simps)
+apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric])
+done
+
+lemma lam_supp2:
+  shows "supp (Lam x t) = supp (Abs {atom x} t)"
+apply(simp add: supp_def permute_set_eq atom_eqvt)
+apply(simp add: Lam_pseudo_inject)
+apply(simp add: Abs_eq_iff)
+apply(simp add: alpha_gen  supp_fv)
+done
+
+lemma lam_supp:
+  shows "supp (Lam x t) = ((supp t) - {atom x})"
+apply(simp add: lam_supp2)
+apply(simp add: supp_Abs)
+done
+
+lemma fresh_lam:
+  "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)"
+apply(simp add: fresh_def)
+apply(simp add: lam_supp)
+apply(auto)
+done
+
+lemma lam_induct_strong:
+  fixes a::"'a::fs"
+  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; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
+  shows "P a lam"
+proof -
+  have "\<And>pi 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 a. P a (pi \<bullet> lam1)" by fact
+    have b2: "\<And>pi 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 a. P a (pi \<bullet> lam)" by fact
+    obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
+      apply(rule obtain_atom)
+      apply(auto)
+      sorry
+    from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" 
+      apply -
+      apply(rule a3)
+      apply(blast)
+      apply(simp add: fresh_Pair)
+      done
+    have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
+      apply(rule swap_fresh_fresh)
+      using fr
+      apply(simp add: fresh_lam fresh_Pair)
+      apply(simp add: fresh_lam fresh_Pair)
+      done
+    show "P a (pi \<bullet> Lam name lam)" 
+      apply (simp)
+      apply(subst eq[symmetric])
+      using p
+      apply(simp only: permute_lam)
+      apply(simp add: flip_def)
+      done
+  qed
+  then have "P a (0 \<bullet> lam)" by blast
+  then show "P a lam" by simp 
+qed
+
+
+lemma var_fresh:
+  fixes a::"name"
+  shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)"
+  apply(simp add: fresh_def)
+  apply(simp add: var_supp1)
+  done
+
+
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Nominal2_Atoms.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,251 @@
+(*  Title:      Nominal2_Atoms
+    Authors:    Brian Huffman, Christian Urban
+
+    Definitions for concrete atom types. 
+*)
+theory Nominal2_Atoms
+imports Nominal2_Base
+uses ("nominal_atoms.ML")
+begin
+
+section {* Concrete atom types *}
+
+text {*
+  Class @{text at_base} allows types containing multiple sorts of atoms.
+  Class @{text at} only allows types with a single sort.
+*}
+
+class at_base = pt +
+  fixes atom :: "'a \<Rightarrow> atom"
+  assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b"
+  assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)"
+
+class at = at_base +
+  assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)"
+
+instance at < at_base ..
+
+lemma supp_at_base: 
+  fixes a::"'a::at_base"
+  shows "supp a = {atom a}"
+  by (simp add: supp_atom [symmetric] supp_def atom_eqvt)
+
+lemma fresh_at_base: 
+  shows "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b"
+  unfolding fresh_def by (simp add: supp_at_base)
+
+instance at_base < fs
+proof qed (simp add: supp_at_base)
+
+lemma at_base_infinite [simp]:
+  shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U")
+proof
+  obtain a :: 'a where "True" by auto
+  assume "finite ?U"
+  hence "finite (atom ` ?U)"
+    by (rule finite_imageI)
+  then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)"
+    by (rule obtain_atom)
+  from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)"
+    unfolding atom_eqvt [symmetric]
+    by (simp add: swap_atom)
+  hence "b \<in> atom ` ?U" by simp
+  with b(1) show "False" by simp
+qed
+
+lemma swap_at_base_simps [simp]:
+  fixes x y::"'a::at_base"
+  shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y"
+  and   "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x"
+  and   "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x"
+  unfolding atom_eq_iff [symmetric]
+  unfolding atom_eqvt [symmetric]
+  by simp_all
+
+lemma obtain_at_base:
+  assumes X: "finite X"
+  obtains a::"'a::at_base" where "atom a \<notin> X"
+proof -
+  have "inj (atom :: 'a \<Rightarrow> atom)"
+    by (simp add: inj_on_def)
+  with X have "finite (atom -` X :: 'a set)"
+    by (rule finite_vimageI)
+  with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)"
+    by auto
+  then obtain a :: 'a where "atom a \<notin> X"
+    by auto
+  thus ?thesis ..
+qed
+
+
+section {* A swapping operation for concrete atoms *}
+  
+definition
+  flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
+where
+  "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
+
+lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
+  unfolding flip_def by (rule swap_self)
+
+lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
+  unfolding flip_def by (rule swap_commute)
+
+lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
+  unfolding flip_def by (rule minus_swap)
+
+lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
+  unfolding flip_def by (rule swap_cancel)
+
+lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
+  unfolding permute_plus [symmetric] add_flip_cancel by simp
+
+lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
+  by (simp add: flip_commute)
+
+lemma flip_eqvt: 
+  fixes a b c::"'a::at_base"
+  shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
+  unfolding flip_def
+  by (simp add: swap_eqvt atom_eqvt)
+
+lemma flip_at_base_simps [simp]:
+  shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
+  and   "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
+  and   "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
+  and   "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
+  unfolding flip_def
+  unfolding atom_eq_iff [symmetric]
+  unfolding atom_eqvt [symmetric]
+  by simp_all
+
+text {* the following two lemmas do not hold for at_base, 
+  only for single sort atoms from at *}
+
+lemma permute_flip_at:
+  fixes a b c::"'a::at"
+  shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
+  unfolding flip_def
+  apply (rule atom_eq_iff [THEN iffD1])
+  apply (subst atom_eqvt [symmetric])
+  apply (simp add: swap_atom)
+  done
+
+lemma flip_at_simps [simp]:
+  fixes a b::"'a::at"
+  shows "(a \<leftrightarrow> b) \<bullet> a = b" 
+  and   "(a \<leftrightarrow> b) \<bullet> b = a"
+  unfolding permute_flip_at by simp_all
+
+
+subsection {* Syntax for coercing at-elements to the atom-type *}
+
+syntax
+  "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
+
+translations
+  "_atom_constrain a t" => "atom (_constrain a t)"
+
+
+subsection {* A lemma for proving instances of class @{text at}. *}
+
+setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
+setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
+
+text {*
+  New atom types are defined as subtypes of @{typ atom}.
+*}
+
+lemma exists_eq_simple_sort: 
+  shows "\<exists>a. a \<in> {a. sort_of a = s}"
+  by (rule_tac x="Atom s 0" in exI, simp)
+
+lemma exists_eq_sort: 
+  shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
+  by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
+
+lemma at_base_class:
+  fixes sort_fun :: "'b \<Rightarrow>atom_sort"
+  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
+  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
+  assumes atom_def: "\<And>a. atom a = Rep a"
+  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
+  shows "OFCLASS('a, at_base_class)"
+proof
+  interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
+  have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
+  fix a b :: 'a and p p1 p2 :: perm
+  show "0 \<bullet> a = a"
+    unfolding permute_def by (simp add: Rep_inverse)
+  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
+    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
+  show "atom a = atom b \<longleftrightarrow> a = b"
+    unfolding atom_def by (simp add: Rep_inject)
+  show "p \<bullet> atom a = atom (p \<bullet> a)"
+    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
+qed
+
+(*
+lemma at_class:
+  fixes s :: atom_sort
+  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
+  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
+  assumes atom_def: "\<And>a. atom a = Rep a"
+  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
+  shows "OFCLASS('a, at_class)"
+proof
+  interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
+  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
+  fix a b :: 'a and p p1 p2 :: perm
+  show "0 \<bullet> a = a"
+    unfolding permute_def by (simp add: Rep_inverse)
+  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
+    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
+  show "sort_of (atom a) = sort_of (atom b)"
+    unfolding atom_def by (simp add: sort_of_Rep)
+  show "atom a = atom b \<longleftrightarrow> a = b"
+    unfolding atom_def by (simp add: Rep_inject)
+  show "p \<bullet> atom a = atom (p \<bullet> a)"
+    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
+qed
+*)
+
+lemma at_class:
+  fixes s :: atom_sort
+  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
+  assumes type: "type_definition Rep Abs {a. sort_of a = s}"
+  assumes atom_def: "\<And>a. atom a = Rep a"
+  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
+  shows "OFCLASS('a, at_class)"
+proof
+  interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
+  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
+  fix a b :: 'a and p p1 p2 :: perm
+  show "0 \<bullet> a = a"
+    unfolding permute_def by (simp add: Rep_inverse)
+  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
+    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
+  show "sort_of (atom a) = sort_of (atom b)"
+    unfolding atom_def by (simp add: sort_of_Rep)
+  show "atom a = atom b \<longleftrightarrow> a = b"
+    unfolding atom_def by (simp add: Rep_inject)
+  show "p \<bullet> atom a = atom (p \<bullet> a)"
+    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
+qed
+
+setup {* Sign.add_const_constraint
+  (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
+setup {* Sign.add_const_constraint
+  (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
+
+
+section {* Automation for creating concrete atom types *}
+
+text {* at the moment only single-sort concrete atoms are supported *}
+
+use "nominal_atoms.ML"
+
+
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Nominal2_Base.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,1022 @@
+(*  Title:      Nominal2_Base
+    Authors:    Brian Huffman, Christian Urban
+
+    Basic definitions and lemma infrastructure for 
+    Nominal Isabelle. 
+*)
+theory Nominal2_Base
+imports Main Infinite_Set
+begin
+
+section {* Atoms and Sorts *}
+
+text {* A simple implementation for atom_sorts is strings. *}
+(* types atom_sort = string *)
+
+text {* To deal with Church-like binding we use trees of  
+  strings as sorts. *}
+
+datatype atom_sort = Sort "string" "atom_sort list"
+
+datatype atom = Atom atom_sort nat
+
+
+text {* Basic projection function. *}
+
+primrec
+  sort_of :: "atom \<Rightarrow> atom_sort"
+where
+  "sort_of (Atom s i) = s"
+
+
+text {* There are infinitely many atoms of each sort. *}
+lemma INFM_sort_of_eq: 
+  shows "INFM a. sort_of a = s"
+proof -
+  have "INFM i. sort_of (Atom s i) = s" by simp
+  moreover have "inj (Atom s)" by (simp add: inj_on_def)
+  ultimately show "INFM a. sort_of a = s" by (rule INFM_inj)
+qed
+
+lemma infinite_sort_of_eq:
+  shows "infinite {a. sort_of a = s}"
+  using INFM_sort_of_eq unfolding INFM_iff_infinite .
+
+lemma atom_infinite [simp]: 
+  shows "infinite (UNIV :: atom set)"
+  using subset_UNIV infinite_sort_of_eq
+  by (rule infinite_super)
+
+lemma obtain_atom:
+  fixes X :: "atom set"
+  assumes X: "finite X"
+  obtains a where "a \<notin> X" "sort_of a = s"
+proof -
+  from X have "MOST a. a \<notin> X"
+    unfolding MOST_iff_cofinite by simp
+  with INFM_sort_of_eq
+  have "INFM a. sort_of a = s \<and> a \<notin> X"
+    by (rule INFM_conjI)
+  then obtain a where "a \<notin> X" "sort_of a = s"
+    by (auto elim: INFM_E)
+  then show ?thesis ..
+qed
+
+section {* Sort-Respecting Permutations *}
+
+typedef perm =
+  "{f. bij f \<and> finite {a. f a \<noteq> a} \<and> (\<forall>a. sort_of (f a) = sort_of a)}"
+proof
+  show "id \<in> ?perm" by simp
+qed
+
+lemma permI:
+  assumes "bij f" and "MOST x. f x = x" and "\<And>a. sort_of (f a) = sort_of a"
+  shows "f \<in> perm"
+  using assms unfolding perm_def MOST_iff_cofinite by simp
+
+lemma perm_is_bij: "f \<in> perm \<Longrightarrow> bij f"
+  unfolding perm_def by simp
+
+lemma perm_is_finite: "f \<in> perm \<Longrightarrow> finite {a. f a \<noteq> a}"
+  unfolding perm_def by simp
+
+lemma perm_is_sort_respecting: "f \<in> perm \<Longrightarrow> sort_of (f a) = sort_of a"
+  unfolding perm_def by simp
+
+lemma perm_MOST: "f \<in> perm \<Longrightarrow> MOST x. f x = x"
+  unfolding perm_def MOST_iff_cofinite by simp
+
+lemma perm_id: "id \<in> perm"
+  unfolding perm_def by simp
+
+lemma perm_comp:
+  assumes f: "f \<in> perm" and g: "g \<in> perm"
+  shows "(f \<circ> g) \<in> perm"
+apply (rule permI)
+apply (rule bij_comp)
+apply (rule perm_is_bij [OF g])
+apply (rule perm_is_bij [OF f])
+apply (rule MOST_rev_mp [OF perm_MOST [OF g]])
+apply (rule MOST_rev_mp [OF perm_MOST [OF f]])
+apply (simp)
+apply (simp add: perm_is_sort_respecting [OF f])
+apply (simp add: perm_is_sort_respecting [OF g])
+done
+
+lemma perm_inv:
+  assumes f: "f \<in> perm"
+  shows "(inv f) \<in> perm"
+apply (rule permI)
+apply (rule bij_imp_bij_inv)
+apply (rule perm_is_bij [OF f])
+apply (rule MOST_mono [OF perm_MOST [OF f]])
+apply (erule subst, rule inv_f_f)
+apply (rule bij_is_inj [OF perm_is_bij [OF f]])
+apply (rule perm_is_sort_respecting [OF f, THEN sym, THEN trans])
+apply (simp add: surj_f_inv_f [OF bij_is_surj [OF perm_is_bij [OF f]]])
+done
+
+lemma bij_Rep_perm: "bij (Rep_perm p)"
+  using Rep_perm [of p] unfolding perm_def by simp
+
+lemma finite_Rep_perm: "finite {a. Rep_perm p a \<noteq> a}"
+  using Rep_perm [of p] unfolding perm_def by simp
+
+lemma sort_of_Rep_perm: "sort_of (Rep_perm p a) = sort_of a"
+  using Rep_perm [of p] unfolding perm_def by simp
+
+lemma Rep_perm_ext:
+  "Rep_perm p1 = Rep_perm p2 \<Longrightarrow> p1 = p2"
+  by (simp add: expand_fun_eq Rep_perm_inject [symmetric])
+
+
+subsection {* Permutations form a group *}
+
+instantiation perm :: group_add
+begin
+
+definition
+  "0 = Abs_perm id"
+
+definition
+  "- p = Abs_perm (inv (Rep_perm p))"
+
+definition
+  "p + q = Abs_perm (Rep_perm p \<circ> Rep_perm q)"
+
+definition
+  "(p1::perm) - p2 = p1 + - p2"
+
+lemma Rep_perm_0: "Rep_perm 0 = id"
+  unfolding zero_perm_def
+  by (simp add: Abs_perm_inverse perm_id)
+
+lemma Rep_perm_add:
+  "Rep_perm (p1 + p2) = Rep_perm p1 \<circ> Rep_perm p2"
+  unfolding plus_perm_def
+  by (simp add: Abs_perm_inverse perm_comp Rep_perm)
+
+lemma Rep_perm_uminus:
+  "Rep_perm (- p) = inv (Rep_perm p)"
+  unfolding uminus_perm_def
+  by (simp add: Abs_perm_inverse perm_inv Rep_perm)
+
+instance
+apply default
+unfolding Rep_perm_inject [symmetric]
+unfolding minus_perm_def
+unfolding Rep_perm_add
+unfolding Rep_perm_uminus
+unfolding Rep_perm_0
+by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]])
+
+end
+
+
+section {* Implementation of swappings *}
+
+definition
+  swap :: "atom \<Rightarrow> atom \<Rightarrow> perm" ("'(_ \<rightleftharpoons> _')")
+where
+  "(a \<rightleftharpoons> b) =
+    Abs_perm (if sort_of a = sort_of b 
+              then (\<lambda>c. if a = c then b else if b = c then a else c) 
+              else id)"
+
+lemma Rep_perm_swap:
+  "Rep_perm (a \<rightleftharpoons> b) =
+    (if sort_of a = sort_of b 
+     then (\<lambda>c. if a = c then b else if b = c then a else c)
+     else id)"
+unfolding swap_def
+apply (rule Abs_perm_inverse)
+apply (rule permI)
+apply (auto simp add: bij_def inj_on_def surj_def)[1]
+apply (rule MOST_rev_mp [OF MOST_neq(1) [of a]])
+apply (rule MOST_rev_mp [OF MOST_neq(1) [of b]])
+apply (simp)
+apply (simp)
+done
+
+lemmas Rep_perm_simps =
+  Rep_perm_0
+  Rep_perm_add
+  Rep_perm_uminus
+  Rep_perm_swap
+
+lemma swap_different_sorts [simp]:
+  "sort_of a \<noteq> sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) = 0"
+  by (rule Rep_perm_ext) (simp add: Rep_perm_simps)
+
+lemma swap_cancel:
+  "(a \<rightleftharpoons> b) + (a \<rightleftharpoons> b) = 0"
+by (rule Rep_perm_ext) 
+   (simp add: Rep_perm_simps expand_fun_eq)
+
+lemma swap_self [simp]:
+  "(a \<rightleftharpoons> a) = 0"
+  by (rule Rep_perm_ext, simp add: Rep_perm_simps expand_fun_eq)
+
+lemma minus_swap [simp]:
+  "- (a \<rightleftharpoons> b) = (a \<rightleftharpoons> b)"
+  by (rule minus_unique [OF swap_cancel])
+
+lemma swap_commute:
+  "(a \<rightleftharpoons> b) = (b \<rightleftharpoons> a)"
+  by (rule Rep_perm_ext)
+     (simp add: Rep_perm_swap expand_fun_eq)
+
+lemma swap_triple:
+  assumes "a \<noteq> b" and "c \<noteq> b"
+  assumes "sort_of a = sort_of b" "sort_of b = sort_of c"
+  shows "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
+  using assms
+  by (rule_tac Rep_perm_ext)
+     (auto simp add: Rep_perm_simps expand_fun_eq)
+
+
+section {* Permutation Types *}
+
+text {*
+  Infix syntax for @{text permute} has higher precedence than
+  addition, but lower than unary minus.
+*}
+
+class pt =
+  fixes permute :: "perm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [76, 75] 75)
+  assumes permute_zero [simp]: "0 \<bullet> x = x"
+  assumes permute_plus [simp]: "(p + q) \<bullet> x = p \<bullet> (q \<bullet> x)"
+begin
+
+lemma permute_diff [simp]:
+  shows "(p - q) \<bullet> x = p \<bullet> - q \<bullet> x"
+  unfolding diff_minus by simp
+
+lemma permute_minus_cancel [simp]:
+  shows "p \<bullet> - p \<bullet> x = x"
+  and   "- p \<bullet> p \<bullet> x = x"
+  unfolding permute_plus [symmetric] by simp_all
+
+lemma permute_swap_cancel [simp]:
+  shows "(a \<rightleftharpoons> b) \<bullet> (a \<rightleftharpoons> b) \<bullet> x = x"
+  unfolding permute_plus [symmetric]
+  by (simp add: swap_cancel)
+
+lemma permute_swap_cancel2 [simp]:
+  shows "(a \<rightleftharpoons> b) \<bullet> (b \<rightleftharpoons> a) \<bullet> x = x"
+  unfolding permute_plus [symmetric]
+  by (simp add: swap_commute)
+
+lemma inj_permute [simp]: 
+  shows "inj (permute p)"
+  by (rule inj_on_inverseI)
+     (rule permute_minus_cancel)
+
+lemma surj_permute [simp]: 
+  shows "surj (permute p)"
+  by (rule surjI, rule permute_minus_cancel)
+
+lemma bij_permute [simp]: 
+  shows "bij (permute p)"
+  by (rule bijI [OF inj_permute surj_permute])
+
+lemma inv_permute: 
+  shows "inv (permute p) = permute (- p)"
+  by (rule inv_equality) (simp_all)
+
+lemma permute_minus: 
+  shows "permute (- p) = inv (permute p)"
+  by (simp add: inv_permute)
+
+lemma permute_eq_iff [simp]: 
+  shows "p \<bullet> x = p \<bullet> y \<longleftrightarrow> x = y"
+  by (rule inj_permute [THEN inj_eq])
+
+end
+
+subsection {* Permutations for atoms *}
+
+instantiation atom :: pt
+begin
+
+definition
+  "p \<bullet> a = Rep_perm p a"
+
+instance 
+apply(default)
+apply(simp_all add: permute_atom_def Rep_perm_simps)
+done
+
+end
+
+lemma sort_of_permute [simp]:
+  shows "sort_of (p \<bullet> a) = sort_of a"
+  unfolding permute_atom_def by (rule sort_of_Rep_perm)
+
+lemma swap_atom:
+  shows "(a \<rightleftharpoons> b) \<bullet> c =
+           (if sort_of a = sort_of b
+            then (if c = a then b else if c = b then a else c) else c)"
+  unfolding permute_atom_def
+  by (simp add: Rep_perm_swap)
+
+lemma swap_atom_simps [simp]:
+  "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> a = b"
+  "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> b = a"
+  "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> c = c"
+  unfolding swap_atom by simp_all
+
+lemma expand_perm_eq:
+  fixes p q :: "perm"
+  shows "p = q \<longleftrightarrow> (\<forall>a::atom. p \<bullet> a = q \<bullet> a)"
+  unfolding permute_atom_def
+  by (metis Rep_perm_ext ext)
+
+
+subsection {* Permutations for permutations *}
+
+instantiation perm :: pt
+begin
+
+definition
+  "p \<bullet> q = p + q - p"
+
+instance
+apply default
+apply (simp add: permute_perm_def)
+apply (simp add: permute_perm_def diff_minus minus_add add_assoc)
+done
+
+end
+
+lemma permute_self: "p \<bullet> p = p"
+unfolding permute_perm_def by (simp add: diff_minus add_assoc)
+
+lemma permute_eqvt:
+  shows "p \<bullet> (q \<bullet> x) = (p \<bullet> q) \<bullet> (p \<bullet> x)"
+  unfolding permute_perm_def by simp
+
+lemma zero_perm_eqvt:
+  shows "p \<bullet> (0::perm) = 0"
+  unfolding permute_perm_def by simp
+
+lemma add_perm_eqvt:
+  fixes p p1 p2 :: perm
+  shows "p \<bullet> (p1 + p2) = p \<bullet> p1 + p \<bullet> p2"
+  unfolding permute_perm_def
+  by (simp add: expand_perm_eq)
+
+lemma swap_eqvt:
+  shows "p \<bullet> (a \<rightleftharpoons> b) = (p \<bullet> a \<rightleftharpoons> p \<bullet> b)"
+  unfolding permute_perm_def
+  by (auto simp add: swap_atom expand_perm_eq)
+
+
+subsection {* Permutations for functions *}
+
+instantiation "fun" :: (pt, pt) pt
+begin
+
+definition
+  "p \<bullet> f = (\<lambda>x. p \<bullet> (f (- p \<bullet> x)))"
+
+instance
+apply default
+apply (simp add: permute_fun_def)
+apply (simp add: permute_fun_def minus_add)
+done
+
+end
+
+lemma permute_fun_app_eq:
+  shows "p \<bullet> (f x) = (p \<bullet> f) (p \<bullet> x)"
+unfolding permute_fun_def by simp
+
+
+subsection {* Permutations for booleans *}
+
+instantiation bool :: pt
+begin
+
+definition "p \<bullet> (b::bool) = b"
+
+instance
+apply(default) 
+apply(simp_all add: permute_bool_def)
+done
+
+end
+
+lemma Not_eqvt:
+  shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))"
+by (simp add: permute_bool_def)
+
+
+subsection {* Permutations for sets *}
+
+lemma permute_set_eq:
+  fixes x::"'a::pt"
+  and   p::"perm"
+  shows "(p \<bullet> X) = {p \<bullet> x | x. x \<in> X}"
+  apply(auto simp add: permute_fun_def permute_bool_def mem_def)
+  apply(rule_tac x="- p \<bullet> x" in exI)
+  apply(simp)
+  done
+
+lemma permute_set_eq_image:
+  shows "p \<bullet> X = permute p ` X"
+unfolding permute_set_eq by auto
+
+lemma permute_set_eq_vimage:
+  shows "p \<bullet> X = permute (- p) -` X"
+unfolding permute_fun_def permute_bool_def
+unfolding vimage_def Collect_def mem_def ..
+
+lemma swap_set_not_in:
+  assumes a: "a \<notin> S" "b \<notin> S"
+  shows "(a \<rightleftharpoons> b) \<bullet> S = S"
+  using a by (auto simp add: permute_set_eq swap_atom)
+
+lemma swap_set_in:
+  assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b"
+  shows "(a \<rightleftharpoons> b) \<bullet> S \<noteq> S"
+  using a by (auto simp add: permute_set_eq swap_atom)
+
+
+subsection {* Permutations for units *}
+
+instantiation unit :: pt
+begin
+
+definition "p \<bullet> (u::unit) = u"
+
+instance proof
+qed (simp_all add: permute_unit_def)
+
+end
+
+
+subsection {* Permutations for products *}
+
+instantiation "*" :: (pt, pt) pt
+begin
+
+primrec 
+  permute_prod 
+where
+  Pair_eqvt: "p \<bullet> (x, y) = (p \<bullet> x, p \<bullet> y)"
+
+instance
+by default auto
+
+end
+
+subsection {* Permutations for sums *}
+
+instantiation "+" :: (pt, pt) pt
+begin
+
+primrec 
+  permute_sum 
+where
+  "p \<bullet> (Inl x) = Inl (p \<bullet> x)"
+| "p \<bullet> (Inr y) = Inr (p \<bullet> y)"
+
+instance proof
+qed (case_tac [!] x, simp_all)
+
+end
+
+subsection {* Permutations for lists *}
+
+instantiation list :: (pt) pt
+begin
+
+primrec 
+  permute_list 
+where
+  "p \<bullet> [] = []"
+| "p \<bullet> (x # xs) = p \<bullet> x # p \<bullet> xs"
+
+instance proof
+qed (induct_tac [!] x, simp_all)
+
+end
+
+subsection {* Permutations for options *}
+
+instantiation option :: (pt) pt
+begin
+
+primrec 
+  permute_option 
+where
+  "p \<bullet> None = None"
+| "p \<bullet> (Some x) = Some (p \<bullet> x)"
+
+instance proof
+qed (induct_tac [!] x, simp_all)
+
+end
+
+subsection {* Permutations for @{typ char}, @{typ nat}, and @{typ int} *}
+
+instantiation char :: pt
+begin
+
+definition "p \<bullet> (c::char) = c"
+
+instance proof
+qed (simp_all add: permute_char_def)
+
+end
+
+instantiation nat :: pt
+begin
+
+definition "p \<bullet> (n::nat) = n"
+
+instance proof
+qed (simp_all add: permute_nat_def)
+
+end
+
+instantiation int :: pt
+begin
+
+definition "p \<bullet> (i::int) = i"
+
+instance proof
+qed (simp_all add: permute_int_def)
+
+end
+
+
+section {* Pure types *}
+
+text {* Pure types will have always empty support. *}
+
+class pure = pt +
+  assumes permute_pure: "p \<bullet> x = x"
+
+text {* Types @{typ unit} and @{typ bool} are pure. *}
+
+instance unit :: pure
+proof qed (rule permute_unit_def)
+
+instance bool :: pure
+proof qed (rule permute_bool_def)
+
+text {* Other type constructors preserve purity. *}
+
+instance "fun" :: (pure, pure) pure
+by default (simp add: permute_fun_def permute_pure)
+
+instance "*" :: (pure, pure) pure
+by default (induct_tac x, simp add: permute_pure)
+
+instance "+" :: (pure, pure) pure
+by default (induct_tac x, simp_all add: permute_pure)
+
+instance list :: (pure) pure
+by default (induct_tac x, simp_all add: permute_pure)
+
+instance option :: (pure) pure
+by default (induct_tac x, simp_all add: permute_pure)
+
+
+subsection {* Types @{typ char}, @{typ nat}, and @{typ int} *}
+
+instance char :: pure
+proof qed (rule permute_char_def)
+
+instance nat :: pure
+proof qed (rule permute_nat_def)
+
+instance int :: pure
+proof qed (rule permute_int_def)
+
+
+subsection {* Supp, Freshness and Supports *}
+
+context pt
+begin
+
+definition
+  supp :: "'a \<Rightarrow> atom set"
+where
+  "supp x = {a. infinite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}}"
+
+end
+
+definition
+  fresh :: "atom \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp> _" [55, 55] 55)
+where   
+  "a \<sharp> x \<equiv> a \<notin> supp x"
+
+lemma supp_conv_fresh: 
+  shows "supp x = {a. \<not> a \<sharp> x}"
+  unfolding fresh_def by simp
+
+lemma swap_rel_trans:
+  assumes "sort_of a = sort_of b"
+  assumes "sort_of b = sort_of c"
+  assumes "(a \<rightleftharpoons> c) \<bullet> x = x"
+  assumes "(b \<rightleftharpoons> c) \<bullet> x = x"
+  shows "(a \<rightleftharpoons> b) \<bullet> x = x"
+proof (cases)
+  assume "a = b \<or> c = b"
+  with assms show "(a \<rightleftharpoons> b) \<bullet> x = x" by auto
+next
+  assume *: "\<not> (a = b \<or> c = b)"
+  have "((a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c)) \<bullet> x = x"
+    using assms by simp
+  also have "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
+    using assms * by (simp add: swap_triple)
+  finally show "(a \<rightleftharpoons> b) \<bullet> x = x" .
+qed
+
+lemma swap_fresh_fresh:
+  assumes a: "a \<sharp> x" 
+  and     b: "b \<sharp> x"
+  shows "(a \<rightleftharpoons> b) \<bullet> x = x"
+proof (cases)
+  assume asm: "sort_of a = sort_of b" 
+  have "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}" "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}" 
+    using a b unfolding fresh_def supp_def by simp_all
+  then have "finite ({c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x} \<union> {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x})" by simp
+  then obtain c 
+    where "(a \<rightleftharpoons> c) \<bullet> x = x" "(b \<rightleftharpoons> c) \<bullet> x = x" "sort_of c = sort_of b"
+    by (rule obtain_atom) (auto)
+  then show "(a \<rightleftharpoons> b) \<bullet> x = x" using asm by (rule_tac swap_rel_trans) (simp_all)
+next
+  assume "sort_of a \<noteq> sort_of b"
+  then show "(a \<rightleftharpoons> b) \<bullet> x = x" by simp
+qed
+
+
+subsection {* supp and fresh are equivariant *}
+
+lemma finite_Collect_bij:
+  assumes a: "bij f"
+  shows "finite {x. P (f x)} = finite {x. P x}"
+by (metis a finite_vimage_iff vimage_Collect_eq)
+
+lemma fresh_permute_iff:
+  shows "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x"
+proof -
+  have "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
+    unfolding fresh_def supp_def by simp
+  also have "\<dots> \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> p \<bullet> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
+    using bij_permute by (rule finite_Collect_bij [symmetric])
+  also have "\<dots> \<longleftrightarrow> finite {b. p \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> p \<bullet> x}"
+    by (simp only: permute_eqvt [of p] swap_eqvt)
+  also have "\<dots> \<longleftrightarrow> finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"
+    by (simp only: permute_eq_iff)
+  also have "\<dots> \<longleftrightarrow> a \<sharp> x"
+    unfolding fresh_def supp_def by simp
+  finally show ?thesis .
+qed
+
+lemma fresh_eqvt:
+  shows "p \<bullet> (a \<sharp> x) = (p \<bullet> a) \<sharp> (p \<bullet> x)"
+  by (simp add: permute_bool_def fresh_permute_iff)
+
+lemma supp_eqvt:
+  fixes  p  :: "perm"
+  and    x   :: "'a::pt"
+  shows "p \<bullet> (supp x) = supp (p \<bullet> x)"
+  unfolding supp_conv_fresh
+  unfolding permute_fun_def Collect_def
+  by (simp add: Not_eqvt fresh_eqvt)
+
+subsection {* supports *}
+
+definition
+  supports :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" (infixl "supports" 80)
+where  
+  "S supports x \<equiv> \<forall>a b. (a \<notin> S \<and> b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x)"
+
+lemma supp_is_subset:
+  fixes S :: "atom set"
+  and   x :: "'a::pt"
+  assumes a1: "S supports x"
+  and     a2: "finite S"
+  shows "(supp x) \<subseteq> S"
+proof (rule ccontr)
+  assume "\<not>(supp x \<subseteq> S)"
+  then obtain a where b1: "a \<in> supp x" and b2: "a \<notin> S" by auto
+  from a1 b2 have "\<forall>b. b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" by (unfold supports_def) (auto)
+  hence "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x} \<subseteq> S" by auto
+  with a2 have "finite {b. (a \<rightleftharpoons> b)\<bullet>x \<noteq> x}" by (simp add: finite_subset)
+  then have "a \<notin> (supp x)" unfolding supp_def by simp
+  with b1 show False by simp
+qed
+
+lemma supports_finite:
+  fixes S :: "atom set"
+  and   x :: "'a::pt"
+  assumes a1: "S supports x"
+  and     a2: "finite S"
+  shows "finite (supp x)"
+proof -
+  have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
+  then show "finite (supp x)" using a2 by (simp add: finite_subset)
+qed
+
+lemma supp_supports:
+  fixes x :: "'a::pt"
+  shows "(supp x) supports x"
+proof (unfold supports_def, intro strip)
+  fix a b
+  assume "a \<notin> (supp x) \<and> b \<notin> (supp x)"
+  then have "a \<sharp> x" and "b \<sharp> x" by (simp_all add: fresh_def)
+  then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (rule swap_fresh_fresh)
+qed
+
+lemma supp_is_least_supports:
+  fixes S :: "atom set"
+  and   x :: "'a::pt"
+  assumes  a1: "S supports x"
+  and      a2: "finite S"
+  and      a3: "\<And>S'. finite S' \<Longrightarrow> (S' supports x) \<Longrightarrow> S \<subseteq> S'"
+  shows "(supp x) = S"
+proof (rule equalityI)
+  show "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
+  with a2 have fin: "finite (supp x)" by (rule rev_finite_subset)
+  have "(supp x) supports x" by (rule supp_supports)
+  with fin a3 show "S \<subseteq> supp x" by blast
+qed
+
+lemma subsetCI: 
+  shows "(\<And>x. x \<in> A \<Longrightarrow> x \<notin> B \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> B"
+  by auto
+
+lemma finite_supp_unique:
+  assumes a1: "S supports x"
+  assumes a2: "finite S"
+  assumes a3: "\<And>a b. \<lbrakk>a \<in> S; b \<notin> S; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
+  shows "(supp x) = S"
+  using a1 a2
+proof (rule supp_is_least_supports)
+  fix S'
+  assume "finite S'" and "S' supports x"
+  show "S \<subseteq> S'"
+  proof (rule subsetCI)
+    fix a
+    assume "a \<in> S" and "a \<notin> S'"
+    have "finite (S \<union> S')"
+      using `finite S` `finite S'` by simp
+    then obtain b where "b \<notin> S \<union> S'" and "sort_of b = sort_of a"
+      by (rule obtain_atom)
+    then have "b \<notin> S" and "b \<notin> S'"  and "sort_of a = sort_of b"
+      by simp_all
+    then have "(a \<rightleftharpoons> b) \<bullet> x = x"
+      using `a \<notin> S'` `S' supports x` by (simp add: supports_def)
+    moreover have "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
+      using `a \<in> S` `b \<notin> S` `sort_of a = sort_of b`
+      by (rule a3)
+    ultimately show "False" by simp
+  qed
+qed
+
+section {* Finitely-supported types *}
+
+class fs = pt +
+  assumes finite_supp: "finite (supp x)"
+
+lemma pure_supp: 
+  shows "supp (x::'a::pure) = {}"
+  unfolding supp_def by (simp add: permute_pure)
+
+lemma pure_fresh:
+  fixes x::"'a::pure"
+  shows "a \<sharp> x"
+  unfolding fresh_def by (simp add: pure_supp)
+
+instance pure < fs
+by default (simp add: pure_supp)
+
+
+subsection  {* Type @{typ atom} is finitely-supported. *}
+
+lemma supp_atom:
+  shows "supp a = {a}"
+apply (rule finite_supp_unique)
+apply (clarsimp simp add: supports_def)
+apply simp
+apply simp
+done
+
+lemma fresh_atom: 
+  shows "a \<sharp> b \<longleftrightarrow> a \<noteq> b"
+  unfolding fresh_def supp_atom by simp
+
+instance atom :: fs
+by default (simp add: supp_atom)
+
+
+section {* Type @{typ perm} is finitely-supported. *}
+
+lemma perm_swap_eq:
+  shows "(a \<rightleftharpoons> b) \<bullet> p = p \<longleftrightarrow> (p \<bullet> (a \<rightleftharpoons> b)) = (a \<rightleftharpoons> b)"
+unfolding permute_perm_def
+by (metis add_diff_cancel minus_perm_def)
+
+lemma supports_perm: 
+  shows "{a. p \<bullet> a \<noteq> a} supports p"
+  unfolding supports_def
+  by (simp add: perm_swap_eq swap_eqvt)
+
+lemma finite_perm_lemma: 
+  shows "finite {a::atom. p \<bullet> a \<noteq> a}"
+  using finite_Rep_perm [of p]
+  unfolding permute_atom_def .
+
+lemma supp_perm:
+  shows "supp p = {a. p \<bullet> a \<noteq> a}"
+apply (rule finite_supp_unique)
+apply (rule supports_perm)
+apply (rule finite_perm_lemma)
+apply (simp add: perm_swap_eq swap_eqvt)
+apply (auto simp add: expand_perm_eq swap_atom)
+done
+
+lemma fresh_perm: 
+  shows "a \<sharp> p \<longleftrightarrow> p \<bullet> a = a"
+unfolding fresh_def by (simp add: supp_perm)
+
+lemma supp_swap:
+  shows "supp (a \<rightleftharpoons> b) = (if a = b \<or> sort_of a \<noteq> sort_of b then {} else {a, b})"
+  by (auto simp add: supp_perm swap_atom)
+
+lemma fresh_zero_perm: 
+  shows "a \<sharp> (0::perm)"
+  unfolding fresh_perm by simp
+
+lemma supp_zero_perm: 
+  shows "supp (0::perm) = {}"
+  unfolding supp_perm by simp
+
+lemma fresh_plus_perm:
+  fixes p q::perm
+  assumes "a \<sharp> p" "a \<sharp> q"
+  shows "a \<sharp> (p + q)"
+  using assms
+  unfolding fresh_def
+  by (auto simp add: supp_perm)
+
+lemma supp_plus_perm:
+  fixes p q::perm
+  shows "supp (p + q) \<subseteq> supp p \<union> supp q"
+  by (auto simp add: supp_perm)
+
+lemma fresh_minus_perm:
+  fixes p::perm
+  shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p"
+  unfolding fresh_def
+  apply(auto simp add: supp_perm)
+  apply(metis permute_minus_cancel)+
+  done
+
+lemma supp_minus_perm:
+  fixes p::perm
+  shows "supp (- p) = supp p"
+  unfolding supp_conv_fresh
+  by (simp add: fresh_minus_perm)
+
+instance perm :: fs
+by default (simp add: supp_perm finite_perm_lemma)
+
+
+section {* Finite Support instances for other types *}
+
+subsection {* Type @{typ "'a \<times> 'b"} is finitely-supported. *}
+
+lemma supp_Pair: 
+  shows "supp (x, y) = supp x \<union> supp y"
+  by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
+
+lemma fresh_Pair: 
+  shows "a \<sharp> (x, y) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> y"
+  by (simp add: fresh_def supp_Pair)
+
+instance "*" :: (fs, fs) fs
+apply default
+apply (induct_tac x)
+apply (simp add: supp_Pair finite_supp)
+done
+
+subsection {* Type @{typ "'a + 'b"} is finitely supported *}
+
+lemma supp_Inl: 
+  shows "supp (Inl x) = supp x"
+  by (simp add: supp_def)
+
+lemma supp_Inr: 
+  shows "supp (Inr x) = supp x"
+  by (simp add: supp_def)
+
+lemma fresh_Inl: 
+  shows "a \<sharp> Inl x \<longleftrightarrow> a \<sharp> x"
+  by (simp add: fresh_def supp_Inl)
+
+lemma fresh_Inr: 
+  shows "a \<sharp> Inr y \<longleftrightarrow> a \<sharp> y"
+  by (simp add: fresh_def supp_Inr)
+
+instance "+" :: (fs, fs) fs
+apply default
+apply (induct_tac x)
+apply (simp_all add: supp_Inl supp_Inr finite_supp)
+done
+
+subsection {* Type @{typ "'a option"} is finitely supported *}
+
+lemma supp_None: 
+  shows "supp None = {}"
+by (simp add: supp_def)
+
+lemma supp_Some: 
+  shows "supp (Some x) = supp x"
+  by (simp add: supp_def)
+
+lemma fresh_None: 
+  shows "a \<sharp> None"
+  by (simp add: fresh_def supp_None)
+
+lemma fresh_Some: 
+  shows "a \<sharp> Some x \<longleftrightarrow> a \<sharp> x"
+  by (simp add: fresh_def supp_Some)
+
+instance option :: (fs) fs
+apply default
+apply (induct_tac x)
+apply (simp_all add: supp_None supp_Some finite_supp)
+done
+
+subsubsection {* Type @{typ "'a list"} is finitely supported *}
+
+lemma supp_Nil: 
+  shows "supp [] = {}"
+  by (simp add: supp_def)
+
+lemma supp_Cons: 
+  shows "supp (x # xs) = supp x \<union> supp xs"
+by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
+
+lemma fresh_Nil: 
+  shows "a \<sharp> []"
+  by (simp add: fresh_def supp_Nil)
+
+lemma fresh_Cons: 
+  shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
+  by (simp add: fresh_def supp_Cons)
+
+instance list :: (fs) fs
+apply default
+apply (induct_tac x)
+apply (simp_all add: supp_Nil supp_Cons finite_supp)
+done
+
+section {* Support and freshness for applications *}
+
+lemma supp_fun_app:
+  shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)"
+proof (rule subsetCI)
+  fix a::"atom"
+  assume a: "a \<in> supp (f x)"
+  assume b: "a \<notin> supp f \<union> supp x"
+  then have "finite {b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f}" "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}" 
+    unfolding supp_def by auto
+  then have "finite ({b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f} \<union> {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x})" by simp
+  moreover 
+  have "{b. ((a \<rightleftharpoons> b) \<bullet> f) ((a \<rightleftharpoons> b) \<bullet> x) \<noteq> f x} \<subseteq> ({b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f} \<union> {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x})"
+    by auto
+  ultimately have "finite {b. ((a \<rightleftharpoons> b) \<bullet> f) ((a \<rightleftharpoons> b) \<bullet> x) \<noteq> f x}"
+    using finite_subset by auto
+  then have "a \<notin> supp (f x)" unfolding supp_def
+    by (simp add: permute_fun_app_eq)
+  with a show "False" by simp
+qed
+    
+lemma fresh_fun_app:
+  shows "a \<sharp> (f, x) \<Longrightarrow> a \<sharp> f x"
+unfolding fresh_def
+using supp_fun_app
+by (auto simp add: supp_Pair)
+
+lemma fresh_fun_eqvt_app:
+  assumes a: "\<forall>p. p \<bullet> f = f"
+  shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
+proof -
+  from a have b: "supp f = {}"
+    unfolding supp_def by simp
+  show "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
+    unfolding fresh_def
+    using supp_fun_app b
+    by auto
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Nominal2_Eqvt.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,304 @@
+(*  Title:      Nominal2_Eqvt
+    Authors:    Brian Huffman, Christian Urban
+
+    Equivariance, Supp and Fresh Lemmas for Operators. 
+    (Contains most, but not all such lemmas.)
+*)
+theory Nominal2_Eqvt
+imports Nominal2_Base
+uses ("nominal_thmdecls.ML")
+     ("nominal_permeq.ML")
+begin
+
+section {* Logical Operators *}
+
+
+lemma eq_eqvt:
+  shows "p \<bullet> (x = y) \<longleftrightarrow> (p \<bullet> x) = (p \<bullet> y)"
+  unfolding permute_eq_iff permute_bool_def ..
+
+lemma if_eqvt:
+  shows "p \<bullet> (if b then x else y) = (if p \<bullet> b then p \<bullet> x else p \<bullet> y)"
+  by (simp add: permute_fun_def permute_bool_def)
+
+lemma True_eqvt:
+  shows "p \<bullet> True = True"
+  unfolding permute_bool_def ..
+
+lemma False_eqvt:
+  shows "p \<bullet> False = False"
+  unfolding permute_bool_def ..
+
+lemma imp_eqvt:
+  shows "p \<bullet> (A \<longrightarrow> B) = ((p \<bullet> A) \<longrightarrow> (p \<bullet> B))"
+  by (simp add: permute_bool_def)
+
+lemma conj_eqvt:
+  shows "p \<bullet> (A \<and> B) = ((p \<bullet> A) \<and> (p \<bullet> B))"
+  by (simp add: permute_bool_def)
+
+lemma disj_eqvt:
+  shows "p \<bullet> (A \<or> B) = ((p \<bullet> A) \<or> (p \<bullet> B))"
+  by (simp add: permute_bool_def)
+
+lemma Not_eqvt:
+  shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))"
+  by (simp add: permute_bool_def)
+
+lemma all_eqvt:
+  shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)"
+  unfolding permute_fun_def permute_bool_def
+  by (auto, drule_tac x="p \<bullet> x" in spec, simp)
+
+lemma all_eqvt2:
+  shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. p \<bullet> P (- p \<bullet> x))"
+  unfolding permute_fun_def permute_bool_def
+  by (auto, drule_tac x="p \<bullet> x" in spec, simp)
+
+lemma ex_eqvt:
+  shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)"
+  unfolding permute_fun_def permute_bool_def
+  by (auto, rule_tac x="p \<bullet> x" in exI, simp)
+
+lemma ex_eqvt2:
+  shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. p \<bullet> P (- p \<bullet> x))"
+  unfolding permute_fun_def permute_bool_def
+  by (auto, rule_tac x="p \<bullet> x" in exI, simp)
+
+lemma ex1_eqvt:
+  shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. (p \<bullet> P) x)"
+  unfolding Ex1_def 
+  by (simp add: ex_eqvt permute_fun_def conj_eqvt all_eqvt imp_eqvt eq_eqvt)
+
+lemma ex1_eqvt2:
+  shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. p \<bullet> P (- p \<bullet> x))"
+  unfolding Ex1_def ex_eqvt2 conj_eqvt all_eqvt2 imp_eqvt eq_eqvt
+  by simp
+
+lemma the_eqvt:
+  assumes unique: "\<exists>!x. P x"
+  shows "(p \<bullet> (THE x. P x)) = (THE x. p \<bullet> P (- p \<bullet> x))"
+  apply(rule the1_equality [symmetric])
+  apply(simp add: ex1_eqvt2[symmetric])
+  apply(simp add: permute_bool_def unique)
+  apply(simp add: permute_bool_def)
+  apply(rule theI'[OF unique])
+  done
+
+section {* Set Operations *}
+
+lemma mem_permute_iff:
+  shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X"
+unfolding mem_def permute_fun_def permute_bool_def
+by simp
+
+lemma mem_eqvt:
+  shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
+  unfolding mem_permute_iff permute_bool_def by simp
+
+lemma not_mem_eqvt:
+  shows "p \<bullet> (x \<notin> A) \<longleftrightarrow> (p \<bullet> x) \<notin> (p \<bullet> A)"
+  unfolding mem_def permute_fun_def by (simp add: Not_eqvt)
+
+lemma Collect_eqvt:
+  shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
+  unfolding Collect_def permute_fun_def ..
+
+lemma Collect_eqvt2:
+  shows "p \<bullet> {x. P x} = {x. p \<bullet> (P (-p \<bullet> x))}"
+  unfolding Collect_def permute_fun_def ..
+
+lemma empty_eqvt:
+  shows "p \<bullet> {} = {}"
+  unfolding empty_def Collect_eqvt2 False_eqvt ..
+
+lemma supp_set_empty:
+  shows "supp {} = {}"
+  by (simp add: supp_def empty_eqvt)
+
+lemma fresh_set_empty:
+  shows "a \<sharp> {}"
+  by (simp add: fresh_def supp_set_empty)
+
+lemma UNIV_eqvt:
+  shows "p \<bullet> UNIV = UNIV"
+  unfolding UNIV_def Collect_eqvt2 True_eqvt ..
+
+lemma union_eqvt:
+  shows "p \<bullet> (A \<union> B) = (p \<bullet> A) \<union> (p \<bullet> B)"
+  unfolding Un_def Collect_eqvt2 disj_eqvt mem_eqvt by simp
+
+lemma inter_eqvt:
+  shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)"
+  unfolding Int_def Collect_eqvt2 conj_eqvt mem_eqvt by simp
+
+lemma Diff_eqvt:
+  fixes A B :: "'a::pt set"
+  shows "p \<bullet> (A - B) = p \<bullet> A - p \<bullet> B"
+  unfolding set_diff_eq Collect_eqvt2 conj_eqvt Not_eqvt mem_eqvt by simp
+
+lemma Compl_eqvt:
+  fixes A :: "'a::pt set"
+  shows "p \<bullet> (- A) = - (p \<bullet> A)"
+  unfolding Compl_eq_Diff_UNIV Diff_eqvt UNIV_eqvt ..
+
+lemma insert_eqvt:
+  shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)"
+  unfolding permute_set_eq_image image_insert ..
+
+lemma vimage_eqvt:
+  shows "p \<bullet> (f -` A) = (p \<bullet> f) -` (p \<bullet> A)"
+  unfolding vimage_def permute_fun_def [where f=f]
+  unfolding Collect_eqvt2 mem_eqvt ..
+
+lemma image_eqvt:
+  shows "p \<bullet> (f ` A) = (p \<bullet> f) ` (p \<bullet> A)"
+  unfolding permute_set_eq_image
+  unfolding permute_fun_def [where f=f]
+  by (simp add: image_image)
+
+lemma finite_permute_iff:
+  shows "finite (p \<bullet> A) \<longleftrightarrow> finite A"
+  unfolding permute_set_eq_vimage
+  using bij_permute by (rule finite_vimage_iff)
+
+lemma finite_eqvt:
+  shows "p \<bullet> finite A = finite (p \<bullet> A)"
+  unfolding finite_permute_iff permute_bool_def ..
+
+
+section {* List Operations *}
+
+lemma append_eqvt:
+  shows "p \<bullet> (xs @ ys) = (p \<bullet> xs) @ (p \<bullet> ys)"
+  by (induct xs) auto
+
+lemma supp_append:
+  shows "supp (xs @ ys) = supp xs \<union> supp ys"
+  by (induct xs) (auto simp add: supp_Nil supp_Cons)
+
+lemma fresh_append:
+  shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
+  by (induct xs) (simp_all add: fresh_Nil fresh_Cons)
+
+lemma rev_eqvt:
+  shows "p \<bullet> (rev xs) = rev (p \<bullet> xs)"
+  by (induct xs) (simp_all add: append_eqvt)
+
+lemma supp_rev:
+  shows "supp (rev xs) = supp xs"
+  by (induct xs) (auto simp add: supp_append supp_Cons supp_Nil)
+
+lemma fresh_rev:
+  shows "a \<sharp> rev xs \<longleftrightarrow> a \<sharp> xs"
+  by (induct xs) (auto simp add: fresh_append fresh_Cons fresh_Nil)
+
+lemma set_eqvt:
+  shows "p \<bullet> (set xs) = set (p \<bullet> xs)"
+  by (induct xs) (simp_all add: empty_eqvt insert_eqvt)
+
+(* needs finite support premise
+lemma supp_set:
+  fixes x :: "'a::pt"
+  shows "supp (set xs) = supp xs"
+*)
+
+
+section {* Product Operations *}
+
+lemma fst_eqvt:
+  "p \<bullet> (fst x) = fst (p \<bullet> x)"
+ by (cases x) simp
+
+lemma snd_eqvt:
+  "p \<bullet> (snd x) = snd (p \<bullet> x)"
+ by (cases x) simp
+
+
+section {* Units *}
+
+lemma supp_unit:
+  shows "supp () = {}"
+  by (simp add: supp_def)
+
+lemma fresh_unit:
+  shows "a \<sharp> ()"
+  by (simp add: fresh_def supp_unit)
+
+section {* Equivariance automation *}
+
+text {* Setup of the theorem attributes @{text eqvt} and @{text eqvt_force} *}
+
+use "nominal_thmdecls.ML"
+setup "Nominal_ThmDecls.setup"
+
+lemmas [eqvt] = 
+  (* connectives *)
+  eq_eqvt if_eqvt imp_eqvt disj_eqvt conj_eqvt Not_eqvt 
+  True_eqvt False_eqvt ex_eqvt all_eqvt ex1_eqvt
+  imp_eqvt [folded induct_implies_def]
+
+  (* nominal *)
+  permute_eqvt supp_eqvt fresh_eqvt
+  permute_pure
+
+  (* datatypes *)
+  permute_prod.simps append_eqvt rev_eqvt set_eqvt
+  fst_eqvt snd_eqvt
+
+  (* sets *)
+  empty_eqvt UNIV_eqvt union_eqvt inter_eqvt mem_eqvt
+  Diff_eqvt Compl_eqvt insert_eqvt Collect_eqvt
+
+thm eqvts
+thm eqvts_raw
+
+text {* helper lemmas for the eqvt_tac *}
+
+definition
+  "unpermute p = permute (- p)"
+
+lemma eqvt_apply:
+  fixes f :: "'a::pt \<Rightarrow> 'b::pt" 
+  and x :: "'a::pt"
+  shows "p \<bullet> (f x) \<equiv> (p \<bullet> f) (p \<bullet> x)"
+  unfolding permute_fun_def by simp
+
+lemma eqvt_lambda:
+  fixes f :: "'a::pt \<Rightarrow> 'b::pt"
+  shows "p \<bullet> (\<lambda>x. f x) \<equiv> (\<lambda>x. p \<bullet> (f (unpermute p x)))"
+  unfolding permute_fun_def unpermute_def by simp
+
+lemma eqvt_bound:
+  shows "p \<bullet> unpermute p x \<equiv> x"
+  unfolding unpermute_def by simp
+
+use "nominal_permeq.ML"
+
+
+lemma "p \<bullet> (A \<longrightarrow> B = C)"
+apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *}) 
+oops
+
+lemma "p \<bullet> (\<lambda>(x::'a::pt). A \<longrightarrow> (B::'a \<Rightarrow> bool) x = C) = foo"
+apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
+oops
+
+lemma "p \<bullet> (\<lambda>x y. \<exists>z. x = z \<and> x = y \<longrightarrow> z \<noteq> x) = foo"
+apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
+oops
+
+lemma "p \<bullet> (\<lambda>f x. f (g (f x))) = foo"
+apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
+oops
+
+lemma "p \<bullet> (\<lambda>q. q \<bullet> (r \<bullet> x)) = foo"
+apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
+oops
+
+lemma "p \<bullet> (q \<bullet> r \<bullet> x) = foo"
+apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
+oops
+
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Nominal2_Supp.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,375 @@
+(*  Title:      Nominal2_Supp
+    Authors:    Brian Huffman, Christian Urban
+
+    Supplementary Lemmas and Definitions for 
+    Nominal Isabelle. 
+*)
+theory Nominal2_Supp
+imports Nominal2_Base Nominal2_Eqvt Nominal2_Atoms
+begin
+
+
+section {* Fresh-Star *}
+
+text {* The fresh-star generalisation of fresh is used in strong
+  induction principles. *}
+
+definition 
+  fresh_star :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp>* _" [80,80] 80)
+where 
+  "xs \<sharp>* c \<equiv> \<forall>x \<in> xs. x \<sharp> c"
+
+lemma fresh_star_prod:
+  fixes xs::"atom set"
+  shows "xs \<sharp>* (a, b) = (xs \<sharp>* a \<and> xs \<sharp>* b)"
+  by (auto simp add: fresh_star_def fresh_Pair)
+
+lemma fresh_star_union:
+  shows "(xs \<union> ys) \<sharp>* c = (xs \<sharp>* c \<and> ys \<sharp>* c)"
+  by (auto simp add: fresh_star_def)
+
+lemma fresh_star_insert:
+  shows "(insert x ys) \<sharp>* c = (x \<sharp> c \<and> ys \<sharp>* c)"
+  by (auto simp add: fresh_star_def)
+
+lemma fresh_star_Un_elim:
+  "((S \<union> T) \<sharp>* c \<Longrightarrow> PROP C) \<equiv> (S \<sharp>* c \<Longrightarrow> T \<sharp>* c \<Longrightarrow> PROP C)"
+  unfolding fresh_star_def
+  apply(rule)
+  apply(erule meta_mp)
+  apply(auto)
+  done
+
+lemma fresh_star_insert_elim:
+  "(insert x S \<sharp>* c \<Longrightarrow> PROP C) \<equiv> (x \<sharp> c \<Longrightarrow> S \<sharp>* c \<Longrightarrow> PROP C)"
+  unfolding fresh_star_def
+  by rule (simp_all add: fresh_star_def)
+
+lemma fresh_star_empty_elim:
+  "({} \<sharp>* c \<Longrightarrow> PROP C) \<equiv> PROP C"
+  by (simp add: fresh_star_def)
+
+lemma fresh_star_unit_elim: 
+  shows "(a \<sharp>* () \<Longrightarrow> PROP C) \<equiv> PROP C"
+  by (simp add: fresh_star_def fresh_unit) 
+
+lemma fresh_star_prod_elim: 
+  shows "(a \<sharp>* (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp>* x \<Longrightarrow> a \<sharp>* y \<Longrightarrow> PROP C)"
+  by (rule, simp_all add: fresh_star_prod)
+
+
+section {* Avoiding of atom sets *}
+
+text {* 
+  For every set of atoms, there is another set of atoms
+  avoiding a finitely supported c and there is a permutation
+  which 'translates' between both sets.
+*}
+
+lemma at_set_avoiding_aux:
+  fixes Xs::"atom set"
+  and   As::"atom set"
+  assumes b: "Xs \<subseteq> As"
+  and     c: "finite As"
+  shows "\<exists>p. (p \<bullet> Xs) \<inter> As = {} \<and> (supp p) \<subseteq> (Xs \<union> (p \<bullet> Xs))"
+proof -
+  from b c have "finite Xs" by (rule finite_subset)
+  then show ?thesis using b
+  proof (induct rule: finite_subset_induct)
+    case empty
+    have "0 \<bullet> {} \<inter> As = {}" by simp
+    moreover
+    have "supp (0::perm) \<subseteq> {} \<union> 0 \<bullet> {}" by (simp add: supp_zero_perm)
+    ultimately show ?case by blast
+  next
+    case (insert x Xs)
+    then obtain p where
+      p1: "(p \<bullet> Xs) \<inter> As = {}" and 
+      p2: "supp p \<subseteq> (Xs \<union> (p \<bullet> Xs))" by blast
+    from `x \<in> As` p1 have "x \<notin> p \<bullet> Xs" by fast
+    with `x \<notin> Xs` p2 have "x \<notin> supp p" by fast
+    hence px: "p \<bullet> x = x" unfolding supp_perm by simp
+    have "finite (As \<union> p \<bullet> Xs)"
+      using `finite As` `finite Xs`
+      by (simp add: permute_set_eq_image)
+    then obtain y where "y \<notin> (As \<union> p \<bullet> Xs)" "sort_of y = sort_of x"
+      by (rule obtain_atom)
+    hence y: "y \<notin> As" "y \<notin> p \<bullet> Xs" "sort_of y = sort_of x"
+      by simp_all
+    let ?q = "(x \<rightleftharpoons> y) + p"
+    have q: "?q \<bullet> insert x Xs = insert y (p \<bullet> Xs)"
+      unfolding insert_eqvt
+      using `p \<bullet> x = x` `sort_of y = sort_of x`
+      using `x \<notin> p \<bullet> Xs` `y \<notin> p \<bullet> Xs`
+      by (simp add: swap_atom swap_set_not_in)
+    have "?q \<bullet> insert x Xs \<inter> As = {}"
+      using `y \<notin> As` `p \<bullet> Xs \<inter> As = {}`
+      unfolding q by simp
+    moreover
+    have "supp ?q \<subseteq> insert x Xs \<union> ?q \<bullet> insert x Xs"
+      using p2 unfolding q
+      apply (intro subset_trans [OF supp_plus_perm])
+      apply (auto simp add: supp_swap)
+      done
+    ultimately show ?case by blast
+  qed
+qed
+
+lemma at_set_avoiding:
+  assumes a: "finite Xs"
+  and     b: "finite (supp c)"
+  obtains p::"perm" where "(p \<bullet> Xs)\<sharp>*c" and "(supp p) \<subseteq> (Xs \<union> (p \<bullet> Xs))"
+  using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \<union> supp c"]
+  unfolding fresh_star_def fresh_def by blast
+
+
+section {* The freshness lemma according to Andrew Pitts *}
+
+lemma fresh_conv_MOST: 
+  shows "a \<sharp> x \<longleftrightarrow> (MOST b. (a \<rightleftharpoons> b) \<bullet> x = x)"
+  unfolding fresh_def supp_def MOST_iff_cofinite by simp
+
+lemma fresh_apply:
+  assumes "a \<sharp> f" and "a \<sharp> x" 
+  shows "a \<sharp> f x"
+  using assms unfolding fresh_conv_MOST
+  unfolding permute_fun_app_eq [where f=f]
+  by (elim MOST_rev_mp, simp)
+
+lemma freshness_lemma:
+  fixes h :: "'a::at \<Rightarrow> 'b::pt"
+  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
+  shows  "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
+proof -
+  from a obtain b where a1: "atom b \<sharp> h" and a2: "atom b \<sharp> h b"
+    by (auto simp add: fresh_Pair)
+  show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
+  proof (intro exI allI impI)
+    fix a :: 'a
+    assume a3: "atom a \<sharp> h"
+    show "h a = h b"
+    proof (cases "a = b")
+      assume "a = b"
+      thus "h a = h b" by simp
+    next
+      assume "a \<noteq> b"
+      hence "atom a \<sharp> b" by (simp add: fresh_at_base)
+      with a3 have "atom a \<sharp> h b" by (rule fresh_apply)
+      with a2 have d1: "(atom b \<rightleftharpoons> atom a) \<bullet> (h b) = (h b)"
+        by (rule swap_fresh_fresh)
+      from a1 a3 have d2: "(atom b \<rightleftharpoons> atom a) \<bullet> h = h"
+        by (rule swap_fresh_fresh)
+      from d1 have "h b = (atom b \<rightleftharpoons> atom a) \<bullet> (h b)" by simp
+      also have "\<dots> = ((atom b \<rightleftharpoons> atom a) \<bullet> h) ((atom b \<rightleftharpoons> atom a) \<bullet> b)"
+        by (rule permute_fun_app_eq)
+      also have "\<dots> = h a"
+        using d2 by simp
+      finally show "h a = h b"  by simp
+    qed
+  qed
+qed
+
+lemma freshness_lemma_unique:
+  fixes h :: "'a::at \<Rightarrow> 'b::pt"
+  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
+  shows "\<exists>!x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
+proof (rule ex_ex1I)
+  from a show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
+    by (rule freshness_lemma)
+next
+  fix x y
+  assume x: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
+  assume y: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = y"
+  from a x y show "x = y"
+    by (auto simp add: fresh_Pair)
+qed
+
+text {* packaging the freshness lemma into a function *}
+
+definition
+  fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
+where
+  "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)"
+
+lemma fresh_fun_app:
+  fixes h :: "'a::at \<Rightarrow> 'b::pt"
+  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
+  assumes b: "atom a \<sharp> h"
+  shows "fresh_fun h = h a"
+unfolding fresh_fun_def
+proof (rule the_equality)
+  show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a"
+  proof (intro strip)
+    fix a':: 'a
+    assume c: "atom a' \<sharp> h"
+    from a have "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" by (rule freshness_lemma)
+    with b c show "h a' = h a" by auto
+  qed
+next
+  fix fr :: 'b
+  assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr"
+  with b show "fr = h a" by auto
+qed
+
+lemma fresh_fun_app':
+  fixes h :: "'a::at \<Rightarrow> 'b::pt"
+  assumes a: "atom a \<sharp> h" "atom a \<sharp> h a"
+  shows "fresh_fun h = h a"
+  apply (rule fresh_fun_app)
+  apply (auto simp add: fresh_Pair intro: a)
+  done
+
+lemma fresh_fun_eqvt:
+  fixes h :: "'a::at \<Rightarrow> 'b::pt"
+  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
+  shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)"
+  using a
+  apply (clarsimp simp add: fresh_Pair)
+  apply (subst fresh_fun_app', assumption+)
+  apply (drule fresh_permute_iff [where p=p, THEN iffD2])
+  apply (drule fresh_permute_iff [where p=p, THEN iffD2])
+  apply (simp add: atom_eqvt permute_fun_app_eq [where f=h])
+  apply (erule (1) fresh_fun_app' [symmetric])
+  done
+
+lemma fresh_fun_supports:
+  fixes h :: "'a::at \<Rightarrow> 'b::pt"
+  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
+  shows "(supp h) supports (fresh_fun h)"
+  apply (simp add: supports_def fresh_def [symmetric])
+  apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh)
+  done
+
+notation fresh_fun (binder "FRESH " 10)
+
+lemma FRESH_f_iff:
+  fixes P :: "'a::at \<Rightarrow> 'b::pure"
+  fixes f :: "'b \<Rightarrow> 'c::pure"
+  assumes P: "finite (supp P)"
+  shows "(FRESH x. f (P x)) = f (FRESH x. P x)"
+proof -
+  obtain a::'a where "atom a \<notin> supp P"
+    using P by (rule obtain_at_base)
+  hence "atom a \<sharp> P"
+    by (simp add: fresh_def)
+  show "(FRESH x. f (P x)) = f (FRESH x. P x)"
+    apply (subst fresh_fun_app' [where a=a, OF _ pure_fresh])
+    apply (cut_tac `atom a \<sharp> P`)
+    apply (simp add: fresh_conv_MOST)
+    apply (elim MOST_rev_mp, rule MOST_I, clarify)
+    apply (simp add: permute_fun_def permute_pure expand_fun_eq)
+    apply (subst fresh_fun_app' [where a=a, OF `atom a \<sharp> P` pure_fresh])
+    apply (rule refl)
+    done
+qed
+
+lemma FRESH_binop_iff:
+  fixes P :: "'a::at \<Rightarrow> 'b::pure"
+  fixes Q :: "'a::at \<Rightarrow> 'c::pure"
+  fixes binop :: "'b \<Rightarrow> 'c \<Rightarrow> 'd::pure"
+  assumes P: "finite (supp P)" 
+  and     Q: "finite (supp Q)"
+  shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)"
+proof -
+  from assms have "finite (supp P \<union> supp Q)" by simp
+  then obtain a::'a where "atom a \<notin> (supp P \<union> supp Q)"
+    by (rule obtain_at_base)
+  hence "atom a \<sharp> P" and "atom a \<sharp> Q"
+    by (simp_all add: fresh_def)
+  show ?thesis
+    apply (subst fresh_fun_app' [where a=a, OF _ pure_fresh])
+    apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`)
+    apply (simp add: fresh_conv_MOST)
+    apply (elim MOST_rev_mp, rule MOST_I, clarify)
+    apply (simp add: permute_fun_def permute_pure expand_fun_eq)
+    apply (subst fresh_fun_app' [where a=a, OF `atom a \<sharp> P` pure_fresh])
+    apply (subst fresh_fun_app' [where a=a, OF `atom a \<sharp> Q` pure_fresh])
+    apply (rule refl)
+    done
+qed
+
+lemma FRESH_conj_iff:
+  fixes P Q :: "'a::at \<Rightarrow> bool"
+  assumes P: "finite (supp P)" and Q: "finite (supp Q)"
+  shows "(FRESH x. P x \<and> Q x) \<longleftrightarrow> (FRESH x. P x) \<and> (FRESH x. Q x)"
+using P Q by (rule FRESH_binop_iff)
+
+lemma FRESH_disj_iff:
+  fixes P Q :: "'a::at \<Rightarrow> bool"
+  assumes P: "finite (supp P)" and Q: "finite (supp Q)"
+  shows "(FRESH x. P x \<or> Q x) \<longleftrightarrow> (FRESH x. P x) \<or> (FRESH x. Q x)"
+using P Q by (rule FRESH_binop_iff)
+
+
+section {* An example of a function without finite support *}
+
+primrec
+  nat_of :: "atom \<Rightarrow> nat"
+where
+  "nat_of (Atom s n) = n"
+
+lemma atom_eq_iff:
+  fixes a b :: atom
+  shows "a = b \<longleftrightarrow> sort_of a = sort_of b \<and> nat_of a = nat_of b"
+  by (induct a, induct b, simp)
+
+lemma not_fresh_nat_of:
+  shows "\<not> a \<sharp> nat_of"
+unfolding fresh_def supp_def
+proof (clarsimp)
+  assume "finite {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"
+  hence "finite ({a} \<union> {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of})"
+    by simp
+  then obtain b where
+    b1: "b \<noteq> a" and
+    b2: "sort_of b = sort_of a" and
+    b3: "(a \<rightleftharpoons> b) \<bullet> nat_of = nat_of"
+    by (rule obtain_atom) auto
+  have "nat_of a = (a \<rightleftharpoons> b) \<bullet> (nat_of a)" by (simp add: permute_nat_def)
+  also have "\<dots> = ((a \<rightleftharpoons> b) \<bullet> nat_of) ((a \<rightleftharpoons> b) \<bullet> a)" by (simp add: permute_fun_app_eq)
+  also have "\<dots> = nat_of ((a \<rightleftharpoons> b) \<bullet> a)" using b3 by simp
+  also have "\<dots> = nat_of b" using b2 by simp
+  finally have "nat_of a = nat_of b" by simp
+  with b2 have "a = b" by (simp add: atom_eq_iff)
+  with b1 show "False" by simp
+qed
+
+lemma supp_nat_of:
+  shows "supp nat_of = UNIV"
+  using not_fresh_nat_of [unfolded fresh_def] by auto
+
+
+section {* Support for sets of atoms *}
+
+lemma supp_finite_atom_set:
+  fixes S::"atom set"
+  assumes "finite S"
+  shows "supp S = S"
+  apply(rule finite_supp_unique)
+  apply(simp add: supports_def)
+  apply(simp add: swap_set_not_in)
+  apply(rule assms)
+  apply(simp add: swap_set_in)
+done
+
+
+(*
+lemma supp_infinite:
+  fixes S::"atom set"
+  assumes asm: "finite (UNIV - S)"
+  shows "(supp S) = (UNIV - S)"
+apply(rule finite_supp_unique)
+apply(auto simp add: supports_def permute_set_eq swap_atom)[1]
+apply(rule asm)
+apply(auto simp add: permute_set_eq swap_atom)[1]
+done
+
+lemma supp_infinite_coinfinite:
+  fixes S::"atom set"
+  assumes asm1: "infinite S"
+  and     asm2: "infinite (UNIV-S)"
+  shows "(supp S) = (UNIV::atom set)"
+*)
+
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Perm.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,147 @@
+theory Perm
+imports "Nominal2_Atoms"
+begin
+
+ML {*
+  open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *)
+  fun permute ty = Const (@{const_name permute}, @{typ perm} --> ty --> ty);
+  val minus_perm = Const (@{const_name minus}, @{typ perm} --> @{typ perm});
+*}
+
+ML {*
+fun prove_perm_empty lthy induct perm_def perm_frees =
+let
+  val perm_types = map fastype_of perm_frees;
+  val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types);
+  val gl =
+    HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+      (map (fn ((perm, T), x) => HOLogic.mk_eq
+          (perm $ @{term "0 :: perm"} $ Free (x, T),
+           Free (x, T)))
+       (perm_frees ~~
+        map body_type perm_types ~~ perm_indnames)));
+  fun tac _ =
+    EVERY [
+      indtac induct perm_indnames 1,
+      ALLGOALS (asm_full_simp_tac (HOL_ss addsimps (@{thm permute_zero} :: perm_def)))
+    ];
+in
+  split_conj_thm (Goal.prove lthy perm_indnames [] gl tac)
+end;
+*}
+
+ML {*
+fun prove_perm_append lthy induct perm_def perm_frees =
+let
+  val add_perm = @{term "op + :: (perm \<Rightarrow> perm \<Rightarrow> perm)"}
+  val pi1 = Free ("pi1", @{typ perm});
+  val pi2 = Free ("pi2", @{typ perm});
+  val perm_types = map fastype_of perm_frees
+  val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types);
+  val gl =
+    (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
+      (map (fn ((perm, T), x) =>
+          let
+            val lhs = perm $ (add_perm $ pi1 $ pi2) $ Free (x, T)
+            val rhs = perm $ pi1 $ (perm $ pi2 $ Free (x, T))
+          in HOLogic.mk_eq (lhs, rhs)
+          end)
+        (perm_frees ~~ map body_type perm_types ~~ perm_indnames))))
+  fun tac _ =
+    EVERY [
+      indtac induct perm_indnames 1,
+      ALLGOALS (asm_full_simp_tac (HOL_ss addsimps (@{thm permute_plus} :: perm_def)))
+    ]
+in
+  split_conj_thm (Goal.prove lthy ("pi1" :: "pi2" :: perm_indnames) [] gl tac)
+end;
+*}
+
+ML {*
+(* TODO: full_name can be obtained from new_type_names with Datatype *)
+fun define_raw_perms new_type_names full_tnames thy =
+let
+  val {descr, induct, ...} = Datatype.the_info thy (hd full_tnames);
+  (* TODO: [] should be the sorts that we'll take from the specification *)
+  val sorts = [];
+  fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
+  val perm_names' = Datatype_Prop.indexify_names (map (fn (i, _) =>
+    "permute_" ^ name_of_typ (nth_dtyp i)) descr);
+  val perm_types = map (fn (i, _) =>
+    let val T = nth_dtyp i
+    in @{typ perm} --> T --> T end) descr;
+  val perm_names_types' = perm_names' ~~ perm_types;
+  val pi = Free ("pi", @{typ perm});
+  fun perm_eq_constr i (cname, dts) =
+    let
+      val Ts = map (typ_of_dtyp descr sorts) dts;
+      val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
+      val args = map Free (names ~~ Ts);
+      val c = Const (cname, Ts ---> (nth_dtyp i));
+      fun perm_arg (dt, x) =
+        let val T = type_of x
+        in
+          if is_rec_type dt then
+            let val (Us, _) = strip_type T
+            in list_abs (map (pair "x") Us,
+              Free (nth perm_names_types' (body_index dt)) $ pi $
+                list_comb (x, map (fn (i, U) =>
+                  (permute U) $ (minus_perm $ pi) $ Bound i)
+                  ((length Us - 1 downto 0) ~~ Us)))
+            end
+          else (permute T) $ pi $ x
+        end;
+    in
+      (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
+        (Free (nth perm_names_types' i) $
+           Free ("pi", @{typ perm}) $ list_comb (c, args),
+         list_comb (c, map perm_arg (dts ~~ args)))))
+    end;
+    fun perm_eq (i, (_, _, constrs)) = map (perm_eq_constr i) constrs;
+    val perm_eqs = maps perm_eq descr;
+    val lthy =
+      Theory_Target.instantiation (full_tnames, [], @{sort pt}) thy;
+    (* TODO: Use the version of prmrec that gives the names explicitely. *)
+    val ((_, perm_ldef), lthy') =
+      Primrec.add_primrec
+        (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs lthy;
+    val perm_frees =
+      (distinct (op =)) (map (fst o strip_comb o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) perm_ldef);
+    val perm_empty_thms = List.take (prove_perm_empty lthy' induct perm_ldef perm_frees, length new_type_names);
+    val perm_append_thms = List.take (prove_perm_append lthy' induct perm_ldef perm_frees, length new_type_names)
+    val perms_name = space_implode "_" perm_names'
+    val perms_zero_bind = Binding.name (perms_name ^ "_zero")
+    val perms_append_bind = Binding.name (perms_name ^ "_append")
+    fun tac _ perm_thms =
+      (Class.intro_classes_tac []) THEN (ALLGOALS (
+        simp_tac (HOL_ss addsimps perm_thms
+      )));
+    fun morphism phi = map (Morphism.thm phi);
+  in
+  lthy'
+  |> snd o (Local_Theory.note ((perms_zero_bind, []), perm_empty_thms))
+  |> snd o (Local_Theory.note ((perms_append_bind, []), perm_append_thms))
+  |> Class_Target.prove_instantiation_exit_result morphism tac (perm_empty_thms @ perm_append_thms)
+  end
+
+*}
+
+(* Test
+atom_decl name
+
+datatype rtrm1 =
+  rVr1 "name"
+| rAp1 "rtrm1" "rtrm1 list"
+| rLm1 "name" "rtrm1"
+| rLt1 "bp" "rtrm1" "rtrm1"
+and bp =
+  BUnit
+| BVr "name"
+| BPr "bp" "bp"
+
+
+setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Perm.rtrm1", "Perm.bp"] *}
+print_theorems
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Rsp.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,118 @@
+theory Rsp
+imports Abs
+begin
+
+ML {*
+fun define_quotient_type args tac ctxt =
+let
+  val mthd = Method.SIMPLE_METHOD tac
+  val mthdt = Method.Basic (fn _ => mthd)
+  val bymt = Proof.global_terminal_proof (mthdt, NONE)
+in
+  bymt (Quotient_Type.quotient_type args ctxt)
+end
+*}
+
+ML {*
+fun const_rsp lthy const =
+let
+  val nty = fastype_of (Quotient_Term.quotient_lift_const ("", const) lthy)
+  val rel = Quotient_Term.equiv_relation_chk lthy (fastype_of const, nty);
+in
+  HOLogic.mk_Trueprop (rel $ const $ const)
+end
+*}
+
+(* Replaces bounds by frees and meta implications by implications *)
+ML {*
+fun prepare_goal trm =
+let
+  val vars = strip_all_vars trm
+  val fs = rev (map Free vars)
+  val (fixes, no_alls) = ((map fst vars), subst_bounds (fs, (strip_all_body trm)))
+  val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems no_alls)
+  val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl no_alls)
+in
+  (fixes, fold (curry HOLogic.mk_imp) prems concl)
+end
+*}
+
+ML {*
+fun get_rsp_goal thy trm =
+let
+  val goalstate = Goal.init (cterm_of thy trm);
+  val tac = REPEAT o rtac @{thm fun_rel_id};
+in
+  case (SINGLE (tac 1) goalstate) of
+    NONE => error "rsp_goal failed"
+  | SOME th => prepare_goal (term_of (cprem_of th 1))
+end
+*}
+
+ML {*
+fun repeat_mp thm = repeat_mp (mp OF [thm]) handle THM _ => thm
+*}
+
+ML {*
+fun prove_const_rsp bind consts tac ctxt =
+let
+  val rsp_goals = map (const_rsp ctxt) consts
+  val thy = ProofContext.theory_of ctxt
+  val (fixed, user_goals) = split_list (map (get_rsp_goal thy) rsp_goals)
+  val fixed' = distinct (op =) (flat fixed)
+  val user_goal = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj user_goals)
+  val user_thm = Goal.prove ctxt fixed' [] user_goal tac
+  val user_thms = map repeat_mp (HOLogic.conj_elims user_thm)
+  fun tac _ = (REPEAT o rtac @{thm fun_rel_id} THEN' resolve_tac user_thms THEN_ALL_NEW atac) 1
+  val rsp_thms = map (fn gl => Goal.prove ctxt [] [] gl tac) rsp_goals
+in
+   ctxt
+|> snd o Local_Theory.note 
+  ((Binding.empty, [Attrib.internal (fn _ => Quotient_Info.rsp_rules_add)]), rsp_thms)
+|> snd o Local_Theory.note ((bind, []), user_thms)
+end
+*}
+
+
+
+ML {*
+fun fvbv_rsp_tac induct fvbv_simps =
+  ((((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
+  (TRY o rtac @{thm TrueI})) THEN_ALL_NEW asm_full_simp_tac
+  (HOL_ss addsimps (@{thm alpha_gen} :: fvbv_simps)))
+*}
+
+ML {*
+fun constr_rsp_tac inj rsp equivps =
+let
+  val reflps = map (fn x => @{thm equivp_reflp} OF [x]) equivps
+in
+  REPEAT o rtac impI THEN'
+  simp_tac (HOL_ss addsimps inj) THEN'
+  (TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI)) THEN_ALL_NEW
+  (asm_simp_tac HOL_ss THEN_ALL_NEW (
+   rtac @{thm exI[of _ "0 :: perm"]} THEN'
+   asm_full_simp_tac (HOL_ss addsimps (rsp @ reflps @
+     @{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv}))
+  ))
+end
+*}
+
+(* Testing code
+local_setup {* prove_const_rsp @{binding fv_rtrm2_rsp} [@{term rbv2}]
+  (fn _ => fv_rsp_tac @{thm alpha_rtrm2_alpha_rassign.inducts(2)} @{thms fv_rtrm2_fv_rassign.simps} 1) *}*)
+
+(*ML {*
+  val rsp_goals = map (const_rsp @{context}) [@{term rbv2}]
+  val (fixed, user_goals) = split_list (map (get_rsp_goal @{theory}) rsp_goals)
+  val fixed' = distinct (op =) (flat fixed)
+  val user_goal = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj user_goals)
+*}
+prove ug: {* user_goal *}
+ML_prf {*
+val induct = @{thm alpha_rtrm2_alpha_rassign.inducts(2)}
+val fv_simps = @{thms rbv2.simps}
+*} 
+*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Terms.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,1043 @@
+theory Terms
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../../Attic/Prove"
+begin
+
+atom_decl name
+
+text {* primrec seems to be genarally faster than fun *}
+
+section {*** lets with binding patterns ***}
+
+datatype rtrm1 =
+  rVr1 "name"
+| rAp1 "rtrm1" "rtrm1"
+| rLm1 "name" "rtrm1"        --"name is bound in trm1"
+| rLt1 "bp" "rtrm1" "rtrm1"   --"all variables in bp are bound in the 2nd trm1"
+and bp =
+  BUnit
+| BVr "name"
+| BPr "bp" "bp"
+
+print_theorems
+
+(* to be given by the user *)
+
+primrec 
+  bv1
+where
+  "bv1 (BUnit) = {}"
+| "bv1 (BVr x) = {atom x}"
+| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)"
+
+setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Terms.rtrm1", "Terms.bp"] *}
+thm permute_rtrm1_permute_bp.simps
+
+local_setup {*
+  snd o define_fv_alpha "Terms.rtrm1"
+  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(SOME @{term bv1}, 0)], [], [(SOME @{term bv1}, 0)]]],
+  [[], [[]], [[], []]]] *}
+
+notation
+  alpha_rtrm1 ("_ \<approx>1 _" [100, 100] 100) and
+  alpha_bp ("_ \<approx>1b _" [100, 100] 100)
+thm alpha_rtrm1_alpha_bp.intros
+thm fv_rtrm1_fv_bp.simps
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_inj}, []), (build_alpha_inj @{thms alpha_rtrm1_alpha_bp.intros} @{thms rtrm1.distinct rtrm1.inject bp.distinct bp.inject} @{thms alpha_rtrm1.cases alpha_bp.cases} ctxt)) ctxt)) *}
+thm alpha1_inj
+
+lemma alpha_bp_refl: "alpha_bp a a"
+apply induct
+apply (simp_all  add: alpha1_inj)
+done
+
+lemma alpha_bp_eq_eq: "alpha_bp a b = (a = b)"
+apply rule
+apply (induct a b rule: alpha_rtrm1_alpha_bp.inducts(2))
+apply (simp_all add: alpha_bp_refl)
+done
+
+lemma alpha_bp_eq: "alpha_bp = (op =)"
+apply (rule ext)+
+apply (rule alpha_bp_eq_eq)
+done
+
+lemma bv1_eqvt[eqvt]:
+  shows "(pi \<bullet> bv1 x) = bv1 (pi \<bullet> x)"
+  apply (induct x)
+  apply (simp_all add: atom_eqvt eqvts)
+  done
+
+lemma fv_rtrm1_eqvt[eqvt]:
+    "(pi\<bullet>fv_rtrm1 t) = fv_rtrm1 (pi\<bullet>t)"
+    "(pi\<bullet>fv_bp b) = fv_bp (pi\<bullet>b)"
+  apply (induct t and b)
+  apply (simp_all add: insert_eqvt atom_eqvt empty_eqvt union_eqvt Diff_eqvt bv1_eqvt)
+  done
+
+lemma alpha1_eqvt:
+  "t \<approx>1 s \<Longrightarrow> (pi \<bullet> t) \<approx>1 (pi \<bullet> s)"
+  "alpha_bp a b \<Longrightarrow> alpha_bp (pi \<bullet> a) (pi \<bullet> b)"
+  apply (induct t s and a b rule: alpha_rtrm1_alpha_bp.inducts)
+  apply (simp_all add:eqvts alpha1_inj)
+  apply (erule exE)
+  apply (rule_tac x="pi \<bullet> pia" in exI)
+  apply (simp add: alpha_gen)
+  apply(erule conjE)+
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
+  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt empty_eqvt fv_rtrm1_eqvt)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: atom_eqvt Diff_eqvt fv_rtrm1_eqvt insert_eqvt empty_eqvt)
+  apply(simp add: permute_eqvt[symmetric])
+  apply (erule exE)
+  apply (erule exE)
+  apply (rule conjI)
+  apply (rule_tac x="pi \<bullet> pia" in exI)
+  apply (simp add: alpha_gen)
+  apply(erule conjE)+
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
+  apply(simp add: fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: atom_eqvt fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
+  apply(simp add: permute_eqvt[symmetric])
+  apply (rule_tac x="pi \<bullet> piaa" in exI)
+  apply (simp add: alpha_gen)
+  apply(erule conjE)+
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
+  apply(simp add: fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: atom_eqvt fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
+  apply(simp add: permute_eqvt[symmetric])
+  done
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_equivp}, []),
+  (build_equivps [@{term alpha_rtrm1}, @{term alpha_bp}] @{thm rtrm1_bp.induct} @{thm alpha_rtrm1_alpha_bp.induct} @{thms rtrm1.inject bp.inject} @{thms alpha1_inj} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} ctxt)) ctxt)) *}
+thm alpha1_equivp
+
+local_setup  {* define_quotient_type [(([], @{binding trm1}, NoSyn), (@{typ rtrm1}, @{term alpha_rtrm1}))]
+  (rtac @{thm alpha1_equivp(1)} 1) *}
+
+local_setup {*
+(fn ctxt => ctxt
+ |> snd o (Quotient_Def.quotient_lift_const ("Vr1", @{term rVr1}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Ap1", @{term rAp1}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lm1", @{term rLm1}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lt1", @{term rLt1}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_trm1", @{term fv_rtrm1})))
+*}
+print_theorems
+
+thm alpha_rtrm1_alpha_bp.induct
+local_setup {* prove_const_rsp @{binding fv_rtrm1_rsp} [@{term fv_rtrm1}]
+  (fn _ => fvbv_rsp_tac @{thm alpha_rtrm1_alpha_bp.inducts(1)} @{thms fv_rtrm1_fv_bp.simps} 1) *}
+local_setup {* prove_const_rsp @{binding rVr1_rsp} [@{term rVr1}]
+  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding rAp1_rsp} [@{term rAp1}]
+  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding rLm1_rsp} [@{term rLm1}]
+  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding rLt1_rsp} [@{term rLt1}]
+  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding permute_rtrm1_rsp} [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"}]
+  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha1_eqvt}) 1) *}
+
+lemmas trm1_bp_induct = rtrm1_bp.induct[quot_lifted]
+lemmas trm1_bp_inducts = rtrm1_bp.inducts[quot_lifted]
+
+instantiation trm1 and bp :: pt
+begin
+
+quotient_definition
+  "permute_trm1 :: perm \<Rightarrow> trm1 \<Rightarrow> trm1"
+is
+  "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"
+
+instance by default 
+  (simp_all add: permute_rtrm1_permute_bp_zero[quot_lifted] permute_rtrm1_permute_bp_append[quot_lifted])
+
+end
+
+lemmas
+    permute_trm1 = permute_rtrm1_permute_bp.simps[quot_lifted]
+and fv_trm1 = fv_rtrm1_fv_bp.simps[quot_lifted]
+and fv_trm1_eqvt = fv_rtrm1_eqvt[quot_lifted]
+and alpha1_INJ = alpha1_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
+
+lemma supports:
+  "(supp (atom x)) supports (Vr1 x)"
+  "(supp t \<union> supp s) supports (Ap1 t s)"
+  "(supp (atom x) \<union> supp t) supports (Lm1 x t)"
+  "(supp b \<union> supp t \<union> supp s) supports (Lt1 b t s)"
+  "{} supports BUnit"
+  "(supp (atom x)) supports (BVr x)"
+  "(supp a \<union> supp b) supports (BPr a b)"
+apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh permute_trm1)
+apply(rule_tac [!] allI)+
+apply(rule_tac [!] impI)
+apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
+apply(simp_all add: fresh_atom)
+done
+
+lemma rtrm1_bp_fs:
+  "finite (supp (x :: trm1))"
+  "finite (supp (b :: bp))"
+  apply (induct x and b rule: trm1_bp_inducts)
+  apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
+  apply(simp_all add: supp_atom)
+  done
+
+instance trm1 :: fs
+apply default
+apply (rule rtrm1_bp_fs(1))
+done
+
+lemma fv_eq_bv: "fv_bp bp = bv1 bp"
+apply(induct bp rule: trm1_bp_inducts(2))
+apply(simp_all)
+done
+
+lemma helper2: "{b. \<forall>pi. pi \<bullet> (a \<rightleftharpoons> b) \<bullet> bp \<noteq> bp} = {}"
+apply auto
+apply (rule_tac x="(x \<rightleftharpoons> a)" in exI)
+apply auto
+done
+
+lemma supp_fv:
+  "supp t = fv_trm1 t"
+  "supp b = fv_bp b"
+apply(induct t and b rule: trm1_bp_inducts)
+apply(simp_all)
+apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1)
+apply(simp only: supp_at_base[simplified supp_def])
+apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1)
+apply(simp add: Collect_imp_eq Collect_neg_eq)
+apply(subgoal_tac "supp (Lm1 name rtrm1) = supp (Abs {atom name} rtrm1)")
+apply(simp add: supp_Abs fv_trm1)
+apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt permute_trm1)
+apply(simp add: alpha1_INJ)
+apply(simp add: Abs_eq_iff)
+apply(simp add: alpha_gen.simps)
+apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric])
+apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp(rtrm11) \<union> supp (Abs (bv1 bp) rtrm12)")
+apply(simp add: supp_Abs fv_trm1 fv_eq_bv)
+apply(simp (no_asm) add: supp_def permute_trm1)
+apply(simp add: alpha1_INJ alpha_bp_eq)
+apply(simp add: Abs_eq_iff)
+apply(simp add: alpha_gen)
+apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv)
+apply(simp add: Collect_imp_eq Collect_neg_eq fresh_star_def helper2)
+apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
+apply(simp (no_asm) add: supp_def eqvts)
+apply(fold supp_def)
+apply(simp add: supp_at_base)
+apply(simp (no_asm) add: supp_def Collect_imp_eq Collect_neg_eq)
+apply(simp add: Collect_imp_eq[symmetric] Collect_neg_eq[symmetric] supp_def[symmetric])
+done
+
+lemma trm1_supp:
+  "supp (Vr1 x) = {atom x}"
+  "supp (Ap1 t1 t2) = supp t1 \<union> supp t2"
+  "supp (Lm1 x t) = (supp t) - {atom x}"
+  "supp (Lt1 b t s) = supp t \<union> (supp s - bv1 b)"
+by (simp_all add: supp_fv fv_trm1 fv_eq_bv)
+
+lemma trm1_induct_strong:
+  assumes "\<And>name b. P b (Vr1 name)"
+  and     "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)"
+  and     "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)"
+  and     "\<And>bp rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12; bv1 bp \<sharp>* b\<rbrakk> \<Longrightarrow> P b (Lt1 bp rtrm11 rtrm12)"
+  shows   "P a rtrma"
+sorry
+
+section {*** lets with single assignments ***}
+
+datatype rtrm2 =
+  rVr2 "name"
+| rAp2 "rtrm2" "rtrm2"
+| rLm2 "name" "rtrm2" --"bind (name) in (rtrm2)"
+| rLt2 "rassign" "rtrm2" --"bind (bv2 rassign) in (rtrm2)"
+and rassign =
+  rAs "name" "rtrm2"
+
+(* to be given by the user *)
+primrec 
+  rbv2
+where
+  "rbv2 (rAs x t) = {atom x}"
+
+setup {* snd o define_raw_perms ["rtrm2", "rassign"] ["Terms.rtrm2", "Terms.rassign"] *}
+
+local_setup {* snd o define_fv_alpha "Terms.rtrm2"
+  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv2}, 0)]]],
+   [[[], []]]] *}
+
+notation
+  alpha_rtrm2 ("_ \<approx>2 _" [100, 100] 100) and
+  alpha_rassign ("_ \<approx>2b _" [100, 100] 100)
+thm alpha_rtrm2_alpha_rassign.intros
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha2_inj}, []), (build_alpha_inj @{thms alpha_rtrm2_alpha_rassign.intros} @{thms rtrm2.distinct rtrm2.inject rassign.distinct rassign.inject} @{thms alpha_rtrm2.cases alpha_rassign.cases} ctxt)) ctxt)) *}
+thm alpha2_inj
+
+lemma alpha2_eqvt:
+  "t \<approx>2 s \<Longrightarrow> (pi \<bullet> t) \<approx>2 (pi \<bullet> s)"
+  "a \<approx>2b b \<Longrightarrow> (pi \<bullet> a) \<approx>2b (pi \<bullet> b)"
+sorry
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha2_equivp}, []),
+  (build_equivps [@{term alpha_rtrm2}, @{term alpha_rassign}] @{thm rtrm2_rassign.induct} @{thm alpha_rtrm2_alpha_rassign.induct} @{thms rtrm2.inject rassign.inject} @{thms alpha2_inj} @{thms rtrm2.distinct rassign.distinct} @{thms alpha_rtrm2.cases alpha_rassign.cases} @{thms alpha2_eqvt} ctxt)) ctxt)) *}
+thm alpha2_equivp
+
+local_setup  {* define_quotient_type 
+  [(([], @{binding trm2}, NoSyn), (@{typ rtrm2}, @{term alpha_rtrm2})),
+   (([], @{binding assign}, NoSyn), (@{typ rassign}, @{term alpha_rassign}))]
+  ((rtac @{thm alpha2_equivp(1)} 1) THEN (rtac @{thm alpha2_equivp(2)}) 1) *}
+
+local_setup {*
+(fn ctxt => ctxt
+ |> snd o (Quotient_Def.quotient_lift_const ("Vr2", @{term rVr2}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Ap2", @{term rAp2}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lm2", @{term rLm2}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lt2", @{term rLt2}))
+ |> snd o (Quotient_Def.quotient_lift_const ("As", @{term rAs}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_trm2", @{term fv_rtrm2}))
+ |> snd o (Quotient_Def.quotient_lift_const ("bv2", @{term rbv2})))
+*}
+print_theorems
+
+local_setup {* prove_const_rsp @{binding fv_rtrm2_rsp} [@{term fv_rtrm2}, @{term fv_rassign}]
+  (fn _ => fvbv_rsp_tac @{thm alpha_rtrm2_alpha_rassign.induct} @{thms fv_rtrm2_fv_rassign.simps} 1) *}
+local_setup {* prove_const_rsp @{binding rbv2_rsp} [@{term rbv2}]
+  (fn _ => fvbv_rsp_tac @{thm alpha_rtrm2_alpha_rassign.inducts(2)} @{thms rbv2.simps} 1) *}
+local_setup {* prove_const_rsp @{binding rVr2_rsp} [@{term rVr2}]
+  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp} @{thms alpha2_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding rAp2_rsp} [@{term rAp2}]
+  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp} @{thms alpha2_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding rLm2_rsp} [@{term rLm2}]
+  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp} @{thms alpha2_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding rLt2_rsp} [@{term rLt2}]
+  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp rbv2_rsp} @{thms alpha2_equivp} 1) *}
+local_setup {* prove_const_rsp @{binding permute_rtrm2_rsp} [@{term "permute :: perm \<Rightarrow> rtrm2 \<Rightarrow> rtrm2"}]
+  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha2_eqvt}) 1) *}
+
+
+section {*** lets with many assignments ***}
+
+datatype rtrm3 =
+  rVr3 "name"
+| rAp3 "rtrm3" "rtrm3"
+| rLm3 "name" "rtrm3" --"bind (name) in (trm3)"
+| rLt3 "rassigns" "rtrm3" --"bind (bv3 assigns) in (trm3)"
+and rassigns =
+  rANil
+| rACons "name" "rtrm3" "rassigns"
+
+(* to be given by the user *)
+primrec 
+  bv3
+where
+  "bv3 rANil = {}"
+| "bv3 (rACons x t as) = {atom x} \<union> (bv3 as)"
+
+setup {* snd o define_raw_perms ["rtrm3", "rassigns"] ["Terms.rtrm3", "Terms.rassigns"] *}
+
+local_setup {* snd o define_fv_alpha "Terms.rtrm3"
+  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term bv3}, 0)]]],
+   [[], [[], [], []]]] *}
+print_theorems
+
+notation
+  alpha_rtrm3 ("_ \<approx>3 _" [100, 100] 100) and
+  alpha_rassigns ("_ \<approx>3a _" [100, 100] 100)
+thm alpha_rtrm3_alpha_rassigns.intros
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha3_inj}, []), (build_alpha_inj @{thms alpha_rtrm3_alpha_rassigns.intros} @{thms rtrm3.distinct rtrm3.inject rassigns.distinct rassigns.inject} @{thms alpha_rtrm3.cases alpha_rassigns.cases} ctxt)) ctxt)) *}
+thm alpha3_inj
+
+lemma alpha3_eqvt:
+  "t \<approx>3 s \<Longrightarrow> (pi \<bullet> t) \<approx>3 (pi \<bullet> s)"
+  "a \<approx>3a b \<Longrightarrow> (pi \<bullet> a) \<approx>3a (pi \<bullet> b)"
+sorry
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha3_equivp}, []),
+  (build_equivps [@{term alpha_rtrm3}, @{term alpha_rassigns}] @{thm rtrm3_rassigns.induct} @{thm alpha_rtrm3_alpha_rassigns.induct} @{thms rtrm3.inject rassigns.inject} @{thms alpha3_inj} @{thms rtrm3.distinct rassigns.distinct} @{thms alpha_rtrm3.cases alpha_rassigns.cases} @{thms alpha3_eqvt} ctxt)) ctxt)) *}
+thm alpha3_equivp
+
+quotient_type
+  trm3 = rtrm3 / alpha_rtrm3
+and
+  assigns = rassigns / alpha_rassigns
+  by (rule alpha3_equivp(1)) (rule alpha3_equivp(2))
+
+
+section {*** lam with indirect list recursion ***}
+
+datatype rtrm4 =
+  rVr4 "name"
+| rAp4 "rtrm4" "rtrm4 list"
+| rLm4 "name" "rtrm4"  --"bind (name) in (trm)"
+print_theorems
+
+thm rtrm4.recs
+
+(* there cannot be a clause for lists, as *)
+(* permutations are  already defined in Nominal (also functions, options, and so on) *)
+setup {* snd o define_raw_perms ["rtrm4"] ["Terms.rtrm4"] *}
+
+(* "repairing" of the permute function *)
+lemma repaired:
+  fixes ts::"rtrm4 list"
+  shows "permute_rtrm4_list p ts = p \<bullet> ts"
+  apply(induct ts)
+  apply(simp_all)
+  done
+
+thm permute_rtrm4_permute_rtrm4_list.simps
+thm permute_rtrm4_permute_rtrm4_list.simps[simplified repaired]
+
+local_setup {* snd o define_fv_alpha "Terms.rtrm4" [
+  [[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]]], [[], [[], []]]  ] *}
+print_theorems
+
+notation
+  alpha_rtrm4 ("_ \<approx>4 _" [100, 100] 100) and
+  alpha_rtrm4_list ("_ \<approx>4l _" [100, 100] 100)
+thm alpha_rtrm4_alpha_rtrm4_list.intros
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha4_inj}, []), (build_alpha_inj @{thms alpha_rtrm4_alpha_rtrm4_list.intros} @{thms rtrm4.distinct rtrm4.inject list.distinct list.inject} @{thms alpha_rtrm4.cases alpha_rtrm4_list.cases} ctxt)) ctxt)) *}
+thm alpha4_inj
+
+lemma alpha4_eqvt:
+  "t \<approx>4 s \<Longrightarrow> (pi \<bullet> t) \<approx>4 (pi \<bullet> s)"
+  "a \<approx>4l b \<Longrightarrow> (pi \<bullet> a) \<approx>4l (pi \<bullet> b)"
+sorry
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha4_equivp}, []),
+  (build_equivps [@{term alpha_rtrm4}, @{term alpha_rtrm4_list}] @{thm rtrm4.induct} @{thm alpha_rtrm4_alpha_rtrm4_list.induct} @{thms rtrm4.inject list.inject} @{thms alpha4_inj} @{thms rtrm4.distinct list.distinct} @{thms alpha_rtrm4_list.cases alpha_rtrm4.cases} @{thms alpha4_eqvt} ctxt)) ctxt)) *}
+thm alpha4_equivp
+
+quotient_type 
+  qrtrm4 = rtrm4 / alpha_rtrm4 and
+  qrtrm4list = "rtrm4 list" / alpha_rtrm4_list
+  by (simp_all add: alpha4_equivp)
+
+
+datatype rtrm5 =
+  rVr5 "name"
+| rAp5 "rtrm5" "rtrm5"
+| rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)"
+and rlts =
+  rLnil
+| rLcons "name" "rtrm5" "rlts"
+
+primrec
+  rbv5
+where
+  "rbv5 rLnil = {}"
+| "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)"
+
+
+setup {* snd o define_raw_perms ["rtrm5", "rlts"] ["Terms.rtrm5", "Terms.rlts"] *}
+print_theorems
+
+local_setup {* snd o define_fv_alpha "Terms.rtrm5" [
+  [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[], [], []]]  ] *}
+print_theorems
+
+(* Alternate version with additional binding of name in rlts in rLcons *)
+(*local_setup {* snd o define_fv_alpha "Terms.rtrm5" [
+  [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[(NONE,0)], [], [(NONE,0)]]]  ] *}
+print_theorems*)
+
+notation
+  alpha_rtrm5 ("_ \<approx>5 _" [100, 100] 100) and
+  alpha_rlts ("_ \<approx>l _" [100, 100] 100)
+thm alpha_rtrm5_alpha_rlts.intros
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_alpha_inj @{thms alpha_rtrm5_alpha_rlts.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases} ctxt)) ctxt)) *}
+thm alpha5_inj
+
+lemma rbv5_eqvt:
+  "pi \<bullet> (rbv5 x) = rbv5 (pi \<bullet> x)"
+sorry
+
+lemma fv_rtrm5_eqvt:
+  "pi \<bullet> (fv_rtrm5 x) = fv_rtrm5 (pi \<bullet> x)"
+sorry
+
+lemma fv_rlts_eqvt:
+  "pi \<bullet> (fv_rlts x) = fv_rlts (pi \<bullet> x)"
+sorry
+
+lemma alpha5_eqvt:
+  "xa \<approx>5 y \<Longrightarrow> (x \<bullet> xa) \<approx>5 (x \<bullet> y)"
+  "xb \<approx>l ya \<Longrightarrow> (x \<bullet> xb) \<approx>l (x \<bullet> ya)"
+  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts)
+  apply (simp_all add: alpha5_inj)
+  apply (erule exE)+
+  apply(unfold alpha_gen)
+  apply (erule conjE)+
+  apply (rule conjI)
+  apply (rule_tac x="x \<bullet> pi" in exI)
+  apply (rule conjI)
+  apply(rule_tac ?p1="- x" in permute_eq_iff[THEN iffD1])
+  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rlts_eqvt)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- x" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rlts_eqvt)
+  apply (subst permute_eqvt[symmetric])
+  apply (simp)
+  apply (rule_tac x="x \<bullet> pia" in exI)
+  apply (rule conjI)
+  apply(rule_tac ?p1="- x" in permute_eq_iff[THEN iffD1])
+  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rtrm5_eqvt)
+  apply(rule conjI)
+  apply(rule_tac ?p1="- x" in fresh_star_permute_iff[THEN iffD1])
+  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rtrm5_eqvt)
+  apply (subst permute_eqvt[symmetric])
+  apply (simp)
+  done
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_equivp}, []),
+  (build_equivps [@{term alpha_rtrm5}, @{term alpha_rlts}] @{thm rtrm5_rlts.induct} @{thm alpha_rtrm5_alpha_rlts.induct} @{thms rtrm5.inject rlts.inject} @{thms alpha5_inj} @{thms rtrm5.distinct rlts.distinct} @{thms alpha_rtrm5.cases alpha_rlts.cases} @{thms alpha5_eqvt} ctxt)) ctxt)) *}
+thm alpha5_equivp
+
+quotient_type
+  trm5 = rtrm5 / alpha_rtrm5
+and
+  lts = rlts / alpha_rlts
+  by (auto intro: alpha5_equivp)
+
+local_setup {*
+(fn ctxt => ctxt
+ |> snd o (Quotient_Def.quotient_lift_const ("Vr5", @{term rVr5}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Ap5", @{term rAp5}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lt5", @{term rLt5}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lnil", @{term rLnil}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lcons", @{term rLcons}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_trm5", @{term fv_rtrm5}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_lts", @{term fv_rlts}))
+ |> snd o (Quotient_Def.quotient_lift_const ("bv5", @{term rbv5})))
+*}
+print_theorems
+
+lemma alpha5_rfv:
+  "(t \<approx>5 s \<Longrightarrow> fv_rtrm5 t = fv_rtrm5 s)"
+  "(l \<approx>l m \<Longrightarrow> fv_rlts l = fv_rlts m)"
+  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts)
+  apply(simp_all add: alpha_gen)
+  done
+
+lemma bv_list_rsp:
+  shows "x \<approx>l y \<Longrightarrow> rbv5 x = rbv5 y"
+  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts(2))
+  apply(simp_all)
+  done
+
+lemma [quot_respect]:
+  "(alpha_rlts ===> op =) fv_rlts fv_rlts"
+  "(alpha_rtrm5 ===> op =) fv_rtrm5 fv_rtrm5"
+  "(alpha_rlts ===> op =) rbv5 rbv5"
+  "(op = ===> alpha_rtrm5) rVr5 rVr5"
+  "(alpha_rtrm5 ===> alpha_rtrm5 ===> alpha_rtrm5) rAp5 rAp5"
+  "(alpha_rlts ===> alpha_rtrm5 ===> alpha_rtrm5) rLt5 rLt5"
+  "(op = ===> alpha_rtrm5 ===> alpha_rlts ===> alpha_rlts) rLcons rLcons"
+  "(op = ===> alpha_rtrm5 ===> alpha_rtrm5) permute permute"
+  "(op = ===> alpha_rlts ===> alpha_rlts) permute permute"
+  apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp)
+  apply (clarify) apply (rule conjI)
+  apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
+  apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
+  done
+
+lemma
+  shows "(alpha_rlts ===> op =) rbv5 rbv5"
+  by (simp add: bv_list_rsp)
+
+lemmas trm5_lts_inducts = rtrm5_rlts.inducts[quot_lifted]
+
+instantiation trm5 and lts :: pt
+begin
+
+quotient_definition
+  "permute_trm5 :: perm \<Rightarrow> trm5 \<Rightarrow> trm5"
+is
+  "permute :: perm \<Rightarrow> rtrm5 \<Rightarrow> rtrm5"
+
+quotient_definition
+  "permute_lts :: perm \<Rightarrow> lts \<Rightarrow> lts"
+is
+  "permute :: perm \<Rightarrow> rlts \<Rightarrow> rlts"
+
+instance by default
+  (simp_all add: permute_rtrm5_permute_rlts_zero[quot_lifted] permute_rtrm5_permute_rlts_append[quot_lifted])
+
+end
+
+lemmas
+    permute_trm5_lts = permute_rtrm5_permute_rlts.simps[quot_lifted]
+and alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
+and bv5[simp] = rbv5.simps[quot_lifted]
+and fv_trm5_lts[simp] = fv_rtrm5_fv_rlts.simps[quot_lifted]
+
+lemma lets_ok:
+  "(Lt5 (Lcons x (Vr5 x) Lnil) (Vr5 x)) = (Lt5 (Lcons y (Vr5 y) Lnil) (Vr5 y))"
+apply (subst alpha5_INJ)
+apply (rule conjI)
+apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
+apply (simp only: alpha_gen)
+apply (simp add: permute_trm5_lts fresh_star_def)
+apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
+apply (simp only: alpha_gen)
+apply (simp add: permute_trm5_lts fresh_star_def)
+done
+
+lemma lets_ok2:
+  "(Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) =
+   (Lt5 (Lcons y (Vr5 y) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
+apply (subst alpha5_INJ)
+apply (rule conjI)
+apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
+apply (simp only: alpha_gen)
+apply (simp add: permute_trm5_lts fresh_star_def)
+apply (rule_tac x="0 :: perm" in exI)
+apply (simp only: alpha_gen)
+apply (simp add: permute_trm5_lts fresh_star_def)
+done
+
+
+lemma lets_not_ok1:
+  "x \<noteq> y \<Longrightarrow> (Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
+             (Lt5 (Lcons y (Vr5 x) (Lcons x (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
+apply (simp add: alpha5_INJ(3) alpha_gen)
+apply (simp add: permute_trm5_lts fresh_star_def alpha5_INJ(5) alpha5_INJ(2) alpha5_INJ(1))
+done
+
+lemma distinct_helper:
+  shows "\<not>(rVr5 x \<approx>5 rAp5 y z)"
+  apply auto
+  apply (erule alpha_rtrm5.cases)
+  apply (simp_all only: rtrm5.distinct)
+  done
+
+lemma distinct_helper2:
+  shows "(Vr5 x) \<noteq> (Ap5 y z)"
+  by (lifting distinct_helper)
+
+lemma lets_nok:
+  "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow>
+   (Lt5 (Lcons x (Ap5 (Vr5 z) (Vr5 z)) (Lcons y (Vr5 z) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
+   (Lt5 (Lcons y (Vr5 z) (Lcons x (Ap5 (Vr5 z) (Vr5 z)) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
+apply (simp only: alpha5_INJ(3) alpha5_INJ(5) alpha_gen permute_trm5_lts fresh_star_def)
+apply (simp add: distinct_helper2)
+done
+
+
+(* example with a bn function defined over the type itself *)
+datatype rtrm6 =
+  rVr6 "name"
+| rLm6 "name" "rtrm6" --"bind name in rtrm6"
+| rLt6 "rtrm6" "rtrm6" --"bind (bv6 left) in (right)"
+
+primrec
+  rbv6
+where
+  "rbv6 (rVr6 n) = {}"
+| "rbv6 (rLm6 n t) = {atom n} \<union> rbv6 t"
+| "rbv6 (rLt6 l r) = rbv6 l \<union> rbv6 r"
+
+setup {* snd o define_raw_perms ["rtrm6"] ["Terms.rtrm6"] *}
+print_theorems
+
+local_setup {* snd o define_fv_alpha "Terms.rtrm6" [
+  [[[]], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv6}, 0)]]]] *}
+notation alpha_rtrm6 ("_ \<approx>6 _" [100, 100] 100)
+(* HERE THE RULES DIFFER *)
+thm alpha_rtrm6.intros
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha6_inj}, []), (build_alpha_inj @{thms alpha_rtrm6.intros} @{thms rtrm6.distinct rtrm6.inject} @{thms alpha_rtrm6.cases} ctxt)) ctxt)) *}
+thm alpha6_inj
+
+lemma alpha6_eqvt:
+  "xa \<approx>6 y \<Longrightarrow> (x \<bullet> xa) \<approx>6 (x \<bullet> y)"
+sorry
+
+local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha6_equivp}, []),
+  (build_equivps [@{term alpha_rtrm6}] @{thm rtrm6.induct} @{thm alpha_rtrm6.induct} @{thms rtrm6.inject} @{thms alpha6_inj} @{thms rtrm6.distinct} @{thms alpha_rtrm6.cases} @{thms alpha6_eqvt} ctxt)) ctxt)) *}
+thm alpha6_equivp
+
+quotient_type
+  trm6 = rtrm6 / alpha_rtrm6
+  by (auto intro: alpha6_equivp)
+
+local_setup {*
+(fn ctxt => ctxt
+ |> snd o (Quotient_Def.quotient_lift_const ("Vr6", @{term rVr6}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lm6", @{term rLm6}))
+ |> snd o (Quotient_Def.quotient_lift_const ("Lt6", @{term rLt6}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_trm6", @{term fv_rtrm6}))
+ |> snd o (Quotient_Def.quotient_lift_const ("bv6", @{term rbv6})))
+*}
+print_theorems
+
+lemma [quot_respect]:
+  "(op = ===> alpha_rtrm6 ===> alpha_rtrm6) permute permute"
+by (auto simp add: alpha6_eqvt)
+
+(* Definitely not true , see lemma below *)
+lemma [quot_respect]:"(alpha_rtrm6 ===> op =) rbv6 rbv6"
+apply simp apply clarify
+apply (erule alpha_rtrm6.induct)
+oops
+
+lemma "(a :: name) \<noteq> b \<Longrightarrow> \<not> (alpha_rtrm6 ===> op =) rbv6 rbv6"
+apply simp
+apply (rule_tac x="rLm6 (a::name) (rVr6 (a :: name))" in  exI)
+apply (rule_tac x="rLm6 (b::name) (rVr6 (b :: name))" in  exI)
+apply simp
+apply (simp add: alpha6_inj)
+apply (rule_tac x="(a \<leftrightarrow> b)" in  exI)
+apply (simp add: alpha_gen fresh_star_def)
+apply (simp add: alpha6_inj)
+done
+
+lemma fv6_rsp: "x \<approx>6 y \<Longrightarrow> fv_rtrm6 x = fv_rtrm6 y"
+apply (induct_tac x y rule: alpha_rtrm6.induct)
+apply simp_all
+apply (erule exE)
+apply (simp_all add: alpha_gen)
+done
+
+lemma [quot_respect]:"(alpha_rtrm6 ===> op =) fv_rtrm6 fv_rtrm6"
+by (simp add: fv6_rsp)
+
+lemma [quot_respect]:
+ "(op = ===> alpha_rtrm6) rVr6 rVr6"
+ "(op = ===> alpha_rtrm6 ===> alpha_rtrm6) rLm6 rLm6"
+apply auto
+apply (simp_all add: alpha6_inj)
+apply (rule_tac x="0::perm" in exI)
+apply (simp add: alpha_gen fv6_rsp fresh_star_def fresh_zero_perm)
+done
+
+lemma [quot_respect]:
+ "(alpha_rtrm6 ===> alpha_rtrm6 ===> alpha_rtrm6) rLt6 rLt6"
+apply auto
+apply (simp_all add: alpha6_inj)
+apply (rule_tac [!] x="0::perm" in exI)
+apply (simp_all add: alpha_gen fresh_star_def fresh_zero_perm)
+(* needs rbv6_rsp *)
+oops
+
+instantiation trm6 :: pt begin
+
+quotient_definition
+  "permute_trm6 :: perm \<Rightarrow> trm6 \<Rightarrow> trm6"
+is
+  "permute :: perm \<Rightarrow> rtrm6 \<Rightarrow> rtrm6"
+
+instance
+apply default
+sorry
+end
+
+lemma lifted_induct:
+"\<lbrakk>x1 = x2; \<And>name namea. name = namea \<Longrightarrow> P (Vr6 name) (Vr6 namea);
+ \<And>name rtrm6 namea rtrm6a.
+    \<lbrakk>True;
+     \<exists>pi. fv_trm6 rtrm6 - {atom name} = fv_trm6 rtrm6a - {atom namea} \<and>
+          (fv_trm6 rtrm6 - {atom name}) \<sharp>* pi \<and> pi \<bullet> rtrm6 = rtrm6a \<and> P (pi \<bullet> rtrm6) rtrm6a\<rbrakk>
+    \<Longrightarrow> P (Lm6 name rtrm6) (Lm6 namea rtrm6a);
+ \<And>rtrm61 rtrm61a rtrm62 rtrm62a.
+    \<lbrakk>rtrm61 = rtrm61a; P rtrm61 rtrm61a;
+     \<exists>pi. fv_trm6 rtrm62 - bv6 rtrm61 = fv_trm6 rtrm62a - bv6 rtrm61a \<and>
+          (fv_trm6 rtrm62 - bv6 rtrm61) \<sharp>* pi \<and> pi \<bullet> rtrm62 = rtrm62a \<and> P (pi \<bullet> rtrm62) rtrm62a\<rbrakk>
+    \<Longrightarrow> P (Lt6 rtrm61 rtrm62) (Lt6 rtrm61a rtrm62a)\<rbrakk>
+\<Longrightarrow> P x1 x2"
+apply (lifting alpha_rtrm6.induct[unfolded alpha_gen])
+apply injection
+(* notice unsolvable goals: (alpha_rtrm6 ===> op =) rbv6 rbv6 *)
+oops
+
+lemma lifted_inject_a3:
+"(Lt6 rtrm61 rtrm62 = Lt6 rtrm61a rtrm62a) =
+(rtrm61 = rtrm61a \<and>
+ (\<exists>pi. fv_trm6 rtrm62 - bv6 rtrm61 = fv_trm6 rtrm62a - bv6 rtrm61a \<and>
+       (fv_trm6 rtrm62 - bv6 rtrm61) \<sharp>* pi \<and> pi \<bullet> rtrm62 = rtrm62a))"
+apply(lifting alpha6_inj(3)[unfolded alpha_gen])
+apply injection
+(* notice unsolvable goals: (alpha_rtrm6 ===> op =) rbv6 rbv6 *)
+oops
+
+
+
+
+(* example with a respectful bn function defined over the type itself *)
+
+datatype rtrm7 =
+  rVr7 "name"
+| rLm7 "name" "rtrm7" --"bind left in right"
+| rLt7 "rtrm7" "rtrm7" --"bind (bv7 left) in (right)"
+
+primrec
+  rbv7
+where
+  "rbv7 (rVr7 n) = {atom n}"
+| "rbv7 (rLm7 n t) = rbv7 t - {atom n}"
+| "rbv7 (rLt7 l r) = rbv7 l \<union> rbv7 r"
+
+setup {* snd o define_raw_perms ["rtrm7"] ["Terms.rtrm7"] *}
+thm permute_rtrm7.simps
+
+local_setup {* snd o define_fv_alpha "Terms.rtrm7" [
+  [[[]], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv7}, 0)]]]] *}
+print_theorems
+notation
+  alpha_rtrm7 ("_ \<approx>7a _" [100, 100] 100)
+(* HERE THE RULES DIFFER *)
+thm alpha_rtrm7.intros
+thm fv_rtrm7.simps
+inductive
+  alpha7 :: "rtrm7 \<Rightarrow> rtrm7 \<Rightarrow> bool" ("_ \<approx>7 _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (rVr7 a) \<approx>7 (rVr7 b)"
+| a2: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha7 fv_rtrm7 pi ({atom b}, s))) \<Longrightarrow> rLm7 a t \<approx>7 rLm7 b s"
+| a3: "(\<exists>pi. (((rbv7 t1), s1) \<approx>gen alpha7 fv_rtrm7 pi ((rbv7 t2), s2))) \<Longrightarrow> rLt7 t1 s1 \<approx>7 rLt7 t2 s2"
+
+lemma "(x::name) \<noteq> y \<Longrightarrow> \<not> (alpha7 ===> op =) rbv7 rbv7"
+  apply simp
+  apply (rule_tac x="rLt7 (rVr7 x) (rVr7 x)" in exI)
+  apply (rule_tac x="rLt7 (rVr7 y) (rVr7 y)" in exI)
+  apply simp
+  apply (rule a3)
+  apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
+  apply (simp_all add: alpha_gen fresh_star_def)
+  apply (rule a1)
+  apply (rule refl)
+done
+
+
+
+
+
+datatype rfoo8 =
+  Foo0 "name"
+| Foo1 "rbar8" "rfoo8" --"bind bv(bar) in foo"
+and rbar8 =
+  Bar0 "name"
+| Bar1 "name" "name" "rbar8" --"bind second name in b"
+
+primrec
+  rbv8
+where
+  "rbv8 (Bar0 x) = {}"
+| "rbv8 (Bar1 v x b) = {atom v}"
+
+setup {* snd o define_raw_perms ["rfoo8", "rbar8"] ["Terms.rfoo8", "Terms.rbar8"] *}
+print_theorems
+
+local_setup {* snd o define_fv_alpha "Terms.rfoo8" [
+  [[[]], [[], [(SOME @{term rbv8}, 0)]]], [[[]], [[], [(NONE, 1)], [(NONE, 1)]]]] *}
+notation
+  alpha_rfoo8 ("_ \<approx>f' _" [100, 100] 100) and
+  alpha_rbar8 ("_ \<approx>b' _" [100, 100] 100)
+(* HERE THE RULE DIFFERS *)
+thm alpha_rfoo8_alpha_rbar8.intros
+
+
+inductive
+  alpha8f :: "rfoo8 \<Rightarrow> rfoo8 \<Rightarrow> bool" ("_ \<approx>f _" [100, 100] 100)
+and
+  alpha8b :: "rbar8 \<Rightarrow> rbar8 \<Rightarrow> bool" ("_ \<approx>b _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (Foo0 a) \<approx>f (Foo0 b)"
+| a2: "a = b \<Longrightarrow> (Bar0 a) \<approx>b (Bar0 b)"
+| a3: "b1 \<approx>b b2 \<Longrightarrow> (\<exists>pi. (((rbv8 b1), t1) \<approx>gen alpha8f fv_rfoo8 pi ((rbv8 b2), t2))) \<Longrightarrow> Foo1 b1 t1 \<approx>f Foo1 b2 t2"
+| a4: "v1 = v2 \<Longrightarrow> (\<exists>pi. (({atom x1}, t1) \<approx>gen alpha8b fv_rbar8 pi ({atom x2}, t2))) \<Longrightarrow> Bar1 v1 x1 t1 \<approx>b Bar1 v2 x2 t2"
+
+lemma "(alpha8b ===> op =) rbv8 rbv8"
+  apply simp apply clarify
+  apply (erule alpha8f_alpha8b.inducts(2))
+  apply (simp_all)
+done
+
+lemma fv_rbar8_rsp_hlp: "x \<approx>b y \<Longrightarrow> fv_rbar8 x = fv_rbar8 y"
+  apply (erule alpha8f_alpha8b.inducts(2))
+  apply (simp_all add: alpha_gen)
+done
+lemma "(alpha8b ===> op =) fv_rbar8 fv_rbar8"
+  apply simp apply clarify apply (simp add: fv_rbar8_rsp_hlp)
+done
+
+lemma "(alpha8f ===> op =) fv_rfoo8 fv_rfoo8"
+  apply simp apply clarify
+  apply (erule alpha8f_alpha8b.inducts(1))
+  apply (simp_all add: alpha_gen fv_rbar8_rsp_hlp)
+done
+
+
+
+
+
+
+datatype rlam9 =
+  Var9 "name"
+| Lam9 "name" "rlam9" --"bind name in rlam"
+and rbla9 =
+  Bla9 "rlam9" "rlam9" --"bind bv(first) in second"
+
+primrec
+  rbv9
+where
+  "rbv9 (Var9 x) = {}"
+| "rbv9 (Lam9 x b) = {atom x}"
+
+setup {* snd o define_raw_perms ["rlam9", "rbla9"] ["Terms.rlam9", "Terms.rbla9"] *}
+print_theorems
+
+local_setup {* snd o define_fv_alpha "Terms.rlam9" [
+  [[[]], [[(NONE, 0)], [(NONE, 0)]]], [[[], [(SOME @{term rbv9}, 0)]]]] *}
+notation
+  alpha_rlam9 ("_ \<approx>9l' _" [100, 100] 100) and
+  alpha_rbla9 ("_ \<approx>9b' _" [100, 100] 100)
+(* HERE THE RULES DIFFER *)
+thm alpha_rlam9_alpha_rbla9.intros
+
+
+inductive
+  alpha9l :: "rlam9 \<Rightarrow> rlam9 \<Rightarrow> bool" ("_ \<approx>9l _" [100, 100] 100)
+and
+  alpha9b :: "rbla9 \<Rightarrow> rbla9 \<Rightarrow> bool" ("_ \<approx>9b _" [100, 100] 100)
+where
+  a1: "a = b \<Longrightarrow> (Var9 a) \<approx>9l (Var9 b)"
+| a4: "(\<exists>pi. (({atom x1}, t1) \<approx>gen alpha9l fv_rlam9 pi ({atom x2}, t2))) \<Longrightarrow> Lam9 x1 t1 \<approx>9l Lam9 x2 t2"
+| a3: "b1 \<approx>9l b2 \<Longrightarrow> (\<exists>pi. (((rbv9 b1), t1) \<approx>gen alpha9l fv_rlam9 pi ((rbv9 b2), t2))) \<Longrightarrow> Bla9 b1 t1 \<approx>9b Bla9 b2 t2"
+
+quotient_type
+  lam9 = rlam9 / alpha9l and bla9 = rbla9 / alpha9b
+sorry
+
+local_setup {*
+(fn ctxt => ctxt
+ |> snd o (Quotient_Def.quotient_lift_const ("qVar9", @{term Var9}))
+ |> snd o (Quotient_Def.quotient_lift_const ("qLam9", @{term Lam9}))
+ |> snd o (Quotient_Def.quotient_lift_const ("qBla9", @{term Bla9}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_lam9", @{term fv_rlam9}))
+ |> snd o (Quotient_Def.quotient_lift_const ("fv_bla9", @{term fv_rbla9}))
+ |> snd o (Quotient_Def.quotient_lift_const ("bv9", @{term rbv9})))
+*}
+print_theorems
+
+instantiation lam9 and bla9 :: pt
+begin
+
+quotient_definition
+  "permute_lam9 :: perm \<Rightarrow> lam9 \<Rightarrow> lam9"
+is
+  "permute :: perm \<Rightarrow> rlam9 \<Rightarrow> rlam9"
+
+quotient_definition
+  "permute_bla9 :: perm \<Rightarrow> bla9 \<Rightarrow> bla9"
+is
+  "permute :: perm \<Rightarrow> rbla9 \<Rightarrow> rbla9"
+
+instance
+sorry
+
+end
+
+lemma "\<lbrakk>b1 = b2; \<exists>pi. fv_lam9 t1 - bv9 b1 = fv_lam9 t2 - bv9 b2 \<and> (fv_lam9 t1 - bv9 b1) \<sharp>* pi \<and> pi \<bullet> t1 = t2\<rbrakk>
+ \<Longrightarrow> qBla9 b1 t1 = qBla9 b2 t2"
+apply (lifting a3[unfolded alpha_gen])
+apply injection
+sorry
+
+
+
+
+
+
+
+
+text {* type schemes *} 
+datatype ty = 
+  Var "name" 
+| Fun "ty" "ty"
+
+setup {* snd o define_raw_perms ["ty"] ["Terms.ty"] *}
+print_theorems
+
+datatype tyS = 
+  All "name set" "ty" 
+
+setup {* snd o define_raw_perms ["tyS"] ["Terms.tyS"] *}
+print_theorems
+
+local_setup {* snd o define_fv_alpha "Terms.ty" [[[[]], [[], []]]] *}
+print_theorems 
+
+(*
+Doesnot work yet since we do not refer to fv_ty
+local_setup {* define_raw_fv "Terms.tyS" [[[[], []]]] *}
+print_theorems
+*)
+
+primrec
+  fv_tyS
+where 
+  "fv_tyS (All xs T) = (fv_ty T - atom ` xs)"
+
+inductive
+  alpha_tyS :: "tyS \<Rightarrow> tyS \<Rightarrow> bool" ("_ \<approx>tyS _" [100, 100] 100)
+where
+  a1: "\<exists>pi. ((atom ` xs1, T1) \<approx>gen (op =) fv_ty pi (atom ` xs2, T2)) 
+        \<Longrightarrow> All xs1 T1 \<approx>tyS All xs2 T2"
+
+lemma
+  shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {b, a} (Fun (Var a) (Var b))"
+  apply(rule a1)
+  apply(simp add: alpha_gen)
+  apply(rule_tac x="0::perm" in exI)
+  apply(simp add: fresh_star_def)
+  done
+
+lemma
+  shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var b) (Var a))"
+  apply(rule a1)
+  apply(simp add: alpha_gen)
+  apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI)
+  apply(simp add: fresh_star_def)
+  done
+
+lemma
+  shows "All {a, b, c} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var a) (Var b))"
+  apply(rule a1)
+  apply(simp add: alpha_gen)
+  apply(rule_tac x="0::perm" in exI)
+  apply(simp add: fresh_star_def)
+  done
+
+lemma
+  assumes a: "a \<noteq> b"
+  shows "\<not>(All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {c} (Fun (Var c) (Var c)))"
+  using a
+  apply(clarify)
+  apply(erule alpha_tyS.cases)
+  apply(simp add: alpha_gen)
+  apply(erule conjE)+
+  apply(erule exE)
+  apply(erule conjE)+
+  apply(clarify)
+  apply(simp)
+  apply(simp add: fresh_star_def)
+  apply(auto)
+  done
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/Test.thy	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,431 @@
+theory Test
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp"
+begin
+
+atom_decl name
+
+
+section{* Interface for nominal_datatype *}
+
+text {*
+
+Nominal-Datatype-part:
+
+1st Arg: string list  
+         ^^^^^^^^^^^
+         strings of the types to be defined
+
+2nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list
+         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^   ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
+               type(s) to be defined             constructors list
+               (ty args, name, syn)              (name, typs, syn)
+
+Binder-Function-part:
+
+3rd Arg: (binding * typ option * mixfix) list 
+         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^    
+            binding function(s)           
+              to be defined               
+            (name, type, syn)             
+
+4th Arg:  term list 
+          ^^^^^^^^^
+          the equations of the binding functions
+          (Trueprop equations)
+*}
+
+text {*****************************************************}
+ML {* 
+(* nominal datatype parser *)
+local
+  structure P = OuterParse
+in
+
+val _ = OuterKeyword.keyword "bind"
+val anno_typ = Scan.option (P.name --| P.$$$ "::") -- P.typ
+
+(* binding specification *)
+(* should use and_list *)
+val bind_parser = 
+  P.enum "," ((P.$$$ "bind" |-- P.term) -- (P.$$$ "in" |-- P.name))
+
+val constr_parser =
+  P.binding -- Scan.repeat anno_typ
+
+(* datatype parser *)
+val dt_parser =
+  ((P.type_args -- P.binding -- P.opt_infix) >> P.triple1) -- 
+    (P.$$$ "=" |-- P.enum1 "|" ((constr_parser -- bind_parser -- P.opt_mixfix) >> P.triple_swap))
+
+(* function equation parser *)
+val fun_parser = 
+  Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[])
+
+(* main parser *)
+val main_parser =
+  (P.and_list1 dt_parser) -- fun_parser
+
+end
+*}
+
+(* adds "_raw" to the end of constants and types *)
+ML {*
+fun add_raw s = s ^ "_raw"
+fun add_raws ss = map add_raw ss
+fun raw_bind bn = Binding.suffix_name "_raw" bn
+
+fun replace_str ss s = 
+  case (AList.lookup (op=) ss s) of 
+     SOME s' => s'
+   | NONE => s
+
+fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts)
+  | replace_typ ty_ss T = T  
+
+fun raw_dts ty_ss dts =
+let
+  val ty_ss' = ty_ss ~~ (add_raws ty_ss)
+
+  fun raw_dts_aux1 (bind, tys, mx) =
+    (raw_bind bind, map (replace_typ ty_ss') tys, mx)
+
+  fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
+    (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
+in
+  map raw_dts_aux2 dts
+end
+
+fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)
+  | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)
+  | replace_aterm trm_ss trm = trm
+
+fun replace_term trm_ss ty_ss trm =
+  trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) 
+*}
+
+ML {*
+fun get_constrs dts =
+  flat (map (fn (_, _, _, constrs) => constrs) dts)
+
+fun get_typed_constrs dts =
+  flat (map (fn (_, bn, _, constrs) => 
+   (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)
+
+fun get_constr_strs dts =
+  map (fn (bn, _, _) => Binding.name_of bn) (get_constrs dts)
+
+fun get_bn_fun_strs bn_funs =
+  map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
+*}
+
+ML {*
+fun raw_dts_decl dt_names dts lthy =
+let
+  val thy = ProofContext.theory_of lthy
+  val conf = Datatype.default_config
+
+  val dt_names' = add_raws dt_names
+  val dt_full_names = map (Sign.full_bname thy) dt_names 
+  val dts' = raw_dts dt_full_names dts
+in
+  lthy
+  |> Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts')
+end 
+*}
+
+ML {*
+fun raw_bn_fun_decl dt_names dts bn_funs bn_eqs lthy =
+let
+  val thy = ProofContext.theory_of lthy
+
+  val dt_names' = add_raws dt_names
+  val dt_full_names = map (Sign.full_bname thy) dt_names 
+  val dt_full_names' = map (Sign.full_bname thy) dt_names' 
+  
+  val ctrs_names = map (Sign.full_bname thy) (get_constr_strs dts)
+  val ctrs_names' = map (fn (x, y) => (Sign.full_bname_path thy (add_raw x) (add_raw y))) 
+    (get_typed_constrs dts)
+
+  val bn_fun_strs = get_bn_fun_strs bn_funs
+  val bn_fun_strs' = add_raws bn_fun_strs
+  
+  val bn_funs' = map (fn (bn, opt_ty, mx) => 
+    (raw_bind bn, Option.map (replace_typ (dt_full_names ~~ dt_full_names')) opt_ty, mx)) bn_funs
+  
+  val bn_eqs' = map (fn trm => 
+    (Attrib.empty_binding, 
+      (replace_term ((ctrs_names ~~ ctrs_names') @ (bn_fun_strs ~~ bn_fun_strs')) (dt_full_names ~~ dt_full_names') trm))) bn_eqs
+in
+  if null bn_eqs 
+  then (([], []), lthy)
+  else Primrec.add_primrec bn_funs' bn_eqs' lthy 
+end 
+*}
+
+ML {* 
+fun nominal_datatype2 dts bn_funs bn_eqs lthy =
+let
+  val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts
+in
+  lthy
+  |> raw_dts_decl dts_names dts
+  ||>> raw_bn_fun_decl dts_names dts bn_funs bn_eqs
+end
+*}
+
+ML {*
+(* makes a full named type out of a binding with tvars applied to it *)
+fun mk_type thy bind tvrs =
+  Type (Sign.full_name thy bind, map (fn s => TVar ((s, 0), [])) tvrs)
+
+fun get_constrs2 thy dts =
+let
+  val dts' = map (fn (tvrs, tname, _, constrs) => (mk_type thy tname tvrs, constrs)) dts
+in
+  flat (map (fn (ty, constrs) => map (fn (bn, tys, mx) =>  (bn, tys ---> ty, mx)) constrs) dts')
+end
+*}
+
+ML {*
+fun nominal_datatype2_cmd (dt_strs, (bn_fun_strs, bn_eq_strs)) lthy =
+let
+  val thy = ProofContext.theory_of lthy
+
+  fun prep_typ ((tvs, tname, mx), _) = (tname, length tvs, mx);
+
+  (* adding the types for parsing datatypes *)
+  val lthy_tmp = lthy
+    |> Local_Theory.theory (Sign.add_types (map prep_typ dt_strs))
+
+  fun prep_cnstr lthy (((cname, atys), mx), binders) =
+    (cname, map (Syntax.read_typ lthy o snd) atys, mx)
+  
+  fun prep_dt lthy ((tvs, tname, mx), cnstrs) = 
+    (tvs, tname, mx, map (prep_cnstr lthy) cnstrs)
+
+  (* parsing the datatypes *)
+  val dts_prep = map (prep_dt lthy_tmp) dt_strs
+
+  (* adding constructors for parsing functions *)
+  val lthy_tmp2 = lthy_tmp
+    |> Local_Theory.theory (Sign.add_consts_i (get_constrs2 thy dts_prep))
+
+  val (bn_fun_aux, bn_eq_aux) = fst (Specification.read_spec bn_fun_strs bn_eq_strs lthy_tmp2)
+
+  fun prep_bn_fun ((bn, T), mx) = (bn, SOME T, mx) 
+
+  fun prep_bn_eq (attr, t) = t
+
+  val bn_fun_prep = map prep_bn_fun bn_fun_aux
+  val bn_eq_prep = map prep_bn_eq bn_eq_aux 
+in
+  nominal_datatype2 dts_prep bn_fun_prep bn_eq_prep lthy |> snd
+end
+*}
+
+(* Command Keyword *)
+ML {*
+let
+   val kind = OuterKeyword.thy_decl
+in
+   OuterSyntax.local_theory "nominal_datatype" "test" kind 
+     (main_parser >> nominal_datatype2_cmd)
+end
+*}
+
+text {* example 1 *}
+
+nominal_datatype lam =
+  VAR "name"
+| APP "lam" "lam"
+| LET bp::"bp" t::"lam"   bind "bi bp" in t ("Let _ in _" [100,100] 100)
+and bp = 
+  BP "name" "lam"  ("_ ::= _" [100,100] 100)
+binder
+  bi::"bp \<Rightarrow> name set"
+where
+  "bi (BP x t) = {x}"
+
+typ lam_raw
+term VAR_raw
+term Test.BP_raw
+thm bi_raw.simps
+
+print_theorems
+
+text {* examples 2 *}
+nominal_datatype trm =
+  Var "name"
+| App "trm" "trm"
+| Lam x::"name" t::"trm"        bind x in t 
+| Let p::"pat" "trm" t::"trm"   bind "f p" in t
+and pat =
+  PN
+| PS "name"
+| PD "name" "name"
+binder
+  f::"pat \<Rightarrow> name set"
+where 
+  "f PN = {}"
+| "f (PS x) = {x}"
+| "f (PD x y) = {x,y}"
+
+thm f_raw.simps
+
+nominal_datatype trm0 =
+  Var0 "name"
+| App0 "trm0" "trm0"
+| Lam0 x::"name" t::"trm0"          bind x in t 
+| Let0 p::"pat0" "trm0" t::"trm0"   bind "f0 p" in t
+and pat0 =
+  PN0
+| PS0 "name"
+| PD0 "pat0" "pat0"
+binder
+  f0::"pat0 \<Rightarrow> name set"
+where 
+  "f0 PN0 = {}"
+| "f0 (PS0 x) = {x}"
+| "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)"
+
+thm f0_raw.simps
+
+text {* example type schemes *}
+datatype ty = 
+  Var "name" 
+| Fun "ty" "ty"
+
+nominal_datatype tyS = 
+  All xs::"name list" ty::"ty" bind xs in ty
+
+
+
+(* example 1 from Terms.thy *)
+
+nominal_datatype trm1 =
+  Vr1 "name"
+| Ap1 "trm1" "trm1"
+| Lm1 x::"name" t::"trm1"      bind x in t 
+| Lt1 p::"bp1" "trm1" t::"trm1" bind "bv1 p" in t 
+and bp1 =
+  BUnit1
+| BV1 "name"
+| BP1 "bp1" "bp1"
+binder
+  bv1
+where
+  "bv1 (BUnit1) = {}"
+| "bv1 (BV1 x) = {atom x}"
+| "bv1 (BP1 bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)"
+
+thm bv1_raw.simps
+
+(* example 2 from Terms.thy *)
+
+nominal_datatype trm2 =
+  Vr2 "name"
+| Ap2 "trm2" "trm2"
+| Lm2 x::"name" t::"trm2"       bind x in t
+| Lt2 r::"rassign" t::"trm2"    bind "bv2 r" in t
+and rassign = 
+  As "name" "trm2"
+binder
+  bv2
+where
+  "bv2 (As x t) = {atom x}"
+
+(* example 3 from Terms.thy *)
+
+nominal_datatype trm3 =
+  Vr3 "name"
+| Ap3 "trm3" "trm3"
+| Lm3 x::"name" t::"trm3"        bind x in t
+| Lt3 r::"rassigns3" t::"trm3"   bind "bv3 r" in t
+and rassigns3 =
+  ANil
+| ACons "name" "trm3" "rassigns3"
+binder
+  bv3
+where
+  "bv3 ANil = {}"
+| "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)"
+
+(* example 4 from Terms.thy *)
+
+nominal_datatype trm4 =
+  Vr4 "name"
+| Ap4 "trm4" "trm4 list"
+| Lm4 x::"name" t::"trm4"  bind x in t
+
+(* example 5 from Terms.thy *)
+
+nominal_datatype trm5 =
+  Vr5 "name"
+| Ap5 "trm5" "trm5"
+| Lt5 l::"lts" t::"trm5"  bind "bv5 l" in t
+and lts =
+  Lnil
+| Lcons "name" "trm5" "lts"
+binder
+  bv5
+where
+  "bv5 Lnil = {}"
+| "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)"
+
+(* example 6 from Terms.thy *)
+
+nominal_datatype trm6 =
+  Vr6 "name"
+| Lm6 x::"name" t::"trm6"         bind x in t
+| Lt6 left::"trm6" right::"trm6"  bind "bv6 left" in right
+binder
+  bv6
+where
+  "bv6 (Vr6 n) = {}"
+| "bv6 (Lm6 n t) = {atom n} \<union> bv6 t"
+| "bv6 (Lt6 l r) = bv6 l \<union> bv6 r"
+
+(* example 7 from Terms.thy *)
+
+nominal_datatype trm7 =
+  Vr7 "name"
+| Lm7 l::"name" r::"trm7"   bind l in r
+| Lt7 l::"trm7" r::"trm7"   bind "bv7 l" in r
+binder 
+  bv7 
+where
+  "bv7 (Vr7 n) = {atom n}"
+| "bv7 (Lm7 n t) = bv7 t - {atom n}"
+| "bv7 (Lt7 l r) = bv7 l \<union> bv7 r"
+
+(* example 8 from Terms.thy *)
+
+nominal_datatype foo8 =
+  Foo0 "name"
+| Foo1 b::"bar8" f::"foo8" bind "bv8 b" in foo
+and bar8 =
+  Bar0 "name"
+| Bar1 "name" s::"name" b::"bar8" bind s in b
+binder 
+  bv8
+where
+  "bv8 (Bar0 x) = {}"
+| "bv8 (Bar1 v x b) = {atom v}"
+
+(* example 9 from Terms.thy *)
+
+nominal_datatype lam9 =
+  Var9 "name"
+| Lam9 n::"name" l::"lam9" bind n in l
+and bla9 =
+  Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s
+binder
+  bv9
+where
+  "bv9 (Var9 x) = {}"
+| "bv9 (Lam9 x b) = {atom x}"
+
+end
+
+
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/nominal_atoms.ML	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,94 @@
+(*  Title:      nominal_atoms/ML
+    Authors:    Brian Huffman, Christian Urban
+
+    Command for defining concrete atom types. 
+    
+    At the moment, only single-sorted atom types
+    are supported. 
+*)
+
+signature ATOM_DECL =
+sig
+  val add_atom_decl: (binding * (binding option)) -> theory -> theory
+end;
+
+structure Atom_Decl :> ATOM_DECL =
+struct
+
+val atomT = @{typ atom};
+val permT = @{typ perm};
+
+val sort_of_const = @{term sort_of};
+fun atom_const T = Const (@{const_name atom}, T --> atomT);
+fun permute_const T = Const (@{const_name permute}, permT --> T --> T);
+
+fun mk_sort_of t = sort_of_const $ t;
+fun mk_atom t = atom_const (fastype_of t) $ t;
+fun mk_permute (p, t) = permute_const (fastype_of t) $ p $ t;
+
+fun atom_decl_set (str : string) : term =
+  let
+    val a = Free ("a", atomT);
+    val s = Const (@{const_name "Sort"}, @{typ "string => atom_sort list => atom_sort"})
+              $ HOLogic.mk_string str $ HOLogic.nil_const @{typ "atom_sort"};
+  in
+    HOLogic.mk_Collect ("a", atomT, HOLogic.mk_eq (mk_sort_of a, s))
+  end
+
+fun add_atom_decl (name : binding, arg : binding option) (thy : theory) =
+  let
+    val _ = Theory.requires thy "Nominal2_Atoms" "nominal logic";
+    val str = Sign.full_name thy name;
+
+    (* typedef *)
+    val set = atom_decl_set str;
+    val tac = rtac @{thm exists_eq_simple_sort} 1;
+    val ((full_tname, info as {type_definition, Rep_name, Abs_name, ...}), thy) =
+      Typedef.add_typedef false NONE (name, [], NoSyn) set NONE tac thy;
+
+    (* definition of atom and permute *)
+    val newT = #abs_type info;
+    val RepC = Const (Rep_name, newT --> atomT);
+    val AbsC = Const (Abs_name, atomT --> newT);
+    val a = Free ("a", newT);
+    val p = Free ("p", permT);
+    val atom_eqn =
+      HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_atom a, RepC $ a));
+    val permute_eqn =
+      HOLogic.mk_Trueprop (HOLogic.mk_eq
+        (mk_permute (p, a), AbsC $ (mk_permute (p, RepC $ a))));
+    val atom_def_name =
+      Binding.prefix_name "atom_" (Binding.suffix_name "_def" name);
+    val permute_def_name =
+      Binding.prefix_name "permute_" (Binding.suffix_name "_def" name);
+
+    (* at class instance *)
+    val lthy =
+      Theory_Target.instantiation ([full_tname], [], @{sort at}) thy;
+    val ((_, (_, permute_ldef)), lthy) =
+      Specification.definition (NONE, ((permute_def_name, []), permute_eqn)) lthy;
+    val ((_, (_, atom_ldef)), lthy) =
+      Specification.definition (NONE, ((atom_def_name, []), atom_eqn)) lthy;
+    val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
+    val permute_def = singleton (ProofContext.export lthy ctxt_thy) permute_ldef;
+    val atom_def = singleton (ProofContext.export lthy ctxt_thy) atom_ldef;
+    val class_thm = @{thm at_class} OF [type_definition, atom_def, permute_def];
+    val thy = lthy
+      |> Class.prove_instantiation_instance (K (Tactic.rtac class_thm 1))
+      |> Local_Theory.exit_global;
+  in
+    thy
+  end;
+
+(** outer syntax **)
+
+local structure P = OuterParse and K = OuterKeyword in
+
+val _ =
+  OuterSyntax.command "atom_decl" "declaration of a concrete atom type" K.thy_decl
+    ((P.binding -- Scan.option (Args.parens (P.binding))) >>
+      (Toplevel.print oo (Toplevel.theory o add_atom_decl)));
+
+end;
+
+end;
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/nominal_permeq.ML	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,71 @@
+(*  Title:      nominal_thmdecls.ML
+    Author:     Brian Huffman, Christian Urban
+*)
+
+signature NOMINAL_PERMEQ =
+sig
+  val eqvt_tac: Proof.context -> int -> tactic 
+
+end;
+
+(* TODO:
+
+ - provide a method interface with the usual add and del options
+
+ - print a warning if for a constant no eqvt lemma is stored
+
+ - there seems to be too much simplified in case of multiple 
+   permutations, like
+
+      p o q o r o x 
+
+   we usually only want the outermost permutation to "float" in
+*)
+
+
+structure Nominal_Permeq: NOMINAL_PERMEQ =
+struct
+
+local
+
+fun eqvt_apply_conv ctxt ct =
+  case (term_of ct) of
+    (Const (@{const_name "permute"}, _) $ _ $ (_ $ _)) =>
+      let
+        val (perm, t) = Thm.dest_comb ct
+        val (_, p) = Thm.dest_comb perm
+        val (f, x) = Thm.dest_comb t
+        val a = ctyp_of_term x;
+        val b = ctyp_of_term t;
+        val ty_insts = map SOME [b, a]
+        val term_insts = map SOME [p, f, x]
+      in
+        Drule.instantiate' ty_insts term_insts @{thm eqvt_apply}
+      end
+  | _ => Conv.no_conv ct
+
+fun eqvt_lambda_conv ctxt ct =
+  case (term_of ct) of
+    (Const (@{const_name "permute"}, _) $ _ $ Abs _) =>
+      Conv.rewr_conv @{thm eqvt_lambda} ct
+  | _ => Conv.no_conv ct
+
+in
+
+fun eqvt_conv ctxt ct =
+  Conv.first_conv
+    [ Conv.rewr_conv @{thm eqvt_bound},
+      eqvt_apply_conv ctxt
+        then_conv Conv.comb_conv (eqvt_conv ctxt),
+      eqvt_lambda_conv ctxt
+        then_conv Conv.abs_conv (fn (v, ctxt) => eqvt_conv ctxt) ctxt,
+      More_Conv.rewrs_conv (Nominal_ThmDecls.get_eqvts_raw_thms ctxt),
+      Conv.all_conv
+    ] ct
+
+fun eqvt_tac ctxt = 
+  CONVERSION (More_Conv.bottom_conv (fn ctxt => eqvt_conv ctxt) ctxt)
+
+end
+
+end; (* structure *)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/Nominal/nominal_thmdecls.ML	Thu Feb 25 07:48:33 2010 +0100
@@ -0,0 +1,134 @@
+(*  Title:      nominal_thmdecls.ML
+    Author:     Christian Urban
+
+  Infrastructure for the lemma collection "eqvts".
+
+  Provides the attributes [eqvt] and [eqvt_raw], and the theorem
+  lists eqvts and eqvts_raw. The first attribute will store the 
+  theorem in the eqvts list and also in the eqvts_raw list. For 
+  the latter the theorem is expected to be of the form
+
+    p o (c x1 x2 ...) = c (p o x1) (p o x2) ...
+
+  and it is stored in the form
+
+    p o c == c
+
+  The [eqvt_raw] attribute just adds the theorem to eqvts_raw.
+
+  TODO:
+
+   - deal with eqvt-lemmas of the form 
+
+       c x1 x2 ... ==> c (p o x1) (p o x2) ..
+*)
+
+signature NOMINAL_THMDECLS =
+sig
+  val eqvt_add: attribute
+  val eqvt_del: attribute
+  val eqvt_raw_add: attribute
+  val eqvt_raw_del: attribute
+  val setup: theory -> theory
+  val get_eqvts_thms: Proof.context -> thm list
+  val get_eqvts_raw_thms: Proof.context -> thm list
+
+end;
+
+structure Nominal_ThmDecls: NOMINAL_THMDECLS =
+struct
+
+
+structure EqvtData = Generic_Data
+( type T = thm Item_Net.T;
+  val empty = Thm.full_rules;
+  val extend = I;
+  val merge = Item_Net.merge );
+
+structure EqvtRawData = Generic_Data
+( type T = thm Item_Net.T;
+  val empty = Thm.full_rules;
+  val extend = I;
+  val merge = Item_Net.merge );
+
+val eqvts = Item_Net.content o EqvtData.get;
+val eqvts_raw = Item_Net.content o EqvtRawData.get;
+
+val get_eqvts_thms = eqvts o Context.Proof; 
+val get_eqvts_raw_thms = eqvts_raw o Context.Proof; 
+
+val add_thm = EqvtData.map o Item_Net.update;
+val del_thm = EqvtData.map o Item_Net.remove;
+
+val add_raw_thm = EqvtRawData.map o Item_Net.update;
+val del_raw_thm = EqvtRawData.map o Item_Net.remove;
+
+fun dest_perm (Const (@{const_name "permute"}, _) $ p $ t) = (p, t)
+  | dest_perm t = raise TERM("dest_perm", [t])
+
+fun mk_perm p trm =
+let
+  val ty = fastype_of trm
+in
+  Const (@{const_name "permute"}, @{typ "perm"} --> ty --> ty) $ p $ trm
+end
+
+fun eqvt_transform_tac thm = REPEAT o FIRST' 
+  [CHANGED o simp_tac (HOL_basic_ss addsimps @{thms permute_minus_cancel}),
+   rtac (thm RS @{thm trans}),
+   rtac @{thm trans[OF permute_fun_def]} THEN' rtac @{thm ext}]
+
+(* transform equations into the required form *)
+fun transform_eq ctxt thm lhs rhs = 
+let
+  val (p, t) = dest_perm lhs
+  val (c, args) = strip_comb t
+  val (c', args') = strip_comb rhs 
+  val eargs = map Envir.eta_contract args 
+  val eargs' = map Envir.eta_contract args'
+  val p_str = fst (fst (dest_Var p))
+  val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_perm p c, c))
+in
+  if c <> c' 
+    then error "eqvt lemma is not of the right form (constants do not agree)"
+  else if eargs' <> map (mk_perm p) eargs 
+    then error "eqvt lemma is not of the right form (arguments do not agree)"
+  else if args = [] 
+    then thm
+  else Goal.prove ctxt [p_str] [] goal
+    (fn _ => eqvt_transform_tac thm 1)
+end
+
+fun transform addel_fun thm context = 
+let
+  val ctxt = Context.proof_of context
+  val trm = HOLogic.dest_Trueprop (prop_of thm)
+in
+  case trm of
+    Const (@{const_name "op ="}, _) $ lhs $ rhs =>
+      let 
+        val thm' = transform_eq ctxt thm lhs rhs RS @{thm eq_reflection}
+      in
+        addel_fun thm' context
+      end
+  | _ => raise (error "only (op=) case implemented yet")
+end 
+
+val eqvt_add = Thm.declaration_attribute (fn thm => (add_thm thm) o (transform add_raw_thm thm));
+val eqvt_del = Thm.declaration_attribute (fn thm => (del_thm thm) o (transform del_raw_thm thm));
+
+val eqvt_raw_add = Thm.declaration_attribute add_raw_thm;
+val eqvt_raw_del = Thm.declaration_attribute del_raw_thm;
+
+val setup =
+  Attrib.setup @{binding "eqvt"} (Attrib.add_del eqvt_add eqvt_del) 
+    (cat_lines ["declaration of equivariance lemmas - they will automtically be",  
+       "brought into the form p o c = c"]) #>
+  Attrib.setup @{binding "eqvt_raw"} (Attrib.add_del eqvt_raw_add eqvt_raw_del) 
+    (cat_lines ["declaration of equivariance lemmas - no",
+       "transformation is performed"]) #>
+  PureThy.add_thms_dynamic (@{binding "eqvts"}, eqvts) #>
+  PureThy.add_thms_dynamic (@{binding "eqvts_raw"}, eqvts_raw);
+
+
+end;
--- a/Quot/Nominal/Abs.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,506 +0,0 @@
-theory Abs
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "../Quotient" "../Quotient_Product"
-begin
-
-(* the next three lemmas that should be in Nominal \<dots>\<dots>must be cleaned *)
-lemma ball_image: 
-  shows "(\<forall>x \<in> p \<bullet> S. P x) = (\<forall>x \<in> S. P (p \<bullet> x))"
-apply(auto)
-apply(drule_tac x="p \<bullet> x" in bspec)
-apply(simp add: mem_permute_iff)
-apply(simp)
-apply(drule_tac x="(- p) \<bullet> x" in bspec)
-apply(rule_tac p1="p" in mem_permute_iff[THEN iffD1])
-apply(simp)
-apply(simp)
-done
-
-lemma fresh_star_plus:
-  fixes p q::perm
-  shows "\<lbrakk>a \<sharp>* p;  a \<sharp>* q\<rbrakk> \<Longrightarrow> a \<sharp>* (p + q)"
-  unfolding fresh_star_def
-  by (simp add: fresh_plus_perm)
-
-lemma fresh_star_permute_iff:
-  shows "(p \<bullet> a) \<sharp>* (p \<bullet> x) \<longleftrightarrow> a \<sharp>* x"
-apply(simp add: fresh_star_def)
-apply(simp add: ball_image)
-apply(simp add: fresh_permute_iff)
-done
-
-fun
-  alpha_gen 
-where
-  alpha_gen[simp del]:
-  "alpha_gen (bs, x) R f pi (cs, y) \<longleftrightarrow> f x - bs = f y - cs \<and> (f x - bs) \<sharp>* pi \<and> R (pi \<bullet> x) y"
-
-notation
-  alpha_gen ("_ \<approx>gen _ _ _ _" [100, 100, 100, 100, 100] 100)
-
-lemma [mono]: "R1 \<le> R2 \<Longrightarrow> alpha_gen x R1 \<le> alpha_gen x R2"
-  by (cases x) (auto simp add: le_fun_def le_bool_def alpha_gen.simps)
-
-lemma alpha_gen_refl:
-  assumes a: "R x x"
-  shows "(bs, x) \<approx>gen R f 0 (bs, x)"
-  using a by (simp add: alpha_gen fresh_star_def fresh_zero_perm)
-
-lemma alpha_gen_sym:
-  assumes a: "(bs, x) \<approx>gen R f p (cs, y)"
-  and     b: "R (p \<bullet> x) y \<Longrightarrow> R (- p \<bullet> y) x"
-  shows "(cs, y) \<approx>gen R f (- p) (bs, x)"
-  using a b by (simp add: alpha_gen fresh_star_def fresh_def supp_minus_perm)
-
-lemma alpha_gen_trans:
-  assumes a: "(bs, x) \<approx>gen R f p1 (cs, y)"
-  and     b: "(cs, y) \<approx>gen R f p2 (ds, z)"
-  and     c: "\<lbrakk>R (p1 \<bullet> x) y; R (p2 \<bullet> y) z\<rbrakk> \<Longrightarrow> R ((p2 + p1) \<bullet> x) z"
-  shows "(bs, x) \<approx>gen R f (p2 + p1) (ds, z)"
-  using a b c using supp_plus_perm
-  apply(simp add: alpha_gen fresh_star_def fresh_def)
-  apply(blast)
-  done
-
-lemma alpha_gen_eqvt:
-  assumes a: "(bs, x) \<approx>gen R f q (cs, y)"
-  and     b: "R (q \<bullet> x) y \<Longrightarrow> R (p \<bullet> (q \<bullet> x)) (p \<bullet> y)"
-  and     c: "p \<bullet> (f x) = f (p \<bullet> x)"
-  and     d: "p \<bullet> (f y) = f (p \<bullet> y)"
-  shows "(p \<bullet> bs, p \<bullet> x) \<approx>gen R f (p \<bullet> q) (p \<bullet> cs, p \<bullet> y)"
-  using a b
-  apply(simp add: alpha_gen c[symmetric] d[symmetric] Diff_eqvt[symmetric])
-  apply(simp add: permute_eqvt[symmetric])
-  apply(simp add: fresh_star_permute_iff)
-  apply(clarsimp)
-  done
-
-lemma alpha_gen_compose_sym:
-  assumes b: "\<exists>pi. (aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R x2 x1) f pi (ab, s)"
-  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
-  shows "\<exists>pi. (ab, s) \<approx>gen R f pi (aa, t)"
-  using b apply -
-  apply(erule exE)
-  apply(rule_tac x="- pi" in exI)
-  apply(simp add: alpha_gen.simps)
-  apply(erule conjE)+
-  apply(rule conjI)
-  apply(simp add: fresh_star_def fresh_minus_perm)
-  apply(subgoal_tac "R (- pi \<bullet> s) ((- pi) \<bullet> (pi \<bullet> t))")
-  apply simp
-  apply(rule a)
-  apply assumption
-  done
-
-lemma alpha_gen_compose_trans:
-  assumes b: "\<exists>pi\<Colon>perm. (aa, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> (\<forall>x. R x2 x \<longrightarrow> R x1 x)) f pi (ab, ta)"
-  and c: "\<exists>pi\<Colon>perm. (ab, ta) \<approx>gen R f pi (ac, sa)"
-  and a: "\<And>pi t s. (R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s))"
-  shows "\<exists>pi\<Colon>perm. (aa, t) \<approx>gen R f pi (ac, sa)"
-  using b c apply -
-  apply(simp add: alpha_gen.simps)
-  apply(erule conjE)+
-  apply(erule exE)+
-  apply(erule conjE)+
-  apply(rule_tac x="pia + pi" in exI)
-  apply(simp add: fresh_star_plus)
-  apply(drule_tac x="- pia \<bullet> sa" in spec)
-  apply(drule mp)
-  apply(rotate_tac 4)
-  apply(drule_tac pi="- pia" in a)
-  apply(simp)
-  apply(rotate_tac 6)
-  apply(drule_tac pi="pia" in a)
-  apply(simp)
-  done
-
-lemma alpha_gen_atom_eqvt:
-  assumes a: "\<And>x. pi \<bullet> (f x) = f (pi \<bullet> x)"
-  and     b: "\<exists>pia. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. R x1 x2 \<and> R (pi \<bullet> x1) (pi \<bullet> x2)) f pia ({atom b}, s)"
-  shows  "\<exists>pia. ({atom (pi \<bullet> a)}, pi \<bullet> t) \<approx>gen R f pia ({atom (pi \<bullet> b)}, pi \<bullet> s)"
-  using b 
-  apply -
-  apply(erule exE)
-  apply(rule_tac x="pi \<bullet> pia" in exI)
-  apply(simp add: alpha_gen.simps)
-  apply(erule conjE)+
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
-  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt)
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
-  apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt)
-  apply(subst permute_eqvt[symmetric])
-  apply(simp)
-  done
-
-fun
-  alpha_abs 
-where
-  "alpha_abs (bs, x) (cs, y) = (\<exists>p. (bs, x) \<approx>gen (op=) supp p (cs, y))"
-
-notation
-  alpha_abs ("_ \<approx>abs _")
-
-lemma alpha_abs_swap:
-  assumes a1: "a \<notin> (supp x) - bs"
-  and     a2: "b \<notin> (supp x) - bs"
-  shows "(bs, x) \<approx>abs ((a \<rightleftharpoons> b) \<bullet> bs, (a \<rightleftharpoons> b) \<bullet> x)"
-  apply(simp)
-  apply(rule_tac x="(a \<rightleftharpoons> b)" in exI)
-  apply(simp add: alpha_gen)
-  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
-  apply(simp add: swap_set_not_in[OF a1 a2])
-  apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
-  using a1 a2
-  apply(simp add: fresh_star_def fresh_def)
-  apply(blast)
-  apply(simp add: supp_swap)
-  done
-
-fun
-  supp_abs_fun
-where
-  "supp_abs_fun (bs, x) = (supp x) - bs"
-
-lemma supp_abs_fun_lemma:
-  assumes a: "x \<approx>abs y" 
-  shows "supp_abs_fun x = supp_abs_fun y"
-  using a
-  apply(induct rule: alpha_abs.induct)
-  apply(simp add: alpha_gen)
-  done
-  
-quotient_type 'a abs = "(atom set \<times> 'a::pt)" / "alpha_abs"
-  apply(rule equivpI)
-  unfolding reflp_def symp_def transp_def
-  apply(simp_all)
-  (* refl *)
-  apply(clarify)
-  apply(rule exI)
-  apply(rule alpha_gen_refl)
-  apply(simp)
-  (* symm *)
-  apply(clarify)
-  apply(rule exI)
-  apply(rule alpha_gen_sym)
-  apply(assumption)
-  apply(clarsimp)
-  (* trans *)
-  apply(clarify)
-  apply(rule exI)
-  apply(rule alpha_gen_trans)
-  apply(assumption)
-  apply(assumption)
-  apply(simp)
-  done
-
-quotient_definition
-  "Abs::atom set \<Rightarrow> ('a::pt) \<Rightarrow> 'a abs"
-is
-  "Pair::atom set \<Rightarrow> ('a::pt) \<Rightarrow> (atom set \<times> 'a)"
-
-lemma [quot_respect]:
-  shows "((op =) ===> (op =) ===> alpha_abs) Pair Pair"
-  apply(clarsimp)
-  apply(rule exI)
-  apply(rule alpha_gen_refl)
-  apply(simp)
-  done
-
-lemma [quot_respect]:
-  shows "((op =) ===> alpha_abs ===> alpha_abs) permute permute"
-  apply(clarsimp)
-  apply(rule exI)
-  apply(rule alpha_gen_eqvt)
-  apply(assumption)
-  apply(simp_all add: supp_eqvt)
-  done
-
-lemma [quot_respect]:
-  shows "(alpha_abs ===> (op =)) supp_abs_fun supp_abs_fun"
-  apply(simp add: supp_abs_fun_lemma)
-  done
-
-lemma abs_induct:
-  "\<lbrakk>\<And>as (x::'a::pt). P (Abs as x)\<rbrakk> \<Longrightarrow> P t"
-  apply(lifting prod.induct[where 'a="atom set" and 'b="'a"])
-  done
-
-(* TEST case *)
-lemmas abs_induct2 = prod.induct[where 'a="atom set" and 'b="'a::pt", quot_lifted]
-thm abs_induct abs_induct2
-
-instantiation abs :: (pt) pt
-begin
-
-quotient_definition
-  "permute_abs::perm \<Rightarrow> ('a::pt abs) \<Rightarrow> 'a abs"
-is
-  "permute:: perm \<Rightarrow> (atom set \<times> 'a::pt) \<Rightarrow> (atom set \<times> 'a::pt)"
-
-lemma permute_ABS [simp]:
-  fixes x::"'a::pt"  (* ??? has to be 'a \<dots> 'b does not work *)
-  shows "(p \<bullet> (Abs as x)) = Abs (p \<bullet> as) (p \<bullet> x)"
-  by (lifting permute_prod.simps(1)[where 'a="atom set" and 'b="'a"])
-
-instance
-  apply(default)
-  apply(induct_tac [!] x rule: abs_induct)
-  apply(simp_all)
-  done
-
-end
-
-quotient_definition
-  "supp_Abs_fun :: ('a::pt) abs \<Rightarrow> atom \<Rightarrow> bool"
-is
-  "supp_abs_fun"
-
-lemma supp_Abs_fun_simp:
-  shows "supp_Abs_fun (Abs bs x) = (supp x) - bs"
-  by (lifting supp_abs_fun.simps(1))
-
-lemma supp_Abs_fun_eqvt [eqvt]:
-  shows "(p \<bullet> supp_Abs_fun x) = supp_Abs_fun (p \<bullet> x)"
-  apply(induct_tac x rule: abs_induct)
-  apply(simp add: supp_Abs_fun_simp supp_eqvt Diff_eqvt)
-  done
-
-lemma supp_Abs_fun_fresh:
-  shows "a \<sharp> Abs bs x \<Longrightarrow> a \<sharp> supp_Abs_fun (Abs bs x)"
-  apply(rule fresh_fun_eqvt_app)
-  apply(simp add: eqvts_raw)
-  apply(simp)
-  done
-
-lemma Abs_swap:
-  assumes a1: "a \<notin> (supp x) - bs"
-  and     a2: "b \<notin> (supp x) - bs"
-  shows "(Abs bs x) = (Abs ((a \<rightleftharpoons> b) \<bullet> bs) ((a \<rightleftharpoons> b) \<bullet> x))"
-  using a1 a2 by (lifting alpha_abs_swap)
-
-lemma Abs_supports:
-  shows "((supp x) - as) supports (Abs as x)"
-  unfolding supports_def
-  apply(clarify)
-  apply(simp (no_asm))
-  apply(subst Abs_swap[symmetric])
-  apply(simp_all)
-  done
-
-lemma supp_Abs_subset1:
-  fixes x::"'a::fs"
-  shows "(supp x) - as \<subseteq> supp (Abs as x)"
-  apply(simp add: supp_conv_fresh)
-  apply(auto)
-  apply(drule_tac supp_Abs_fun_fresh)
-  apply(simp only: supp_Abs_fun_simp)
-  apply(simp add: fresh_def)
-  apply(simp add: supp_finite_atom_set finite_supp)
-  done
-
-lemma supp_Abs_subset2:
-  fixes x::"'a::fs"
-  shows "supp (Abs as x) \<subseteq> (supp x) - as"
-  apply(rule supp_is_subset)
-  apply(rule Abs_supports)
-  apply(simp add: finite_supp)
-  done
-
-lemma supp_Abs:
-  fixes x::"'a::fs"
-  shows "supp (Abs as x) = (supp x) - as"
-  apply(rule_tac subset_antisym)
-  apply(rule supp_Abs_subset2)
-  apply(rule supp_Abs_subset1)
-  done
-
-instance abs :: (fs) fs
-  apply(default)
-  apply(induct_tac x rule: abs_induct)
-  apply(simp add: supp_Abs)
-  apply(simp add: finite_supp)
-  done
-
-lemma Abs_fresh_iff:
-  fixes x::"'a::fs"
-  shows "a \<sharp> Abs bs x \<longleftrightarrow> a \<in> bs \<or> (a \<notin> bs \<and> a \<sharp> x)"
-  apply(simp add: fresh_def)
-  apply(simp add: supp_Abs)
-  apply(auto)
-  done
-
-lemma Abs_eq_iff:
-  shows "Abs bs x = Abs cs y \<longleftrightarrow> (\<exists>p. (bs, x) \<approx>gen (op =) supp p (cs, y))"
-  by (lifting alpha_abs.simps(1))
-
-
-
-(* 
-  below is a construction site for showing that in the
-  single-binder case, the old and new alpha equivalence 
-  coincide
-*)
-
-fun
-  alpha1
-where
-  "alpha1 (a, x) (b, y) \<longleftrightarrow> (a = b \<and> x = y) \<or> (a \<noteq> b \<and> x = (a \<rightleftharpoons> b) \<bullet> y \<and> a \<sharp> y)"
-
-notation 
-  alpha1 ("_ \<approx>abs1 _")
-
-thm swap_set_not_in
-
-lemma qq:
-  fixes S::"atom set"
-  assumes a: "supp p \<inter> S = {}"
-  shows "p \<bullet> S = S"
-using a
-apply(simp add: supp_perm permute_set_eq)
-apply(auto)
-apply(simp only: disjoint_iff_not_equal)
-apply(simp)
-apply (metis permute_atom_def_raw)
-apply(rule_tac x="(- p) \<bullet> x" in exI)
-apply(simp)
-apply(simp only: disjoint_iff_not_equal)
-apply(simp)
-apply(metis permute_minus_cancel)
-done
-
-lemma alpha_abs_swap:
-  assumes a1: "(supp x - bs) \<sharp>* p"
-  and     a2: "(supp x - bs) \<sharp>* p"
-  shows "(bs, x) \<approx>abs (p \<bullet> bs, p \<bullet> x)"
-  apply(simp)
-  apply(rule_tac x="p" in exI)
-  apply(simp add: alpha_gen)
-  apply(simp add: supp_eqvt[symmetric] Diff_eqvt[symmetric])
-  apply(rule conjI)
-  apply(rule sym)
-  apply(rule qq)
-  using a1 a2
-  apply(auto)[1]
-  oops
-
-
-
-lemma
-  assumes a: "(a, x) \<approx>abs1 (b, y)" "sort_of a = sort_of b"
-  shows "({a}, x) \<approx>abs ({b}, y)"
-using a
-apply(simp)
-apply(erule disjE)
-apply(simp)
-apply(rule exI)
-apply(rule alpha_gen_refl)
-apply(simp)
-apply(rule_tac x="(a \<rightleftharpoons> b)" in  exI)
-apply(simp add: alpha_gen)
-apply(simp add: fresh_def)
-apply(rule conjI)
-apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in  permute_eq_iff[THEN iffD1])
-apply(rule trans)
-apply(simp add: Diff_eqvt supp_eqvt)
-apply(subst swap_set_not_in)
-back
-apply(simp)
-apply(simp)
-apply(simp add: permute_set_eq)
-apply(rule_tac ?p1="(a \<rightleftharpoons> b)" in fresh_star_permute_iff[THEN iffD1])
-apply(simp add: permute_self)
-apply(simp add: Diff_eqvt supp_eqvt)
-apply(simp add: permute_set_eq)
-apply(subgoal_tac "supp (a \<rightleftharpoons> b) \<subseteq> {a, b}")
-apply(simp add: fresh_star_def fresh_def)
-apply(blast)
-apply(simp add: supp_swap)
-done
-
-thm supp_perm
-
-lemma perm_induct_test:
-  fixes P :: "perm => bool"
-  assumes zero: "P 0"
-  assumes swap: "\<And>a b. \<lbrakk>sort_of a = sort_of b; a \<noteq> b\<rbrakk> \<Longrightarrow> P (a \<rightleftharpoons> b)"
-  assumes plus: "\<And>p1 p2. \<lbrakk>supp (p1 + p2) = (supp p1 \<union> supp p2); P p1; P p2\<rbrakk> \<Longrightarrow> P (p1 + p2)"
-  shows "P p"
-sorry
-
-lemma tt1:
-  assumes a: "finite (supp p)"
-  shows "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
-using a
-unfolding fresh_star_def fresh_def
-apply(induct F\<equiv>"supp p" arbitrary: p rule: finite.induct)
-apply(simp add: supp_perm)
-defer
-apply(case_tac "a \<in> A")
-apply(simp add: insert_absorb)
-apply(subgoal_tac "A = supp p - {a}")
-prefer 2
-apply(blast)
-apply(case_tac "p \<bullet> a = a")
-apply(simp add: supp_perm)
-apply(drule_tac x="p + (((- p) \<bullet> a) \<rightleftharpoons> a)" in meta_spec)
-apply(simp)
-apply(drule meta_mp)
-apply(rule subset_antisym)
-apply(rule subsetI)
-apply(simp)
-apply(simp add: supp_perm)
-apply(case_tac "xa = p \<bullet> a")
-apply(simp)
-apply(case_tac "p \<bullet> a = (- p) \<bullet> a")
-apply(simp)
-defer
-apply(simp)
-oops
-
-lemma tt:
-  "(supp x) \<sharp>* p \<Longrightarrow> p \<bullet> x = x"
-apply(induct p rule: perm_induct_test)
-apply(simp)
-apply(rule swap_fresh_fresh)
-apply(case_tac "a \<in> supp x")
-apply(simp add: fresh_star_def)
-apply(drule_tac x="a" in bspec)
-apply(simp)
-apply(simp add: fresh_def)
-apply(simp add: supp_swap)
-apply(simp add: fresh_def)
-apply(case_tac "b \<in> supp x")
-apply(simp add: fresh_star_def)
-apply(drule_tac x="b" in bspec)
-apply(simp)
-apply(simp add: fresh_def)
-apply(simp add: supp_swap)
-apply(simp add: fresh_def)
-apply(simp)
-apply(drule meta_mp)
-apply(simp add: fresh_star_def fresh_def)
-apply(drule meta_mp)
-apply(simp add: fresh_star_def fresh_def)
-apply(simp)
-done
-
-lemma yy:
-  assumes "S1 - {x} = S2 - {x}" "x \<in> S1" "x \<in> S2"
-  shows "S1 = S2"
-using assms
-apply (metis insert_Diff_single insert_absorb)
-done
-
-
-lemma
-  assumes a: "({a}, x) \<approx>abs ({b}, y)" "sort_of a = sort_of b"
-  shows "(a, x) \<approx>abs1 (b, y)"
-using a
-apply(case_tac "a = b")
-apply(simp)
-oops
-
-
-end
-
--- a/Quot/Nominal/Fv.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,399 +0,0 @@
-theory Fv
-imports "Nominal2_Atoms" "Abs"
-begin
-
-(* Bindings are given as a list which has a length being equal
-   to the length of the number of constructors.
-
-   Each element is a list whose length is equal to the number
-   of arguents.
-
-   Every element specifies bindings of this argument given as
-   a tuple: function, bound argument.
-
-  Eg:
-nominal_datatype
-
-   C1
- | C2 x y z bind x in z
- | C3 x y z bind f x in z bind g y in z
-
-yields:
-[
- [],
- [[], [], [(NONE, 0)]],
- [[], [], [(SOME (Const f), 0), (Some (Const g), 1)]]]
-
-A SOME binding has to have a function returning an atom set,
-and a NONE binding has to be on an argument that is an atom
-or an atom set.
-
-How the procedure works:
-  For each of the defined datatypes,
-  For each of the constructors,
-  It creates a union of free variables for each argument.
-
-  For an argument the free variables are the variables minus
-  bound variables.
-
-  The variables are:
-    For an atom, a singleton set with the atom itself.
-    For an atom set, the atom set itself.
-    For a recursive argument, the appropriate fv function applied to it.
-    (* TODO: This one is not implemented *)
-    For other arguments it should be an appropriate fv function stored
-      in the database.
-  The bound variables are a union of results of all bindings that
-  involve the given argument. For a paricular binding the result is:
-    For a function applied to an argument this function with the argument.
-    For an atom, a singleton set with the atom itself.
-    For an atom set, the atom set itself.
-    For a recursive argument, the appropriate fv function applied to it.
-    (* TODO: This one is not implemented *)
-    For other arguments it should be an appropriate fv function stored
-      in the database.
-*)
-
-ML {*
-  open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *);
-  (* TODO: It is the same as one in 'nominal_atoms' *)
-  fun mk_atom ty = Const (@{const_name atom}, ty --> @{typ atom});
-  val noatoms = @{term "{} :: atom set"};
-  fun mk_single_atom x = HOLogic.mk_set @{typ atom} [mk_atom (type_of x) $ x];
-  fun mk_union sets =
-    fold (fn a => fn b =>
-      if a = noatoms then b else
-      if b = noatoms then a else
-      HOLogic.mk_binop @{const_name union} (a, b)) (rev sets) noatoms;
-  fun mk_diff a b =
-    if b = noatoms then a else
-    if b = a then noatoms else
-    HOLogic.mk_binop @{const_name minus} (a, b);
-  fun mk_atoms t =
-    let
-      val ty = fastype_of t;
-      val atom_ty = HOLogic.dest_setT ty --> @{typ atom};
-      val img_ty = atom_ty --> ty --> @{typ "atom set"};
-    in
-      (Const (@{const_name image}, img_ty) $ Const (@{const_name atom}, atom_ty) $ t)
-    end;
-  (* Copy from Term *)
-  fun is_funtype (Type ("fun", [_, _])) = true
-    | is_funtype _ = false;
-  (* Similar to one in USyntax *)
-  fun mk_pair (fst, snd) =
-    let val ty1 = fastype_of fst
-      val ty2 = fastype_of snd
-      val c = HOLogic.pair_const ty1 ty2
-    in c $ fst $ snd
-    end;
-
-*}
-
-(* TODO: Notice datatypes without bindings and replace alpha with equality *)
-ML {*
-(* Currently needs just one full_tname to access Datatype *)
-fun define_fv_alpha full_tname bindsall lthy =
-let
-  val thy = ProofContext.theory_of lthy;
-  val {descr, ...} = Datatype.the_info thy full_tname;
-  val sorts = []; (* TODO *)
-  fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
-  val fv_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
-    "fv_" ^ name_of_typ (nth_dtyp i)) descr);
-  val fv_types = map (fn (i, _) => nth_dtyp i --> @{typ "atom set"}) descr;
-  val fv_frees = map Free (fv_names ~~ fv_types);
-  val alpha_names = Datatype_Prop.indexify_names (map (fn (i, _) =>
-    "alpha_" ^ name_of_typ (nth_dtyp i)) descr);
-  val alpha_types = map (fn (i, _) => nth_dtyp i --> nth_dtyp i --> @{typ bool}) descr;
-  val alpha_frees = map Free (alpha_names ~~ alpha_types);
-  fun fv_alpha_constr i (cname, dts) bindcs =
-    let
-      val Ts = map (typ_of_dtyp descr sorts) dts;
-      val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
-      val args = map Free (names ~~ Ts);
-      val names2 = Name.variant_list ("pi" :: names) (Datatype_Prop.make_tnames Ts);
-      val args2 = map Free (names2 ~~ Ts);
-      val c = Const (cname, Ts ---> (nth_dtyp i));
-      val fv_c = nth fv_frees i;
-      val alpha = nth alpha_frees i;
-      fun fv_bind args (NONE, i) =
-            if is_rec_type (nth dts i) then (nth fv_frees (body_index (nth dts i))) $ (nth args i) else
-            (* TODO we assume that all can be 'atomized' *)
-            if (is_funtype o fastype_of) (nth args i) then mk_atoms (nth args i) else
-            mk_single_atom (nth args i)
-        | fv_bind args (SOME f, i) = f $ (nth args i);
-      fun fv_arg ((dt, x), bindxs) =
-        let
-          val arg =
-            if is_rec_type dt then nth fv_frees (body_index dt) $ x else
-            (* TODO: we just assume everything can be 'atomized' *)
-            if (is_funtype o fastype_of) x then mk_atoms x else
-            HOLogic.mk_set @{typ atom} [mk_atom (fastype_of x) $ x]
-          val sub = mk_union (map (fv_bind args) bindxs)
-        in
-          mk_diff arg sub
-        end;
-      val fv_eq = HOLogic.mk_Trueprop (HOLogic.mk_eq
-        (fv_c $ list_comb (c, args), mk_union (map fv_arg (dts ~~ args ~~ bindcs))))
-      val alpha_rhs =
-        HOLogic.mk_Trueprop (alpha $ (list_comb (c, args)) $ (list_comb (c, args2)));
-      fun alpha_arg ((dt, bindxs), (arg, arg2)) =
-        if bindxs = [] then (
-          if is_rec_type dt then (nth alpha_frees (body_index dt) $ arg $ arg2)
-          else (HOLogic.mk_eq (arg, arg2)))
-        else
-          if is_rec_type dt then let
-            (* THE HARD CASE *)
-            val lhs_binds = mk_union (map (fv_bind args) bindxs);
-            val lhs = mk_pair (lhs_binds, arg);
-            val rhs_binds = mk_union (map (fv_bind args2) bindxs);
-            val rhs = mk_pair (rhs_binds, arg2);
-            val alpha = nth alpha_frees (body_index dt);
-            val fv = nth fv_frees (body_index dt);
-            val alpha_gen_pre = Const (@{const_name alpha_gen}, dummyT) $ lhs $ alpha $ fv $ (Free ("pi", @{typ perm})) $ rhs;
-            val alpha_gen_t = Syntax.check_term lthy alpha_gen_pre
-          in
-            HOLogic.mk_exists ("pi", @{typ perm}, alpha_gen_t)
-          (* TODO Add some test that is makes sense *)
-          end else @{term "True"}
-      val alpha_lhss = map (HOLogic.mk_Trueprop o alpha_arg) (dts ~~ bindcs ~~ (args ~~ args2))
-      val alpha_eq = Logic.list_implies (alpha_lhss, alpha_rhs)
-    in
-      (fv_eq, alpha_eq)
-    end;
-  fun fv_alpha_eq (i, (_, _, constrs)) binds = map2 (fv_alpha_constr i) constrs binds;
-  val (fv_eqs, alpha_eqs) = split_list (flat (map2 fv_alpha_eq descr bindsall))
-  val add_binds = map (fn x => (Attrib.empty_binding, x))
-  val (fvs, lthy') = (Primrec.add_primrec
-    (map (fn s => (Binding.name s, NONE, NoSyn)) fv_names) (add_binds fv_eqs) lthy)
-  val (alphas, lthy'') = (Inductive.add_inductive_i
-     {quiet_mode = false, verbose = true, alt_name = Binding.empty,
-      coind = false, no_elim = false, no_ind = false, skip_mono = true, fork_mono = false}
-     (map2 (fn x => fn y => ((Binding.name x, y), NoSyn)) alpha_names alpha_types) []
-     (add_binds alpha_eqs) [] lthy')
-in
-  ((fvs, alphas), lthy'')
-end
-*}
-
-(* tests
-atom_decl name
-
-datatype ty =
-  Var "name set"
-
-ML {* Syntax.check_term @{context} (mk_atoms @{term "a :: name set"}) *}
-
-local_setup {* define_fv_alpha "Fv.ty" [[[[]]]] *}
-print_theorems
-
-
-datatype rtrm1 =
-  rVr1 "name"
-| rAp1 "rtrm1" "rtrm1"
-| rLm1 "name" "rtrm1"        --"name is bound in trm1"
-| rLt1 "bp" "rtrm1" "rtrm1"   --"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) = {atom x}"
-| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp1)"
-
-setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Fv.rtrm1", "Fv.bp"] *}
-
-local_setup {* define_fv_alpha "Fv.rtrm1"
-  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(SOME @{term bv1}, 0)], [], [(SOME @{term bv1}, 0)]]],
-   [[], [[]], [[], []]]] *}
-print_theorems
-*)
-
-
-ML {*
-fun alpha_inj_tac dist_inj intrs elims =
-  SOLVED' (asm_full_simp_tac (HOL_ss addsimps intrs)) ORELSE'
-  (rtac @{thm iffI} THEN' RANGE [
-     (eresolve_tac elims THEN_ALL_NEW
-       asm_full_simp_tac (HOL_ss addsimps dist_inj)
-     ),
-     asm_full_simp_tac (HOL_ss addsimps intrs)])
-*}
-
-ML {*
-fun build_alpha_inj_gl thm =
-  let
-    val prop = prop_of thm;
-    val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop);
-    val hyps = map HOLogic.dest_Trueprop (Logic.strip_imp_prems prop);
-    fun list_conj l = foldr1 HOLogic.mk_conj l;
-  in
-    if hyps = [] then concl
-    else HOLogic.mk_eq (concl, list_conj hyps)
-  end;
-*}
-
-ML {*
-fun build_alpha_inj intrs dist_inj elims ctxt =
-let
-  val ((_, thms_imp), ctxt') = Variable.import false intrs ctxt;
-  val gls = map (HOLogic.mk_Trueprop o build_alpha_inj_gl) thms_imp;
-  fun tac _ = alpha_inj_tac dist_inj intrs elims 1;
-  val thms = map (fn gl => Goal.prove ctxt' [] [] gl tac) gls;
-in
-  Variable.export ctxt' ctxt thms
-end
-*}
-
-ML {*
-fun build_alpha_refl_gl alphas (x, y, z) =
-let
-  fun build_alpha alpha =
-    let
-      val ty = domain_type (fastype_of alpha);
-      val var = Free(x, ty);
-      val var2 = Free(y, ty);
-      val var3 = Free(z, ty);
-      val symp = HOLogic.mk_imp (alpha $ var $ var2, alpha $ var2 $ var);
-      val transp = HOLogic.mk_imp (alpha $ var $ var2,
-        HOLogic.mk_all (z, ty,
-          HOLogic.mk_imp (alpha $ var2 $ var3, alpha $ var $ var3)))
-    in
-      ((alpha $ var $ var), (symp, transp))
-    end;
-  val (refl_eqs, eqs) = split_list (map build_alpha alphas)
-  val (sym_eqs, trans_eqs) = split_list eqs
-  fun conj l = @{term Trueprop} $ foldr1 HOLogic.mk_conj l
-in
-  (conj refl_eqs, (conj sym_eqs, conj trans_eqs))
-end
-*}
-
-ML {*
-fun reflp_tac induct inj =
-  rtac induct THEN_ALL_NEW
-  asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
-  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
-  (rtac @{thm exI[of _ "0 :: perm"]} THEN'
-   asm_full_simp_tac (HOL_ss addsimps
-     @{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv}))
-*}
-
-ML {*
-fun symp_tac induct inj eqvt =
-  ((rtac @{thm impI} THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
-  asm_full_simp_tac (HOL_ss addsimps inj) THEN_ALL_NEW
-  TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
-  (etac @{thm alpha_gen_compose_sym} THEN' eresolve_tac eqvt)
-*}
-
-ML {*
-fun imp_elim_tac case_rules =
-  Subgoal.FOCUS (fn {concl, context, ...} =>
-    case term_of concl of
-      _ $ (_ $ asm $ _) =>
-        let
-          fun filter_fn case_rule = (
-            case Logic.strip_assums_hyp (prop_of case_rule) of
-              ((_ $ asmc) :: _) =>
-                let
-                  val thy = ProofContext.theory_of context
-                in
-                  Pattern.matches thy (asmc, asm)
-                end
-            | _ => false)
-          val matching_rules = filter filter_fn case_rules
-        in
-         (rtac impI THEN' rotate_tac (~1) THEN' eresolve_tac matching_rules) 1
-        end
-    | _ => no_tac
-  )
-*}
-
-ML {*
-fun transp_tac ctxt induct alpha_inj term_inj distinct cases eqvt =
-  ((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
-  (TRY o rtac allI THEN' imp_elim_tac cases ctxt) THEN_ALL_NEW
-  (
-    asm_full_simp_tac (HOL_ss addsimps alpha_inj @ term_inj @ distinct) THEN'
-    TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI) THEN_ALL_NEW
-    (etac @{thm alpha_gen_compose_trans} THEN' RANGE [atac, eresolve_tac eqvt])
-  )
-*}
-
-lemma transp_aux:
-  "(\<And>xa ya. R xa ya \<longrightarrow> (\<forall>z. R ya z \<longrightarrow> R xa z)) \<Longrightarrow> transp R"
-  unfolding transp_def
-  by blast
-
-ML {*
-fun equivp_tac reflps symps transps =
-  simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def})
-  THEN' rtac conjI THEN' rtac allI THEN'
-  resolve_tac reflps THEN'
-  rtac conjI THEN' rtac allI THEN' rtac allI THEN'
-  resolve_tac symps THEN'
-  rtac @{thm transp_aux} THEN' resolve_tac transps
-*}
-
-ML {*
-fun build_equivps alphas term_induct alpha_induct term_inj alpha_inj distinct cases eqvt ctxt =
-let
-  val ([x, y, z], ctxt') = Variable.variant_fixes ["x","y","z"] ctxt;
-  val (reflg, (symg, transg)) = build_alpha_refl_gl alphas (x, y, z)
-  fun reflp_tac' _ = reflp_tac term_induct alpha_inj 1;
-  fun symp_tac' _ = symp_tac alpha_induct alpha_inj eqvt 1;
-  fun transp_tac' _ = transp_tac ctxt alpha_induct alpha_inj term_inj distinct cases eqvt 1;
-  val reflt = Goal.prove ctxt' [] [] reflg reflp_tac';
-  val symt = Goal.prove ctxt' [] [] symg symp_tac';
-  val transt = Goal.prove ctxt' [] [] transg transp_tac';
-  val [refltg, symtg, transtg] = Variable.export ctxt' ctxt [reflt, symt, transt]
-  val reflts = HOLogic.conj_elims refltg
-  val symts = HOLogic.conj_elims symtg
-  val transts = HOLogic.conj_elims transtg
-  fun equivp alpha =
-    let
-      val equivp = Const (@{const_name equivp}, fastype_of alpha --> @{typ bool})
-      val goal = @{term Trueprop} $ (equivp $ alpha)
-      fun tac _ = equivp_tac reflts symts transts 1
-    in
-      Goal.prove ctxt [] [] goal tac
-    end
-in
-  map equivp alphas
-end
-*}
-
-(*
-Tests:
-prove alpha1_reflp_aux: {* fst (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *}
-by (tactic {* reflp_tac @{thm rtrm1_bp.induct} @{thms alpha1_inj} 1 *})
-
-prove alpha1_symp_aux: {* (fst o snd) (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *}
-by (tactic {* symp_tac @{thm alpha_rtrm1_alpha_bp.induct} @{thms alpha1_inj} @{thms alpha1_eqvt} 1 *})
-
-prove alpha1_transp_aux: {* (snd o snd) (build_alpha_refl_gl [@{term alpha_rtrm1}, @{term alpha_bp}] ("x","y","z")) *}
-by (tactic {* transp_tac @{context} @{thm alpha_rtrm1_alpha_bp.induct} @{thms alpha1_inj} @{thms rtrm1.inject bp.inject} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} 1 *})
-
-lemma alpha1_equivp:
-  "equivp alpha_rtrm1"
-  "equivp alpha_bp"
-apply (tactic {*
-  (simp_tac (HOL_ss addsimps @{thms equivp_reflp_symp_transp reflp_def symp_def})
-  THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN'
-  resolve_tac (HOLogic.conj_elims @{thm alpha1_reflp_aux})
-  THEN' rtac @{thm conjI} THEN' rtac @{thm allI} THEN' rtac @{thm allI} THEN'
-  resolve_tac (HOLogic.conj_elims @{thm alpha1_symp_aux}) THEN' rtac @{thm transp_aux}
-  THEN' resolve_tac (HOLogic.conj_elims @{thm alpha1_transp_aux})
-)
-1 *})
-done*)
-
-end
--- a/Quot/Nominal/LFex.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,236 +0,0 @@
-theory LFex
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp"
-begin
-
-atom_decl name
-atom_decl ident
-
-datatype rkind =
-    Type
-  | KPi "rty" "name" "rkind"
-and rty =
-    TConst "ident"
-  | TApp "rty" "rtrm"
-  | TPi "rty" "name" "rty"
-and rtrm =
-    Const "ident"
-  | Var "name"
-  | App "rtrm" "rtrm"
-  | Lam "rty" "name" "rtrm"
-
-
-setup {* snd o define_raw_perms ["rkind", "rty", "rtrm"] ["LFex.rkind", "LFex.rty", "LFex.rtrm"] *}
-
-local_setup {*
-  snd o define_fv_alpha "LFex.rkind"
-  [[ [], [[], [(NONE, 1)], [(NONE, 1)]] ],
-   [ [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]] ],
-   [ [[]], [[]], [[], []], [[], [(NONE, 1)], [(NONE, 1)]]]] *}
-notation
-    alpha_rkind  ("_ \<approx>ki _" [100, 100] 100)
-and alpha_rty    ("_ \<approx>ty _" [100, 100] 100)
-and alpha_rtrm   ("_ \<approx>tr _" [100, 100] 100)
-thm fv_rkind_fv_rty_fv_rtrm.simps alpha_rkind_alpha_rty_alpha_rtrm.intros
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_rkind_alpha_rty_alpha_rtrm_inj}, []), (build_alpha_inj @{thms alpha_rkind_alpha_rty_alpha_rtrm.intros} @{thms rkind.distinct rty.distinct rtrm.distinct rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases} ctxt)) ctxt)) *}
-thm alpha_rkind_alpha_rty_alpha_rtrm_inj
-
-lemma rfv_eqvt[eqvt]:
-  "((pi\<bullet>fv_rkind t1) = fv_rkind (pi\<bullet>t1))"
-  "((pi\<bullet>fv_rty t2) = fv_rty (pi\<bullet>t2))"
-  "((pi\<bullet>fv_rtrm t3) = fv_rtrm (pi\<bullet>t3))"
-apply(induct t1 and t2 and t3 rule: rkind_rty_rtrm.inducts)
-apply(simp_all add: union_eqvt Diff_eqvt)
-apply(simp_all add: permute_set_eq atom_eqvt)
-done
-
-lemma alpha_eqvt:
-  "t1 \<approx>ki s1 \<Longrightarrow> (pi \<bullet> t1) \<approx>ki (pi \<bullet> s1)"
-  "t2 \<approx>ty s2 \<Longrightarrow> (pi \<bullet> t2) \<approx>ty (pi \<bullet> s2)"
-  "t3 \<approx>tr s3 \<Longrightarrow> (pi \<bullet> t3) \<approx>tr (pi \<bullet> s3)"
-apply(induct rule: alpha_rkind_alpha_rty_alpha_rtrm.inducts)
-apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm.intros)
-apply (simp_all add: alpha_rkind_alpha_rty_alpha_rtrm_inj)
-apply (rule alpha_gen_atom_eqvt)
-apply (simp add: rfv_eqvt)
-apply assumption
-apply (rule alpha_gen_atom_eqvt)
-apply (simp add: rfv_eqvt)
-apply assumption
-apply (rule alpha_gen_atom_eqvt)
-apply (simp add: rfv_eqvt)
-apply assumption
-done
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha_equivps}, []),
-  (build_equivps [@{term alpha_rkind}, @{term alpha_rty}, @{term alpha_rtrm}]
-     @{thm rkind_rty_rtrm.induct} @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} 
-     @{thms rkind.inject rty.inject rtrm.inject} @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj}
-     @{thms rkind.distinct rty.distinct rtrm.distinct}
-     @{thms alpha_rkind.cases alpha_rty.cases alpha_rtrm.cases}
-     @{thms alpha_eqvt} ctxt)) ctxt)) *}
-thm alpha_equivps
-
-local_setup  {* define_quotient_type
-  [(([], @{binding kind}, NoSyn), (@{typ rkind}, @{term alpha_rkind})),
-   (([], @{binding ty},   NoSyn), (@{typ rty},   @{term alpha_rty}  )),
-   (([], @{binding trm},  NoSyn), (@{typ rtrm},  @{term alpha_rtrm} ))]
-  (ALLGOALS (resolve_tac @{thms alpha_equivps}))
-*}
-
-local_setup {*
-(fn ctxt => ctxt
- |> snd o (Quotient_Def.quotient_lift_const ("TYP", @{term Type}))
- |> snd o (Quotient_Def.quotient_lift_const ("KPI", @{term KPi}))
- |> snd o (Quotient_Def.quotient_lift_const ("TCONST", @{term TConst}))
- |> snd o (Quotient_Def.quotient_lift_const ("TAPP", @{term TApp}))
- |> snd o (Quotient_Def.quotient_lift_const ("TPI", @{term TPi}))
- |> snd o (Quotient_Def.quotient_lift_const ("CONS", @{term Const}))
- |> snd o (Quotient_Def.quotient_lift_const ("VAR", @{term Var}))
- |> snd o (Quotient_Def.quotient_lift_const ("APP", @{term App}))
- |> snd o (Quotient_Def.quotient_lift_const ("LAM", @{term Lam}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_kind", @{term fv_rkind}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_ty", @{term fv_rty}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_trm", @{term fv_rtrm}))) *}
-print_theorems
-
-local_setup {* prove_const_rsp @{binding rfv_rsp} [@{term fv_rkind}, @{term fv_rty}, @{term fv_rtrm}]
-  (fn _ => fvbv_rsp_tac @{thm alpha_rkind_alpha_rty_alpha_rtrm.induct} @{thms fv_rkind_fv_rty_fv_rtrm.simps} 1) *}
-local_setup {* prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"}]
-  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *}
-local_setup {* prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rty \<Rightarrow> rty"}]
-  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *}
-local_setup {* prove_const_rsp Binding.empty [@{term "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"}]
-  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha_eqvt}) 1) *}
-ML {* fun const_rsp_tac _ = constr_rsp_tac @{thms alpha_rkind_alpha_rty_alpha_rtrm_inj}
-  @{thms rfv_rsp} @{thms alpha_equivps} 1 *}
-local_setup {* prove_const_rsp Binding.empty [@{term TConst}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term TApp}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term Var}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term App}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term Const}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term KPi}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term TPi}] const_rsp_tac *}
-local_setup {* prove_const_rsp Binding.empty [@{term Lam}] const_rsp_tac *}
-
-lemmas kind_ty_trm_induct = rkind_rty_rtrm.induct[quot_lifted]
-
-thm rkind_rty_rtrm.inducts
-lemmas kind_ty_trm_inducts = rkind_rty_rtrm.inducts[quot_lifted]
-
-instantiation kind and ty and trm :: pt
-begin
-
-quotient_definition
-  "permute_kind :: perm \<Rightarrow> kind \<Rightarrow> kind"
-is
-  "permute :: perm \<Rightarrow> rkind \<Rightarrow> rkind"
-
-quotient_definition
-  "permute_ty :: perm \<Rightarrow> ty \<Rightarrow> ty"
-is
-  "permute :: perm \<Rightarrow> rty \<Rightarrow> rty"
-
-quotient_definition
-  "permute_trm :: perm \<Rightarrow> trm \<Rightarrow> trm"
-is
-  "permute :: perm \<Rightarrow> rtrm \<Rightarrow> rtrm"
-
-instance by default (simp_all add:
-  permute_rkind_permute_rty_permute_rtrm_zero[quot_lifted]
-  permute_rkind_permute_rty_permute_rtrm_append[quot_lifted])
-
-end
-
-(*
-Lifts, but slow and not needed?.
-lemmas alpha_kind_alpha_ty_alpha_trm_induct = alpha_rkind_alpha_rty_alpha_rtrm.induct[unfolded alpha_gen, quot_lifted, folded alpha_gen]
-*)
-
-lemmas permute_ktt[simp] = permute_rkind_permute_rty_permute_rtrm.simps[quot_lifted]
-
-lemmas kind_ty_trm_inj = alpha_rkind_alpha_rty_alpha_rtrm_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
-
-lemmas fv_kind_ty_trm = fv_rkind_fv_rty_fv_rtrm.simps[quot_lifted]
-
-lemmas fv_eqvt = rfv_eqvt[quot_lifted]
-
-lemma supports:
-  "{} supports TYP"
-  "(supp (atom i)) supports (TCONST i)"
-  "(supp A \<union> supp M) supports (TAPP A M)"
-  "(supp (atom i)) supports (CONS i)"
-  "(supp (atom x)) supports (VAR x)"
-  "(supp M \<union> supp N) supports (APP M N)"
-  "(supp ty \<union> supp (atom na) \<union> supp ki) supports (KPI ty na ki)"
-  "(supp ty \<union> supp (atom na) \<union> supp ty2) supports (TPI ty na ty2)"
-  "(supp ty \<union> supp (atom na) \<union> supp trm) supports (LAM ty na trm)"
-apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh)
-apply(rule_tac [!] allI)+
-apply(rule_tac [!] impI)
-apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
-apply(simp_all add: fresh_atom)
-done
-
-lemma kind_ty_trm_fs:
-  "finite (supp (x\<Colon>kind))"
-  "finite (supp (y\<Colon>ty))"
-  "finite (supp (z\<Colon>trm))"
-apply(induct x and y and z rule: kind_ty_trm_inducts)
-apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
-apply(simp_all add: supp_atom)
-done
-
-instance kind and ty and trm :: fs
-apply(default)
-apply(simp_all only: kind_ty_trm_fs)
-done
-
-lemma supp_eqs:
-  "supp TYP = {}"
-  "supp rkind = fv_kind rkind \<Longrightarrow> supp (KPI rty name rkind) = supp rty \<union> supp (Abs {atom name} rkind)"
-  "supp (TCONST i) = {atom i}"
-  "supp (TAPP A M) = supp A \<union> supp M"
-  "supp rty2 = fv_ty rty2 \<Longrightarrow> supp (TPI rty1 name rty2) = supp rty1 \<union> supp (Abs {atom name} rty2)"
-  "supp (CONS i) = {atom i}"
-  "supp (VAR x) = {atom x}"
-  "supp (APP M N) = supp M \<union> supp N"
-  "supp rtrm = fv_trm rtrm \<Longrightarrow> supp (LAM rty name rtrm) = supp rty \<union> supp (Abs {atom name} rtrm)"
-  apply(simp_all (no_asm) add: supp_def)
-  apply(simp_all only: kind_ty_trm_inj Abs_eq_iff alpha_gen)
-  apply(simp_all only: insert_eqvt empty_eqvt atom_eqvt supp_eqvt[symmetric] fv_eqvt[symmetric])
-  apply(simp_all add: Collect_imp_eq Collect_neg_eq[symmetric] Set.Un_commute)
-  apply(simp_all add: supp_at_base[simplified supp_def])
-  done
-
-lemma supp_fv:
-  "supp t1 = fv_kind t1"
-  "supp t2 = fv_ty t2"
-  "supp t3 = fv_trm t3"
-  apply(induct t1 and t2 and t3 rule: kind_ty_trm_inducts)
-  apply(simp_all (no_asm) only: supp_eqs fv_kind_ty_trm)
-  apply(simp_all)
-  apply(subst supp_eqs)
-  apply(simp_all add: supp_Abs)
-  apply(subst supp_eqs)
-  apply(simp_all add: supp_Abs)
-  apply(subst supp_eqs)
-  apply(simp_all add: supp_Abs)
-  done
-
-lemma supp_rkind_rty_rtrm:
-  "supp TYP = {}"
-  "supp (KPI A x K) = supp A \<union> (supp K - {atom x})"
-  "supp (TCONST i) = {atom i}"
-  "supp (TAPP A M) = supp A \<union> supp M"
-  "supp (TPI A x B) = supp A \<union> (supp B - {atom x})"
-  "supp (CONS i) = {atom i}"
-  "supp (VAR x) = {atom x}"
-  "supp (APP M N) = supp M \<union> supp N"
-  "supp (LAM A x M) = supp A \<union> (supp M - {atom x})"
-  by (simp_all only: supp_fv fv_kind_ty_trm)
-
-end
-
-
-
-
--- a/Quot/Nominal/LamEx.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,624 +0,0 @@
-theory LamEx
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
-begin
-
-atom_decl name
-
-datatype rlam =
-  rVar "name"
-| rApp "rlam" "rlam"
-| rLam "name" "rlam"
-
-fun
-  rfv :: "rlam \<Rightarrow> atom set"
-where
-  rfv_var: "rfv (rVar a) = {atom a}"
-| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
-| rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}"
-
-instantiation rlam :: pt
-begin
-
-primrec
-  permute_rlam
-where
-  "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)"
-| "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)"
-| "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)"
-
-instance
-apply default
-apply(induct_tac [!] x)
-apply(simp_all)
-done
-
-end
-
-instantiation rlam :: fs
-begin
-
-lemma neg_conj:
-  "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)"
-  by simp
-
-lemma infinite_Un:
-  "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
-  by simp
-
-instance
-apply default
-apply(induct_tac x)
-(* var case *)
-apply(simp add: supp_def)
-apply(fold supp_def)[1]
-apply(simp add: supp_at_base)
-(* app case *)
-apply(simp only: supp_def)
-apply(simp only: permute_rlam.simps) 
-apply(simp only: rlam.inject) 
-apply(simp only: neg_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: infinite_Un)
-apply(simp only: Collect_disj_eq)
-apply(simp)
-(* lam case *)
-apply(simp only: supp_def)
-apply(simp only: permute_rlam.simps) 
-apply(simp only: rlam.inject) 
-apply(simp only: neg_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: infinite_Un)
-apply(simp only: Collect_disj_eq)
-apply(simp)
-apply(fold supp_def)[1]
-apply(simp add: supp_at_base)
-done
-
-end
-
-
-(* for the eqvt proof of the alpha-equivalence *)
-declare permute_rlam.simps[eqvt]
-
-lemma rfv_eqvt[eqvt]:
-  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
-apply(induct t)
-apply(simp_all)
-apply(simp add: permute_set_eq atom_eqvt)
-apply(simp add: union_eqvt)
-apply(simp add: Diff_eqvt)
-apply(simp add: permute_set_eq atom_eqvt)
-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. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s)
-       \<Longrightarrow> rLam a t \<approx> rLam b s"
-
-lemma a3_inverse:
-  assumes "rLam a t \<approx> rLam b s"
-  shows "\<exists>pi. (rfv t - {atom a} = rfv s - {atom b} \<and> (rfv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s)"
-using assms
-apply(erule_tac alpha.cases)
-apply(auto)
-done
-
-text {* should be automatic with new version of eqvt-machinery *}
-lemma alpha_eqvt:
-  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 ?p1="- pi" in permute_eq_iff[THEN iffD1])
-apply(simp only: Diff_eqvt rfv_eqvt insert_eqvt atom_eqvt empty_eqvt)
-apply(simp)
-apply(rule conjI)
-apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
-apply(simp add: Diff_eqvt rfv_eqvt atom_eqvt insert_eqvt empty_eqvt)
-apply(subst permute_eqvt[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="0" in exI)
-apply(simp_all add: fresh_star_def fresh_zero_perm)
-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="- pi" in exI)
-apply(simp)
-apply(simp add: fresh_star_def fresh_minus_perm)
-apply(erule conjE)+
-apply(rotate_tac 3)
-apply(drule_tac pi="- pi" in alpha_eqvt)
-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_plus)
-apply(drule_tac x="- pia \<bullet> sa" in spec)
-apply(drule mp)
-apply(rotate_tac 7)
-apply(drule_tac pi="- pia" in alpha_eqvt)
-apply(simp)
-apply(rotate_tac 9)
-apply(drule_tac pi="pia" in alpha_eqvt)
-apply(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_all)
-  done
-
-inductive
-    alpha2 :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx>2 _" [100, 100] 100)
-where
-  a21: "a = b \<Longrightarrow> (rVar a) \<approx>2 (rVar b)"
-| a22: "\<lbrakk>t1 \<approx>2 t2; s1 \<approx>2 s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx>2 rApp t2 s2"
-| a23: "(a = b \<and> t \<approx>2 s) \<or> (a \<noteq> b \<and> ((a \<leftrightarrow> b) \<bullet> t) \<approx>2 s \<and> atom b \<notin> rfv t)\<Longrightarrow> rLam a t \<approx>2 rLam b s"
-
-lemma fv_vars:
-  fixes a::name
-  assumes a1: "\<forall>x \<in> rfv t - {atom a}. pi \<bullet> x = x"
-  shows "(pi \<bullet> t) \<approx>2 ((a \<leftrightarrow> (pi \<bullet> a)) \<bullet> t)"
-using a1
-apply(induct t)
-apply(auto)
-apply(rule a21)
-apply(case_tac "name = a")
-apply(simp)
-apply(simp)
-defer
-apply(rule a22)
-apply(simp)
-apply(simp)
-apply(rule a23)
-apply(case_tac "a = name")
-apply(simp)
-oops
-
-
-lemma 
-  assumes a1: "t \<approx>2 s"
-  shows "t \<approx> s"
-using a1
-apply(induct)
-apply(rule alpha.intros)
-apply(simp)
-apply(rule alpha.intros)
-apply(simp)
-apply(simp)
-apply(rule alpha.intros)
-apply(erule disjE)
-apply(rule_tac x="0" in exI)
-apply(simp add: fresh_star_def fresh_zero_perm)
-apply(erule conjE)+
-apply(drule alpha_rfv)
-apply(simp)
-apply(rule_tac x="(a \<leftrightarrow> b)" in exI)
-apply(simp)
-apply(erule conjE)+
-apply(rule conjI)
-apply(drule alpha_rfv)
-apply(drule sym)
-apply(simp)
-apply(simp add: rfv_eqvt[symmetric])
-defer
-apply(subgoal_tac "atom a \<sharp> (rfv t - {atom a})")
-apply(subgoal_tac "atom b \<sharp> (rfv t - {atom a})")
-
-defer
-sorry
-
-lemma 
-  assumes a1: "t \<approx> s"
-  shows "t \<approx>2 s"
-using a1
-apply(induct)
-apply(rule alpha2.intros)
-apply(simp)
-apply(rule alpha2.intros)
-apply(simp)
-apply(simp)
-apply(clarify)
-apply(rule alpha2.intros)
-apply(frule alpha_rfv)
-apply(rotate_tac 4)
-apply(drule sym)
-apply(simp)
-apply(drule sym)
-apply(simp)
-oops
-
-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> atom set"
-is
-  "rfv"
-
-lemma perm_rsp[quot_respect]:
-  "(op = ===> alpha ===> alpha) permute permute"
-  apply(auto)
-  apply(rule alpha_eqvt)
-  apply(simp)
-  done
-
-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="0" in exI)
-  unfolding fresh_star_def 
-  apply(simp add: fresh_star_def fresh_zero_perm)
-  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"
-  apply (lifting rlam.induct)
-  done
-
-instantiation lam :: pt
-begin
-
-quotient_definition
-  "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam"
-is
-  "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam"
-
-lemma permute_lam [simp]:
-  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 permute_rlam.simps)
-done
-
-instance
-apply default
-apply(induct_tac [!] x rule: lam_induct)
-apply(simp_all)
-done
-
-end
-
-lemma fv_lam [simp]: 
-  shows "fv (Var a) = {atom a}"
-  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
-  and   "fv (Lam a t) = fv t - {atom a}"
-apply(lifting rfv_var rfv_app rfv_lam)
-done
-
-lemma fv_eqvt:
-  shows "(p \<bullet> fv t) = fv (p \<bullet> t)"
-apply(lifting rfv_eqvt)
-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. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s)\<rbrakk> 
-   \<Longrightarrow> Lam a t = Lam b s"
-  apply  (lifting a3)
-  done
-
-lemma a3_inv:
-  assumes "Lam a t = Lam b s"
-  shows "\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s)"
-using assms
-apply(lifting a3_inverse)
-done
-
-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. fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a}) \<sharp>* pi \<and> (pi \<bullet> t) = s\<rbrakk> 
-   \<Longrightarrow> P\<rbrakk>
-    \<Longrightarrow> P"
-  by (lifting alpha.cases)
-
-(* not sure whether needed *)
-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. fv t - {atom a} = fv s - {atom b} \<and>
-         (fv t - {atom a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s)\<rbrakk> 
-     \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
-    \<Longrightarrow> qxb qx qxa"
-  by (lifting alpha.induct)
-
-(* should they lift automatically *)
-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(regularize)
-prefer 2
-apply(regularize)
-prefer 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 Lam_pseudo_inject:
-  shows "(Lam a t = Lam b s) = 
-      (\<exists>pi. (fv t - {atom a} = fv s - {atom b} \<and> (fv t - {atom a})\<sharp>* pi \<and> (pi \<bullet> t) = s))"
-apply(rule iffI)
-apply(rule a3_inv)
-apply(assumption)
-apply(rule a3)
-apply(assumption)
-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 only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-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)"
-  apply (simp add: supp_def)
-  done
-
-lemma var_supp:
-  shows "(supp (Var a)) = {a:::name}"
-  using var_supp1 by (simp add: supp_at_base)
-
-lemma app_supp:
-  shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)"
-apply(simp only: supp_def lam_inject)
-apply(simp add: Collect_imp_eq Collect_neg_eq)
-done
-
-(* supp for lam *)
-lemma lam_supp1:
-  shows "(supp (atom x, t)) supports (Lam x t) "
-apply(simp add: supports_def)
-apply(fold fresh_def)
-apply(simp add: fresh_Pair swap_fresh_fresh)
-apply(clarify)
-apply(subst swap_at_base_simps(3))
-apply(simp_all add: fresh_atom)
-done
-
-lemma lam_fsupp1:
-  assumes a: "finite (supp t)"
-  shows "finite (supp (Lam x t))"
-apply(rule supports_finite)
-apply(rule lam_supp1)
-apply(simp add: a supp_Pair supp_atom)
-done
-
-instance lam :: fs
-apply(default)
-apply(induct_tac x rule: lam_induct)
-apply(simp add: var_supp)
-apply(simp add: app_supp)
-apply(simp add: lam_fsupp1)
-done
-
-lemma supp_fv:
-  shows "supp t = fv t"
-apply(induct t rule: lam_induct)
-apply(simp add: var_supp)
-apply(simp add: app_supp)
-apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)")
-apply(simp add: supp_Abs)
-apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
-apply(simp add: Lam_pseudo_inject)
-apply(simp add: Abs_eq_iff alpha_gen)
-apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric])
-done
-
-lemma lam_supp2:
-  shows "supp (Lam x t) = supp (Abs {atom x} t)"
-apply(simp add: supp_def permute_set_eq atom_eqvt)
-apply(simp add: Lam_pseudo_inject)
-apply(simp add: Abs_eq_iff supp_fv alpha_gen)
-done
-
-lemma lam_supp:
-  shows "supp (Lam x t) = ((supp t) - {atom x})"
-apply(simp add: lam_supp2)
-apply(simp add: supp_Abs)
-done
-
-lemma fresh_lam:
-  "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)"
-apply(simp add: fresh_def)
-apply(simp add: lam_supp)
-apply(auto)
-done
-
-lemma lam_induct_strong:
-  fixes a::"'a::fs"
-  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; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
-  shows "P a lam"
-proof -
-  have "\<And>pi 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 a. P a (pi \<bullet> lam1)" by fact
-    have b2: "\<And>pi 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 a. P a (pi \<bullet> lam)" by fact
-    obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
-      apply(rule obtain_atom)
-      apply(auto)
-      sorry
-    from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" 
-      apply -
-      apply(rule a3)
-      apply(blast)
-      apply(simp add: fresh_Pair)
-      done
-    have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
-      apply(rule swap_fresh_fresh)
-      using fr
-      apply(simp add: fresh_lam fresh_Pair)
-      apply(simp add: fresh_lam fresh_Pair)
-      done
-    show "P a (pi \<bullet> Lam name lam)" 
-      apply (simp)
-      apply(subst eq[symmetric])
-      using p
-      apply(simp only: permute_lam)
-      apply(simp add: flip_def)
-      done
-  qed
-  then have "P a (0 \<bullet> lam)" by blast
-  then show "P a lam" by simp 
-qed
-
-
-lemma var_fresh:
-  fixes a::"name"
-  shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)"
-  apply(simp add: fresh_def)
-  apply(simp add: var_supp1)
-  done
-
-
-
-end
-
--- a/Quot/Nominal/LamEx2.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,567 +0,0 @@
-theory LamEx
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs"
-begin
-
-atom_decl name
-
-datatype rlam =
-  rVar "name"
-| rApp "rlam" "rlam"
-| rLam "name" "rlam"
-
-fun
-  rfv :: "rlam \<Rightarrow> atom set"
-where
-  rfv_var: "rfv (rVar a) = {atom a}"
-| rfv_app: "rfv (rApp t1 t2) = (rfv t1) \<union> (rfv t2)"
-| rfv_lam: "rfv (rLam a t) = (rfv t) - {atom a}"
-
-instantiation rlam :: pt
-begin
-
-primrec
-  permute_rlam
-where
-  "permute_rlam pi (rVar a) = rVar (pi \<bullet> a)"
-| "permute_rlam pi (rApp t1 t2) = rApp (permute_rlam pi t1) (permute_rlam pi t2)"
-| "permute_rlam pi (rLam a t) = rLam (pi \<bullet> a) (permute_rlam pi t)"
-
-instance
-apply default
-apply(induct_tac [!] x)
-apply(simp_all)
-done
-
-end
-
-instantiation rlam :: fs
-begin
-
-lemma neg_conj:
-  "\<not>(P \<and> Q) \<longleftrightarrow> (\<not>P) \<or> (\<not>Q)"
-  by simp
-
-lemma infinite_Un:
-  "infinite (S \<union> T) \<longleftrightarrow> infinite S \<or> infinite T"
-  by simp
-
-instance
-apply default
-apply(induct_tac x)
-(* var case *)
-apply(simp add: supp_def)
-apply(fold supp_def)[1]
-apply(simp add: supp_at_base)
-(* app case *)
-apply(simp only: supp_def)
-apply(simp only: permute_rlam.simps) 
-apply(simp only: rlam.inject) 
-apply(simp only: neg_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: infinite_Un)
-apply(simp only: Collect_disj_eq)
-apply(simp)
-(* lam case *)
-apply(simp only: supp_def)
-apply(simp only: permute_rlam.simps) 
-apply(simp only: rlam.inject) 
-apply(simp only: neg_conj)
-apply(simp only: Collect_disj_eq)
-apply(simp only: infinite_Un)
-apply(simp only: Collect_disj_eq)
-apply(simp)
-apply(fold supp_def)[1]
-apply(simp add: supp_at_base)
-done
-
-end
-
-
-(* for the eqvt proof of the alpha-equivalence *)
-declare permute_rlam.simps[eqvt]
-
-lemma rfv_eqvt[eqvt]:
-  shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)"
-apply(induct t)
-apply(simp_all)
-apply(simp add: permute_set_eq atom_eqvt)
-apply(simp add: union_eqvt)
-apply(simp add: Diff_eqvt)
-apply(simp add: permute_set_eq atom_eqvt)
-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. (({atom a}, t) \<approx>gen alpha rfv pi ({atom b}, s)) \<Longrightarrow> rLam a t \<approx> rLam b s"
-print_theorems
-thm alpha.induct
-
-lemma a3_inverse:
-  assumes "rLam a t \<approx> rLam b s"
-  shows "\<exists>pi. (({atom a}, t) \<approx>gen alpha rfv pi ({atom b}, s))"
-using assms
-apply(erule_tac alpha.cases)
-apply(auto)
-done
-
-text {* should be automatic with new version of eqvt-machinery *}
-lemma alpha_eqvt:
-  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(rule alpha_gen_atom_eqvt)
-apply(rule rfv_eqvt)
-apply assumption
-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="0" in exI)
-apply(rule alpha_gen_refl)
-apply(assumption)
-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 alpha_gen_compose_sym)
-  apply(erule alpha_eqvt)
-  done
-
-lemma alpha_trans:
-  shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3"
-apply(induct arbitrary: t3 rule: alpha.induct)
-apply(simp add: a1)
-apply(rotate_tac 4)
-apply(erule alpha.cases)
-apply(simp_all add: a2)
-apply(erule alpha.cases)
-apply(simp_all)
-apply(rule a3)
-apply(erule alpha_gen_compose_trans)
-apply(assumption)
-apply(erule alpha_eqvt)
-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_all add: alpha_gen.simps)
-  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> atom set"
-is
-  "rfv"
-
-lemma perm_rsp[quot_respect]:
-  "(op = ===> alpha ===> alpha) permute permute"
-  apply(auto)
-  apply(rule alpha_eqvt)
-  apply(simp)
-  done
-
-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="0" in exI)
-  unfolding fresh_star_def 
-  apply(simp add: fresh_star_def fresh_zero_perm alpha_gen.simps)
-  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"
-  apply (lifting rlam.induct)
-  done
-
-instantiation lam :: pt
-begin
-
-quotient_definition
-  "permute_lam :: perm \<Rightarrow> lam \<Rightarrow> lam"
-is
-  "permute :: perm \<Rightarrow> rlam \<Rightarrow> rlam"
-
-lemma permute_lam [simp]:
-  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 permute_rlam.simps)
-done
-
-instance
-apply default
-apply(induct_tac [!] x rule: lam_induct)
-apply(simp_all)
-done
-
-end
-
-lemma fv_lam [simp]: 
-  shows "fv (Var a) = {atom a}"
-  and   "fv (App t1 t2) = fv t1 \<union> fv t2"
-  and   "fv (Lam a t) = fv t - {atom a}"
-apply(lifting rfv_var rfv_app rfv_lam)
-done
-
-lemma fv_eqvt:
-  shows "(p \<bullet> fv t) = fv (p \<bullet> t)"
-apply(lifting rfv_eqvt)
-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 alpha_gen_rsp_pre:
-  assumes a5: "\<And>t s. R t s \<Longrightarrow> R (pi \<bullet> t) (pi \<bullet> s)"
-  and     a1: "R s1 t1"
-  and     a2: "R s2 t2"
-  and     a3: "\<And>a b c d. R a b \<Longrightarrow> R c d \<Longrightarrow> R1 a c = R2 b d"
-  and     a4: "\<And>x y. R x y \<Longrightarrow> fv1 x = fv2 y"
-  shows   "(a, s1) \<approx>gen R1 fv1 pi (b, s2) = (a, t1) \<approx>gen R2 fv2 pi (b, t2)"
-apply (simp add: alpha_gen.simps)
-apply (simp only: a4[symmetric, OF a1] a4[symmetric, OF a2])
-apply auto
-apply (subst a3[symmetric])
-apply (rule a5)
-apply (rule a1)
-apply (rule a2)
-apply (assumption)
-apply (subst a3)
-apply (rule a5)
-apply (rule a1)
-apply (rule a2)
-apply (assumption)
-done
-
-lemma [quot_respect]: "(prod_rel op = alpha ===>
-           (alpha ===> alpha ===> op =) ===> (alpha ===> op =) ===> op = ===> prod_rel op = alpha ===> op =)
-           alpha_gen alpha_gen"
-apply simp
-apply clarify
-apply (rule alpha_gen_rsp_pre[of "alpha",OF alpha_eqvt])
-apply auto
-done
-
-(* pi_abs would be also sufficient to prove the next lemma *)
-lemma replam_eqvt: "pi \<bullet> (rep_lam x) = rep_lam (pi \<bullet> x)"
-apply (unfold rep_lam_def)
-sorry
-
-lemma [quot_preserve]: "(prod_fun id rep_lam --->
-           (abs_lam ---> abs_lam ---> id) ---> (abs_lam ---> id) ---> id ---> (prod_fun id rep_lam) ---> id)
-           alpha_gen = alpha_gen"
-apply (simp add: expand_fun_eq alpha_gen.simps Quotient_abs_rep[OF Quotient_lam])
-apply (simp add: replam_eqvt)
-apply (simp only: Quotient_abs_rep[OF Quotient_lam])
-apply auto
-done
-
-
-lemma alpha_prs [quot_preserve]: "(rep_lam ---> rep_lam ---> id) alpha = (op =)"
-apply (simp add: expand_fun_eq)
-apply (simp add: Quotient_rel_rep[OF Quotient_lam])
-done
-
-lemma a3:
-  "\<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s) \<Longrightarrow> Lam a t = Lam b s"
-  apply (unfold alpha_gen)
-  apply (lifting a3[unfolded alpha_gen])
-  done
-
-
-lemma a3_inv:
-  "Lam a t = Lam b s \<Longrightarrow> \<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s)"
-  apply (unfold alpha_gen)
-  apply (lifting a3_inverse[unfolded alpha_gen])
-  done
-
-lemma alpha_cases: 
-  "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P;
-    \<And>t1 t2 s1 s2. \<lbrakk>a1 = App t1 s1; a2 = App t2 s2; t1 = t2; s1 = s2\<rbrakk> \<Longrightarrow> P;
-    \<And>a t b s. \<lbrakk>a1 = Lam a t; a2 = Lam b s; \<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s)\<rbrakk>
-   \<Longrightarrow> P\<rbrakk>
-    \<Longrightarrow> P"
-unfolding alpha_gen
-apply (lifting alpha.cases[unfolded alpha_gen])
-done
-
-(* not sure whether needed *)
-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>a t b s. \<exists>pi. ({atom a}, t) \<approx>gen (\<lambda>x1 x2. x1 = x2 \<and> qxb x1 x2) fv pi ({atom b}, s) \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk>
-    \<Longrightarrow> qxb qx qxa"
-unfolding alpha_gen by (lifting alpha.induct[unfolded alpha_gen])
-
-(* should they lift automatically *)
-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(regularize)
-prefer 2
-apply(regularize)
-prefer 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
-
-thm a3_inv
-lemma Lam_pseudo_inject:
-  shows "(Lam a t = Lam b s) = (\<exists>pi. ({atom a}, t) \<approx>gen (op =) fv pi ({atom b}, s))"
-apply(rule iffI)
-apply(rule a3_inv)
-apply(assumption)
-apply(rule a3)
-apply(assumption)
-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 only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-apply (erule alpha.cases)
-apply (simp_all only: rlam.distinct)
-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)"
-  apply (simp add: supp_def)
-  done
-
-lemma var_supp:
-  shows "(supp (Var a)) = {a:::name}"
-  using var_supp1 by (simp add: supp_at_base)
-
-lemma app_supp:
-  shows "supp (App t1 t2) = (supp t1) \<union> (supp t2)"
-apply(simp only: supp_def lam_inject)
-apply(simp add: Collect_imp_eq Collect_neg_eq)
-done
-
-(* supp for lam *)
-lemma lam_supp1:
-  shows "(supp (atom x, t)) supports (Lam x t) "
-apply(simp add: supports_def)
-apply(fold fresh_def)
-apply(simp add: fresh_Pair swap_fresh_fresh)
-apply(clarify)
-apply(subst swap_at_base_simps(3))
-apply(simp_all add: fresh_atom)
-done
-
-lemma lam_fsupp1:
-  assumes a: "finite (supp t)"
-  shows "finite (supp (Lam x t))"
-apply(rule supports_finite)
-apply(rule lam_supp1)
-apply(simp add: a supp_Pair supp_atom)
-done
-
-instance lam :: fs
-apply(default)
-apply(induct_tac x rule: lam_induct)
-apply(simp add: var_supp)
-apply(simp add: app_supp)
-apply(simp add: lam_fsupp1)
-done
-
-lemma supp_fv:
-  shows "supp t = fv t"
-apply(induct t rule: lam_induct)
-apply(simp add: var_supp)
-apply(simp add: app_supp)
-apply(subgoal_tac "supp (Lam name lam) = supp (Abs {atom name} lam)")
-apply(simp add: supp_Abs)
-apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
-apply(simp add: Lam_pseudo_inject)
-apply(simp add: Abs_eq_iff)
-apply(simp add: alpha_gen.simps)
-apply(simp add: supp_eqvt[symmetric] fv_eqvt[symmetric])
-done
-
-lemma lam_supp2:
-  shows "supp (Lam x t) = supp (Abs {atom x} t)"
-apply(simp add: supp_def permute_set_eq atom_eqvt)
-apply(simp add: Lam_pseudo_inject)
-apply(simp add: Abs_eq_iff)
-apply(simp add: alpha_gen  supp_fv)
-done
-
-lemma lam_supp:
-  shows "supp (Lam x t) = ((supp t) - {atom x})"
-apply(simp add: lam_supp2)
-apply(simp add: supp_Abs)
-done
-
-lemma fresh_lam:
-  "(atom a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> atom a \<sharp> t)"
-apply(simp add: fresh_def)
-apply(simp add: lam_supp)
-apply(auto)
-done
-
-lemma lam_induct_strong:
-  fixes a::"'a::fs"
-  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; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lam name lam)"
-  shows "P a lam"
-proof -
-  have "\<And>pi 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 a. P a (pi \<bullet> lam1)" by fact
-    have b2: "\<And>pi 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 a. P a (pi \<bullet> lam)" by fact
-    obtain c::name where fr: "atom c\<sharp>(a, pi\<bullet>name, pi\<bullet>lam)"
-      apply(rule obtain_atom)
-      apply(auto)
-      sorry
-    from b fr have p: "P a (Lam c (((c \<leftrightarrow> (pi \<bullet> name)) + pi)\<bullet>lam))" 
-      apply -
-      apply(rule a3)
-      apply(blast)
-      apply(simp add: fresh_Pair)
-      done
-    have eq: "(atom c \<rightleftharpoons> atom (pi\<bullet>name)) \<bullet> Lam (pi \<bullet> name) (pi \<bullet> lam) = Lam (pi \<bullet> name) (pi \<bullet> lam)"
-      apply(rule swap_fresh_fresh)
-      using fr
-      apply(simp add: fresh_lam fresh_Pair)
-      apply(simp add: fresh_lam fresh_Pair)
-      done
-    show "P a (pi \<bullet> Lam name lam)" 
-      apply (simp)
-      apply(subst eq[symmetric])
-      using p
-      apply(simp only: permute_lam)
-      apply(simp add: flip_def)
-      done
-  qed
-  then have "P a (0 \<bullet> lam)" by blast
-  then show "P a lam" by simp 
-qed
-
-
-lemma var_fresh:
-  fixes a::"name"
-  shows "(atom a \<sharp> (Var b)) = (atom a \<sharp> b)"
-  apply(simp add: fresh_def)
-  apply(simp add: var_supp1)
-  done
-
-
-
-end
-
--- a/Quot/Nominal/Nominal2_Atoms.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,251 +0,0 @@
-(*  Title:      Nominal2_Atoms
-    Authors:    Brian Huffman, Christian Urban
-
-    Definitions for concrete atom types. 
-*)
-theory Nominal2_Atoms
-imports Nominal2_Base
-uses ("nominal_atoms.ML")
-begin
-
-section {* Concrete atom types *}
-
-text {*
-  Class @{text at_base} allows types containing multiple sorts of atoms.
-  Class @{text at} only allows types with a single sort.
-*}
-
-class at_base = pt +
-  fixes atom :: "'a \<Rightarrow> atom"
-  assumes atom_eq_iff [simp]: "atom a = atom b \<longleftrightarrow> a = b"
-  assumes atom_eqvt: "p \<bullet> (atom a) = atom (p \<bullet> a)"
-
-class at = at_base +
-  assumes sort_of_atom_eq [simp]: "sort_of (atom a) = sort_of (atom b)"
-
-instance at < at_base ..
-
-lemma supp_at_base: 
-  fixes a::"'a::at_base"
-  shows "supp a = {atom a}"
-  by (simp add: supp_atom [symmetric] supp_def atom_eqvt)
-
-lemma fresh_at_base: 
-  shows "a \<sharp> b \<longleftrightarrow> a \<noteq> atom b"
-  unfolding fresh_def by (simp add: supp_at_base)
-
-instance at_base < fs
-proof qed (simp add: supp_at_base)
-
-lemma at_base_infinite [simp]:
-  shows "infinite (UNIV :: 'a::at_base set)" (is "infinite ?U")
-proof
-  obtain a :: 'a where "True" by auto
-  assume "finite ?U"
-  hence "finite (atom ` ?U)"
-    by (rule finite_imageI)
-  then obtain b where b: "b \<notin> atom ` ?U" "sort_of b = sort_of (atom a)"
-    by (rule obtain_atom)
-  from b(2) have "b = atom ((atom a \<rightleftharpoons> b) \<bullet> a)"
-    unfolding atom_eqvt [symmetric]
-    by (simp add: swap_atom)
-  hence "b \<in> atom ` ?U" by simp
-  with b(1) show "False" by simp
-qed
-
-lemma swap_at_base_simps [simp]:
-  fixes x y::"'a::at_base"
-  shows "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> x = y"
-  and   "sort_of (atom x) = sort_of (atom y) \<Longrightarrow> (atom x \<rightleftharpoons> atom y) \<bullet> y = x"
-  and   "atom x \<noteq> a \<Longrightarrow> atom x \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x"
-  unfolding atom_eq_iff [symmetric]
-  unfolding atom_eqvt [symmetric]
-  by simp_all
-
-lemma obtain_at_base:
-  assumes X: "finite X"
-  obtains a::"'a::at_base" where "atom a \<notin> X"
-proof -
-  have "inj (atom :: 'a \<Rightarrow> atom)"
-    by (simp add: inj_on_def)
-  with X have "finite (atom -` X :: 'a set)"
-    by (rule finite_vimageI)
-  with at_base_infinite have "atom -` X \<noteq> (UNIV :: 'a set)"
-    by auto
-  then obtain a :: 'a where "atom a \<notin> X"
-    by auto
-  thus ?thesis ..
-qed
-
-
-section {* A swapping operation for concrete atoms *}
-  
-definition
-  flip :: "'a::at_base \<Rightarrow> 'a \<Rightarrow> perm" ("'(_ \<leftrightarrow> _')")
-where
-  "(a \<leftrightarrow> b) = (atom a \<rightleftharpoons> atom b)"
-
-lemma flip_self [simp]: "(a \<leftrightarrow> a) = 0"
-  unfolding flip_def by (rule swap_self)
-
-lemma flip_commute: "(a \<leftrightarrow> b) = (b \<leftrightarrow> a)"
-  unfolding flip_def by (rule swap_commute)
-
-lemma minus_flip [simp]: "- (a \<leftrightarrow> b) = (a \<leftrightarrow> b)"
-  unfolding flip_def by (rule minus_swap)
-
-lemma add_flip_cancel: "(a \<leftrightarrow> b) + (a \<leftrightarrow> b) = 0"
-  unfolding flip_def by (rule swap_cancel)
-
-lemma permute_flip_cancel [simp]: "(a \<leftrightarrow> b) \<bullet> (a \<leftrightarrow> b) \<bullet> x = x"
-  unfolding permute_plus [symmetric] add_flip_cancel by simp
-
-lemma permute_flip_cancel2 [simp]: "(a \<leftrightarrow> b) \<bullet> (b \<leftrightarrow> a) \<bullet> x = x"
-  by (simp add: flip_commute)
-
-lemma flip_eqvt: 
-  fixes a b c::"'a::at_base"
-  shows "p \<bullet> (a \<leftrightarrow> b) = (p \<bullet> a \<leftrightarrow> p \<bullet> b)"
-  unfolding flip_def
-  by (simp add: swap_eqvt atom_eqvt)
-
-lemma flip_at_base_simps [simp]:
-  shows "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> a = b"
-  and   "sort_of (atom a) = sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> b = a"
-  and   "\<lbrakk>a \<noteq> c; b \<noteq> c\<rbrakk> \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> c = c"
-  and   "sort_of (atom a) \<noteq> sort_of (atom b) \<Longrightarrow> (a \<leftrightarrow> b) \<bullet> x = x"
-  unfolding flip_def
-  unfolding atom_eq_iff [symmetric]
-  unfolding atom_eqvt [symmetric]
-  by simp_all
-
-text {* the following two lemmas do not hold for at_base, 
-  only for single sort atoms from at *}
-
-lemma permute_flip_at:
-  fixes a b c::"'a::at"
-  shows "(a \<leftrightarrow> b) \<bullet> c = (if c = a then b else if c = b then a else c)"
-  unfolding flip_def
-  apply (rule atom_eq_iff [THEN iffD1])
-  apply (subst atom_eqvt [symmetric])
-  apply (simp add: swap_atom)
-  done
-
-lemma flip_at_simps [simp]:
-  fixes a b::"'a::at"
-  shows "(a \<leftrightarrow> b) \<bullet> a = b" 
-  and   "(a \<leftrightarrow> b) \<bullet> b = a"
-  unfolding permute_flip_at by simp_all
-
-
-subsection {* Syntax for coercing at-elements to the atom-type *}
-
-syntax
-  "_atom_constrain" :: "logic \<Rightarrow> type \<Rightarrow> logic" ("_:::_" [4, 0] 3)
-
-translations
-  "_atom_constrain a t" => "atom (_constrain a t)"
-
-
-subsection {* A lemma for proving instances of class @{text at}. *}
-
-setup {* Sign.add_const_constraint (@{const_name "permute"}, NONE) *}
-setup {* Sign.add_const_constraint (@{const_name "atom"}, NONE) *}
-
-text {*
-  New atom types are defined as subtypes of @{typ atom}.
-*}
-
-lemma exists_eq_simple_sort: 
-  shows "\<exists>a. a \<in> {a. sort_of a = s}"
-  by (rule_tac x="Atom s 0" in exI, simp)
-
-lemma exists_eq_sort: 
-  shows "\<exists>a. a \<in> {a. sort_of a \<in> range sort_fun}"
-  by (rule_tac x="Atom (sort_fun x) y" in exI, simp)
-
-lemma at_base_class:
-  fixes sort_fun :: "'b \<Rightarrow>atom_sort"
-  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
-  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range sort_fun}"
-  assumes atom_def: "\<And>a. atom a = Rep a"
-  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
-  shows "OFCLASS('a, at_base_class)"
-proof
-  interpret type_definition Rep Abs "{a. sort_of a \<in> range sort_fun}" by (rule type)
-  have sort_of_Rep: "\<And>a. sort_of (Rep a) \<in> range sort_fun" using Rep by simp
-  fix a b :: 'a and p p1 p2 :: perm
-  show "0 \<bullet> a = a"
-    unfolding permute_def by (simp add: Rep_inverse)
-  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
-    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
-  show "atom a = atom b \<longleftrightarrow> a = b"
-    unfolding atom_def by (simp add: Rep_inject)
-  show "p \<bullet> atom a = atom (p \<bullet> a)"
-    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
-qed
-
-(*
-lemma at_class:
-  fixes s :: atom_sort
-  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
-  assumes type: "type_definition Rep Abs {a. sort_of a \<in> range (\<lambda>x::unit. s)}"
-  assumes atom_def: "\<And>a. atom a = Rep a"
-  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
-  shows "OFCLASS('a, at_class)"
-proof
-  interpret type_definition Rep Abs "{a. sort_of a \<in> range (\<lambda>x::unit. s)}" by (rule type)
-  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
-  fix a b :: 'a and p p1 p2 :: perm
-  show "0 \<bullet> a = a"
-    unfolding permute_def by (simp add: Rep_inverse)
-  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
-    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
-  show "sort_of (atom a) = sort_of (atom b)"
-    unfolding atom_def by (simp add: sort_of_Rep)
-  show "atom a = atom b \<longleftrightarrow> a = b"
-    unfolding atom_def by (simp add: Rep_inject)
-  show "p \<bullet> atom a = atom (p \<bullet> a)"
-    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
-qed
-*)
-
-lemma at_class:
-  fixes s :: atom_sort
-  fixes Rep :: "'a \<Rightarrow> atom" and Abs :: "atom \<Rightarrow> 'a"
-  assumes type: "type_definition Rep Abs {a. sort_of a = s}"
-  assumes atom_def: "\<And>a. atom a = Rep a"
-  assumes permute_def: "\<And>p a. p \<bullet> a = Abs (p \<bullet> Rep a)"
-  shows "OFCLASS('a, at_class)"
-proof
-  interpret type_definition Rep Abs "{a. sort_of a = s}" by (rule type)
-  have sort_of_Rep: "\<And>a. sort_of (Rep a) = s" using Rep by (simp add: image_def)
-  fix a b :: 'a and p p1 p2 :: perm
-  show "0 \<bullet> a = a"
-    unfolding permute_def by (simp add: Rep_inverse)
-  show "(p1 + p2) \<bullet> a = p1 \<bullet> p2 \<bullet> a"
-    unfolding permute_def by (simp add: Abs_inverse sort_of_Rep)
-  show "sort_of (atom a) = sort_of (atom b)"
-    unfolding atom_def by (simp add: sort_of_Rep)
-  show "atom a = atom b \<longleftrightarrow> a = b"
-    unfolding atom_def by (simp add: Rep_inject)
-  show "p \<bullet> atom a = atom (p \<bullet> a)"
-    unfolding permute_def atom_def by (simp add: Abs_inverse sort_of_Rep)
-qed
-
-setup {* Sign.add_const_constraint
-  (@{const_name "permute"}, SOME @{typ "perm \<Rightarrow> 'a::pt \<Rightarrow> 'a"}) *}
-setup {* Sign.add_const_constraint
-  (@{const_name "atom"}, SOME @{typ "'a::at_base \<Rightarrow> atom"}) *}
-
-
-section {* Automation for creating concrete atom types *}
-
-text {* at the moment only single-sort concrete atoms are supported *}
-
-use "nominal_atoms.ML"
-
-
-
-
-end
--- a/Quot/Nominal/Nominal2_Base.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1022 +0,0 @@
-(*  Title:      Nominal2_Base
-    Authors:    Brian Huffman, Christian Urban
-
-    Basic definitions and lemma infrastructure for 
-    Nominal Isabelle. 
-*)
-theory Nominal2_Base
-imports Main Infinite_Set
-begin
-
-section {* Atoms and Sorts *}
-
-text {* A simple implementation for atom_sorts is strings. *}
-(* types atom_sort = string *)
-
-text {* To deal with Church-like binding we use trees of  
-  strings as sorts. *}
-
-datatype atom_sort = Sort "string" "atom_sort list"
-
-datatype atom = Atom atom_sort nat
-
-
-text {* Basic projection function. *}
-
-primrec
-  sort_of :: "atom \<Rightarrow> atom_sort"
-where
-  "sort_of (Atom s i) = s"
-
-
-text {* There are infinitely many atoms of each sort. *}
-lemma INFM_sort_of_eq: 
-  shows "INFM a. sort_of a = s"
-proof -
-  have "INFM i. sort_of (Atom s i) = s" by simp
-  moreover have "inj (Atom s)" by (simp add: inj_on_def)
-  ultimately show "INFM a. sort_of a = s" by (rule INFM_inj)
-qed
-
-lemma infinite_sort_of_eq:
-  shows "infinite {a. sort_of a = s}"
-  using INFM_sort_of_eq unfolding INFM_iff_infinite .
-
-lemma atom_infinite [simp]: 
-  shows "infinite (UNIV :: atom set)"
-  using subset_UNIV infinite_sort_of_eq
-  by (rule infinite_super)
-
-lemma obtain_atom:
-  fixes X :: "atom set"
-  assumes X: "finite X"
-  obtains a where "a \<notin> X" "sort_of a = s"
-proof -
-  from X have "MOST a. a \<notin> X"
-    unfolding MOST_iff_cofinite by simp
-  with INFM_sort_of_eq
-  have "INFM a. sort_of a = s \<and> a \<notin> X"
-    by (rule INFM_conjI)
-  then obtain a where "a \<notin> X" "sort_of a = s"
-    by (auto elim: INFM_E)
-  then show ?thesis ..
-qed
-
-section {* Sort-Respecting Permutations *}
-
-typedef perm =
-  "{f. bij f \<and> finite {a. f a \<noteq> a} \<and> (\<forall>a. sort_of (f a) = sort_of a)}"
-proof
-  show "id \<in> ?perm" by simp
-qed
-
-lemma permI:
-  assumes "bij f" and "MOST x. f x = x" and "\<And>a. sort_of (f a) = sort_of a"
-  shows "f \<in> perm"
-  using assms unfolding perm_def MOST_iff_cofinite by simp
-
-lemma perm_is_bij: "f \<in> perm \<Longrightarrow> bij f"
-  unfolding perm_def by simp
-
-lemma perm_is_finite: "f \<in> perm \<Longrightarrow> finite {a. f a \<noteq> a}"
-  unfolding perm_def by simp
-
-lemma perm_is_sort_respecting: "f \<in> perm \<Longrightarrow> sort_of (f a) = sort_of a"
-  unfolding perm_def by simp
-
-lemma perm_MOST: "f \<in> perm \<Longrightarrow> MOST x. f x = x"
-  unfolding perm_def MOST_iff_cofinite by simp
-
-lemma perm_id: "id \<in> perm"
-  unfolding perm_def by simp
-
-lemma perm_comp:
-  assumes f: "f \<in> perm" and g: "g \<in> perm"
-  shows "(f \<circ> g) \<in> perm"
-apply (rule permI)
-apply (rule bij_comp)
-apply (rule perm_is_bij [OF g])
-apply (rule perm_is_bij [OF f])
-apply (rule MOST_rev_mp [OF perm_MOST [OF g]])
-apply (rule MOST_rev_mp [OF perm_MOST [OF f]])
-apply (simp)
-apply (simp add: perm_is_sort_respecting [OF f])
-apply (simp add: perm_is_sort_respecting [OF g])
-done
-
-lemma perm_inv:
-  assumes f: "f \<in> perm"
-  shows "(inv f) \<in> perm"
-apply (rule permI)
-apply (rule bij_imp_bij_inv)
-apply (rule perm_is_bij [OF f])
-apply (rule MOST_mono [OF perm_MOST [OF f]])
-apply (erule subst, rule inv_f_f)
-apply (rule bij_is_inj [OF perm_is_bij [OF f]])
-apply (rule perm_is_sort_respecting [OF f, THEN sym, THEN trans])
-apply (simp add: surj_f_inv_f [OF bij_is_surj [OF perm_is_bij [OF f]]])
-done
-
-lemma bij_Rep_perm: "bij (Rep_perm p)"
-  using Rep_perm [of p] unfolding perm_def by simp
-
-lemma finite_Rep_perm: "finite {a. Rep_perm p a \<noteq> a}"
-  using Rep_perm [of p] unfolding perm_def by simp
-
-lemma sort_of_Rep_perm: "sort_of (Rep_perm p a) = sort_of a"
-  using Rep_perm [of p] unfolding perm_def by simp
-
-lemma Rep_perm_ext:
-  "Rep_perm p1 = Rep_perm p2 \<Longrightarrow> p1 = p2"
-  by (simp add: expand_fun_eq Rep_perm_inject [symmetric])
-
-
-subsection {* Permutations form a group *}
-
-instantiation perm :: group_add
-begin
-
-definition
-  "0 = Abs_perm id"
-
-definition
-  "- p = Abs_perm (inv (Rep_perm p))"
-
-definition
-  "p + q = Abs_perm (Rep_perm p \<circ> Rep_perm q)"
-
-definition
-  "(p1::perm) - p2 = p1 + - p2"
-
-lemma Rep_perm_0: "Rep_perm 0 = id"
-  unfolding zero_perm_def
-  by (simp add: Abs_perm_inverse perm_id)
-
-lemma Rep_perm_add:
-  "Rep_perm (p1 + p2) = Rep_perm p1 \<circ> Rep_perm p2"
-  unfolding plus_perm_def
-  by (simp add: Abs_perm_inverse perm_comp Rep_perm)
-
-lemma Rep_perm_uminus:
-  "Rep_perm (- p) = inv (Rep_perm p)"
-  unfolding uminus_perm_def
-  by (simp add: Abs_perm_inverse perm_inv Rep_perm)
-
-instance
-apply default
-unfolding Rep_perm_inject [symmetric]
-unfolding minus_perm_def
-unfolding Rep_perm_add
-unfolding Rep_perm_uminus
-unfolding Rep_perm_0
-by (simp_all add: o_assoc inv_o_cancel [OF bij_is_inj [OF bij_Rep_perm]])
-
-end
-
-
-section {* Implementation of swappings *}
-
-definition
-  swap :: "atom \<Rightarrow> atom \<Rightarrow> perm" ("'(_ \<rightleftharpoons> _')")
-where
-  "(a \<rightleftharpoons> b) =
-    Abs_perm (if sort_of a = sort_of b 
-              then (\<lambda>c. if a = c then b else if b = c then a else c) 
-              else id)"
-
-lemma Rep_perm_swap:
-  "Rep_perm (a \<rightleftharpoons> b) =
-    (if sort_of a = sort_of b 
-     then (\<lambda>c. if a = c then b else if b = c then a else c)
-     else id)"
-unfolding swap_def
-apply (rule Abs_perm_inverse)
-apply (rule permI)
-apply (auto simp add: bij_def inj_on_def surj_def)[1]
-apply (rule MOST_rev_mp [OF MOST_neq(1) [of a]])
-apply (rule MOST_rev_mp [OF MOST_neq(1) [of b]])
-apply (simp)
-apply (simp)
-done
-
-lemmas Rep_perm_simps =
-  Rep_perm_0
-  Rep_perm_add
-  Rep_perm_uminus
-  Rep_perm_swap
-
-lemma swap_different_sorts [simp]:
-  "sort_of a \<noteq> sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) = 0"
-  by (rule Rep_perm_ext) (simp add: Rep_perm_simps)
-
-lemma swap_cancel:
-  "(a \<rightleftharpoons> b) + (a \<rightleftharpoons> b) = 0"
-by (rule Rep_perm_ext) 
-   (simp add: Rep_perm_simps expand_fun_eq)
-
-lemma swap_self [simp]:
-  "(a \<rightleftharpoons> a) = 0"
-  by (rule Rep_perm_ext, simp add: Rep_perm_simps expand_fun_eq)
-
-lemma minus_swap [simp]:
-  "- (a \<rightleftharpoons> b) = (a \<rightleftharpoons> b)"
-  by (rule minus_unique [OF swap_cancel])
-
-lemma swap_commute:
-  "(a \<rightleftharpoons> b) = (b \<rightleftharpoons> a)"
-  by (rule Rep_perm_ext)
-     (simp add: Rep_perm_swap expand_fun_eq)
-
-lemma swap_triple:
-  assumes "a \<noteq> b" and "c \<noteq> b"
-  assumes "sort_of a = sort_of b" "sort_of b = sort_of c"
-  shows "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
-  using assms
-  by (rule_tac Rep_perm_ext)
-     (auto simp add: Rep_perm_simps expand_fun_eq)
-
-
-section {* Permutation Types *}
-
-text {*
-  Infix syntax for @{text permute} has higher precedence than
-  addition, but lower than unary minus.
-*}
-
-class pt =
-  fixes permute :: "perm \<Rightarrow> 'a \<Rightarrow> 'a" ("_ \<bullet> _" [76, 75] 75)
-  assumes permute_zero [simp]: "0 \<bullet> x = x"
-  assumes permute_plus [simp]: "(p + q) \<bullet> x = p \<bullet> (q \<bullet> x)"
-begin
-
-lemma permute_diff [simp]:
-  shows "(p - q) \<bullet> x = p \<bullet> - q \<bullet> x"
-  unfolding diff_minus by simp
-
-lemma permute_minus_cancel [simp]:
-  shows "p \<bullet> - p \<bullet> x = x"
-  and   "- p \<bullet> p \<bullet> x = x"
-  unfolding permute_plus [symmetric] by simp_all
-
-lemma permute_swap_cancel [simp]:
-  shows "(a \<rightleftharpoons> b) \<bullet> (a \<rightleftharpoons> b) \<bullet> x = x"
-  unfolding permute_plus [symmetric]
-  by (simp add: swap_cancel)
-
-lemma permute_swap_cancel2 [simp]:
-  shows "(a \<rightleftharpoons> b) \<bullet> (b \<rightleftharpoons> a) \<bullet> x = x"
-  unfolding permute_plus [symmetric]
-  by (simp add: swap_commute)
-
-lemma inj_permute [simp]: 
-  shows "inj (permute p)"
-  by (rule inj_on_inverseI)
-     (rule permute_minus_cancel)
-
-lemma surj_permute [simp]: 
-  shows "surj (permute p)"
-  by (rule surjI, rule permute_minus_cancel)
-
-lemma bij_permute [simp]: 
-  shows "bij (permute p)"
-  by (rule bijI [OF inj_permute surj_permute])
-
-lemma inv_permute: 
-  shows "inv (permute p) = permute (- p)"
-  by (rule inv_equality) (simp_all)
-
-lemma permute_minus: 
-  shows "permute (- p) = inv (permute p)"
-  by (simp add: inv_permute)
-
-lemma permute_eq_iff [simp]: 
-  shows "p \<bullet> x = p \<bullet> y \<longleftrightarrow> x = y"
-  by (rule inj_permute [THEN inj_eq])
-
-end
-
-subsection {* Permutations for atoms *}
-
-instantiation atom :: pt
-begin
-
-definition
-  "p \<bullet> a = Rep_perm p a"
-
-instance 
-apply(default)
-apply(simp_all add: permute_atom_def Rep_perm_simps)
-done
-
-end
-
-lemma sort_of_permute [simp]:
-  shows "sort_of (p \<bullet> a) = sort_of a"
-  unfolding permute_atom_def by (rule sort_of_Rep_perm)
-
-lemma swap_atom:
-  shows "(a \<rightleftharpoons> b) \<bullet> c =
-           (if sort_of a = sort_of b
-            then (if c = a then b else if c = b then a else c) else c)"
-  unfolding permute_atom_def
-  by (simp add: Rep_perm_swap)
-
-lemma swap_atom_simps [simp]:
-  "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> a = b"
-  "sort_of a = sort_of b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> b = a"
-  "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> c = c"
-  unfolding swap_atom by simp_all
-
-lemma expand_perm_eq:
-  fixes p q :: "perm"
-  shows "p = q \<longleftrightarrow> (\<forall>a::atom. p \<bullet> a = q \<bullet> a)"
-  unfolding permute_atom_def
-  by (metis Rep_perm_ext ext)
-
-
-subsection {* Permutations for permutations *}
-
-instantiation perm :: pt
-begin
-
-definition
-  "p \<bullet> q = p + q - p"
-
-instance
-apply default
-apply (simp add: permute_perm_def)
-apply (simp add: permute_perm_def diff_minus minus_add add_assoc)
-done
-
-end
-
-lemma permute_self: "p \<bullet> p = p"
-unfolding permute_perm_def by (simp add: diff_minus add_assoc)
-
-lemma permute_eqvt:
-  shows "p \<bullet> (q \<bullet> x) = (p \<bullet> q) \<bullet> (p \<bullet> x)"
-  unfolding permute_perm_def by simp
-
-lemma zero_perm_eqvt:
-  shows "p \<bullet> (0::perm) = 0"
-  unfolding permute_perm_def by simp
-
-lemma add_perm_eqvt:
-  fixes p p1 p2 :: perm
-  shows "p \<bullet> (p1 + p2) = p \<bullet> p1 + p \<bullet> p2"
-  unfolding permute_perm_def
-  by (simp add: expand_perm_eq)
-
-lemma swap_eqvt:
-  shows "p \<bullet> (a \<rightleftharpoons> b) = (p \<bullet> a \<rightleftharpoons> p \<bullet> b)"
-  unfolding permute_perm_def
-  by (auto simp add: swap_atom expand_perm_eq)
-
-
-subsection {* Permutations for functions *}
-
-instantiation "fun" :: (pt, pt) pt
-begin
-
-definition
-  "p \<bullet> f = (\<lambda>x. p \<bullet> (f (- p \<bullet> x)))"
-
-instance
-apply default
-apply (simp add: permute_fun_def)
-apply (simp add: permute_fun_def minus_add)
-done
-
-end
-
-lemma permute_fun_app_eq:
-  shows "p \<bullet> (f x) = (p \<bullet> f) (p \<bullet> x)"
-unfolding permute_fun_def by simp
-
-
-subsection {* Permutations for booleans *}
-
-instantiation bool :: pt
-begin
-
-definition "p \<bullet> (b::bool) = b"
-
-instance
-apply(default) 
-apply(simp_all add: permute_bool_def)
-done
-
-end
-
-lemma Not_eqvt:
-  shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))"
-by (simp add: permute_bool_def)
-
-
-subsection {* Permutations for sets *}
-
-lemma permute_set_eq:
-  fixes x::"'a::pt"
-  and   p::"perm"
-  shows "(p \<bullet> X) = {p \<bullet> x | x. x \<in> X}"
-  apply(auto simp add: permute_fun_def permute_bool_def mem_def)
-  apply(rule_tac x="- p \<bullet> x" in exI)
-  apply(simp)
-  done
-
-lemma permute_set_eq_image:
-  shows "p \<bullet> X = permute p ` X"
-unfolding permute_set_eq by auto
-
-lemma permute_set_eq_vimage:
-  shows "p \<bullet> X = permute (- p) -` X"
-unfolding permute_fun_def permute_bool_def
-unfolding vimage_def Collect_def mem_def ..
-
-lemma swap_set_not_in:
-  assumes a: "a \<notin> S" "b \<notin> S"
-  shows "(a \<rightleftharpoons> b) \<bullet> S = S"
-  using a by (auto simp add: permute_set_eq swap_atom)
-
-lemma swap_set_in:
-  assumes a: "a \<in> S" "b \<notin> S" "sort_of a = sort_of b"
-  shows "(a \<rightleftharpoons> b) \<bullet> S \<noteq> S"
-  using a by (auto simp add: permute_set_eq swap_atom)
-
-
-subsection {* Permutations for units *}
-
-instantiation unit :: pt
-begin
-
-definition "p \<bullet> (u::unit) = u"
-
-instance proof
-qed (simp_all add: permute_unit_def)
-
-end
-
-
-subsection {* Permutations for products *}
-
-instantiation "*" :: (pt, pt) pt
-begin
-
-primrec 
-  permute_prod 
-where
-  Pair_eqvt: "p \<bullet> (x, y) = (p \<bullet> x, p \<bullet> y)"
-
-instance
-by default auto
-
-end
-
-subsection {* Permutations for sums *}
-
-instantiation "+" :: (pt, pt) pt
-begin
-
-primrec 
-  permute_sum 
-where
-  "p \<bullet> (Inl x) = Inl (p \<bullet> x)"
-| "p \<bullet> (Inr y) = Inr (p \<bullet> y)"
-
-instance proof
-qed (case_tac [!] x, simp_all)
-
-end
-
-subsection {* Permutations for lists *}
-
-instantiation list :: (pt) pt
-begin
-
-primrec 
-  permute_list 
-where
-  "p \<bullet> [] = []"
-| "p \<bullet> (x # xs) = p \<bullet> x # p \<bullet> xs"
-
-instance proof
-qed (induct_tac [!] x, simp_all)
-
-end
-
-subsection {* Permutations for options *}
-
-instantiation option :: (pt) pt
-begin
-
-primrec 
-  permute_option 
-where
-  "p \<bullet> None = None"
-| "p \<bullet> (Some x) = Some (p \<bullet> x)"
-
-instance proof
-qed (induct_tac [!] x, simp_all)
-
-end
-
-subsection {* Permutations for @{typ char}, @{typ nat}, and @{typ int} *}
-
-instantiation char :: pt
-begin
-
-definition "p \<bullet> (c::char) = c"
-
-instance proof
-qed (simp_all add: permute_char_def)
-
-end
-
-instantiation nat :: pt
-begin
-
-definition "p \<bullet> (n::nat) = n"
-
-instance proof
-qed (simp_all add: permute_nat_def)
-
-end
-
-instantiation int :: pt
-begin
-
-definition "p \<bullet> (i::int) = i"
-
-instance proof
-qed (simp_all add: permute_int_def)
-
-end
-
-
-section {* Pure types *}
-
-text {* Pure types will have always empty support. *}
-
-class pure = pt +
-  assumes permute_pure: "p \<bullet> x = x"
-
-text {* Types @{typ unit} and @{typ bool} are pure. *}
-
-instance unit :: pure
-proof qed (rule permute_unit_def)
-
-instance bool :: pure
-proof qed (rule permute_bool_def)
-
-text {* Other type constructors preserve purity. *}
-
-instance "fun" :: (pure, pure) pure
-by default (simp add: permute_fun_def permute_pure)
-
-instance "*" :: (pure, pure) pure
-by default (induct_tac x, simp add: permute_pure)
-
-instance "+" :: (pure, pure) pure
-by default (induct_tac x, simp_all add: permute_pure)
-
-instance list :: (pure) pure
-by default (induct_tac x, simp_all add: permute_pure)
-
-instance option :: (pure) pure
-by default (induct_tac x, simp_all add: permute_pure)
-
-
-subsection {* Types @{typ char}, @{typ nat}, and @{typ int} *}
-
-instance char :: pure
-proof qed (rule permute_char_def)
-
-instance nat :: pure
-proof qed (rule permute_nat_def)
-
-instance int :: pure
-proof qed (rule permute_int_def)
-
-
-subsection {* Supp, Freshness and Supports *}
-
-context pt
-begin
-
-definition
-  supp :: "'a \<Rightarrow> atom set"
-where
-  "supp x = {a. infinite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}}"
-
-end
-
-definition
-  fresh :: "atom \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp> _" [55, 55] 55)
-where   
-  "a \<sharp> x \<equiv> a \<notin> supp x"
-
-lemma supp_conv_fresh: 
-  shows "supp x = {a. \<not> a \<sharp> x}"
-  unfolding fresh_def by simp
-
-lemma swap_rel_trans:
-  assumes "sort_of a = sort_of b"
-  assumes "sort_of b = sort_of c"
-  assumes "(a \<rightleftharpoons> c) \<bullet> x = x"
-  assumes "(b \<rightleftharpoons> c) \<bullet> x = x"
-  shows "(a \<rightleftharpoons> b) \<bullet> x = x"
-proof (cases)
-  assume "a = b \<or> c = b"
-  with assms show "(a \<rightleftharpoons> b) \<bullet> x = x" by auto
-next
-  assume *: "\<not> (a = b \<or> c = b)"
-  have "((a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c)) \<bullet> x = x"
-    using assms by simp
-  also have "(a \<rightleftharpoons> c) + (b \<rightleftharpoons> c) + (a \<rightleftharpoons> c) = (a \<rightleftharpoons> b)"
-    using assms * by (simp add: swap_triple)
-  finally show "(a \<rightleftharpoons> b) \<bullet> x = x" .
-qed
-
-lemma swap_fresh_fresh:
-  assumes a: "a \<sharp> x" 
-  and     b: "b \<sharp> x"
-  shows "(a \<rightleftharpoons> b) \<bullet> x = x"
-proof (cases)
-  assume asm: "sort_of a = sort_of b" 
-  have "finite {c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x}" "finite {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x}" 
-    using a b unfolding fresh_def supp_def by simp_all
-  then have "finite ({c. (a \<rightleftharpoons> c) \<bullet> x \<noteq> x} \<union> {c. (b \<rightleftharpoons> c) \<bullet> x \<noteq> x})" by simp
-  then obtain c 
-    where "(a \<rightleftharpoons> c) \<bullet> x = x" "(b \<rightleftharpoons> c) \<bullet> x = x" "sort_of c = sort_of b"
-    by (rule obtain_atom) (auto)
-  then show "(a \<rightleftharpoons> b) \<bullet> x = x" using asm by (rule_tac swap_rel_trans) (simp_all)
-next
-  assume "sort_of a \<noteq> sort_of b"
-  then show "(a \<rightleftharpoons> b) \<bullet> x = x" by simp
-qed
-
-
-subsection {* supp and fresh are equivariant *}
-
-lemma finite_Collect_bij:
-  assumes a: "bij f"
-  shows "finite {x. P (f x)} = finite {x. P x}"
-by (metis a finite_vimage_iff vimage_Collect_eq)
-
-lemma fresh_permute_iff:
-  shows "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> a \<sharp> x"
-proof -
-  have "(p \<bullet> a) \<sharp> (p \<bullet> x) \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
-    unfolding fresh_def supp_def by simp
-  also have "\<dots> \<longleftrightarrow> finite {b. (p \<bullet> a \<rightleftharpoons> p \<bullet> b) \<bullet> p \<bullet> x \<noteq> p \<bullet> x}"
-    using bij_permute by (rule finite_Collect_bij [symmetric])
-  also have "\<dots> \<longleftrightarrow> finite {b. p \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> p \<bullet> x}"
-    by (simp only: permute_eqvt [of p] swap_eqvt)
-  also have "\<dots> \<longleftrightarrow> finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"
-    by (simp only: permute_eq_iff)
-  also have "\<dots> \<longleftrightarrow> a \<sharp> x"
-    unfolding fresh_def supp_def by simp
-  finally show ?thesis .
-qed
-
-lemma fresh_eqvt:
-  shows "p \<bullet> (a \<sharp> x) = (p \<bullet> a) \<sharp> (p \<bullet> x)"
-  by (simp add: permute_bool_def fresh_permute_iff)
-
-lemma supp_eqvt:
-  fixes  p  :: "perm"
-  and    x   :: "'a::pt"
-  shows "p \<bullet> (supp x) = supp (p \<bullet> x)"
-  unfolding supp_conv_fresh
-  unfolding permute_fun_def Collect_def
-  by (simp add: Not_eqvt fresh_eqvt)
-
-subsection {* supports *}
-
-definition
-  supports :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" (infixl "supports" 80)
-where  
-  "S supports x \<equiv> \<forall>a b. (a \<notin> S \<and> b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x)"
-
-lemma supp_is_subset:
-  fixes S :: "atom set"
-  and   x :: "'a::pt"
-  assumes a1: "S supports x"
-  and     a2: "finite S"
-  shows "(supp x) \<subseteq> S"
-proof (rule ccontr)
-  assume "\<not>(supp x \<subseteq> S)"
-  then obtain a where b1: "a \<in> supp x" and b2: "a \<notin> S" by auto
-  from a1 b2 have "\<forall>b. b \<notin> S \<longrightarrow> (a \<rightleftharpoons> b) \<bullet> x = x" by (unfold supports_def) (auto)
-  hence "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x} \<subseteq> S" by auto
-  with a2 have "finite {b. (a \<rightleftharpoons> b)\<bullet>x \<noteq> x}" by (simp add: finite_subset)
-  then have "a \<notin> (supp x)" unfolding supp_def by simp
-  with b1 show False by simp
-qed
-
-lemma supports_finite:
-  fixes S :: "atom set"
-  and   x :: "'a::pt"
-  assumes a1: "S supports x"
-  and     a2: "finite S"
-  shows "finite (supp x)"
-proof -
-  have "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
-  then show "finite (supp x)" using a2 by (simp add: finite_subset)
-qed
-
-lemma supp_supports:
-  fixes x :: "'a::pt"
-  shows "(supp x) supports x"
-proof (unfold supports_def, intro strip)
-  fix a b
-  assume "a \<notin> (supp x) \<and> b \<notin> (supp x)"
-  then have "a \<sharp> x" and "b \<sharp> x" by (simp_all add: fresh_def)
-  then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (rule swap_fresh_fresh)
-qed
-
-lemma supp_is_least_supports:
-  fixes S :: "atom set"
-  and   x :: "'a::pt"
-  assumes  a1: "S supports x"
-  and      a2: "finite S"
-  and      a3: "\<And>S'. finite S' \<Longrightarrow> (S' supports x) \<Longrightarrow> S \<subseteq> S'"
-  shows "(supp x) = S"
-proof (rule equalityI)
-  show "(supp x) \<subseteq> S" using a1 a2 by (rule supp_is_subset)
-  with a2 have fin: "finite (supp x)" by (rule rev_finite_subset)
-  have "(supp x) supports x" by (rule supp_supports)
-  with fin a3 show "S \<subseteq> supp x" by blast
-qed
-
-lemma subsetCI: 
-  shows "(\<And>x. x \<in> A \<Longrightarrow> x \<notin> B \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> B"
-  by auto
-
-lemma finite_supp_unique:
-  assumes a1: "S supports x"
-  assumes a2: "finite S"
-  assumes a3: "\<And>a b. \<lbrakk>a \<in> S; b \<notin> S; sort_of a = sort_of b\<rbrakk> \<Longrightarrow> (a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
-  shows "(supp x) = S"
-  using a1 a2
-proof (rule supp_is_least_supports)
-  fix S'
-  assume "finite S'" and "S' supports x"
-  show "S \<subseteq> S'"
-  proof (rule subsetCI)
-    fix a
-    assume "a \<in> S" and "a \<notin> S'"
-    have "finite (S \<union> S')"
-      using `finite S` `finite S'` by simp
-    then obtain b where "b \<notin> S \<union> S'" and "sort_of b = sort_of a"
-      by (rule obtain_atom)
-    then have "b \<notin> S" and "b \<notin> S'"  and "sort_of a = sort_of b"
-      by simp_all
-    then have "(a \<rightleftharpoons> b) \<bullet> x = x"
-      using `a \<notin> S'` `S' supports x` by (simp add: supports_def)
-    moreover have "(a \<rightleftharpoons> b) \<bullet> x \<noteq> x"
-      using `a \<in> S` `b \<notin> S` `sort_of a = sort_of b`
-      by (rule a3)
-    ultimately show "False" by simp
-  qed
-qed
-
-section {* Finitely-supported types *}
-
-class fs = pt +
-  assumes finite_supp: "finite (supp x)"
-
-lemma pure_supp: 
-  shows "supp (x::'a::pure) = {}"
-  unfolding supp_def by (simp add: permute_pure)
-
-lemma pure_fresh:
-  fixes x::"'a::pure"
-  shows "a \<sharp> x"
-  unfolding fresh_def by (simp add: pure_supp)
-
-instance pure < fs
-by default (simp add: pure_supp)
-
-
-subsection  {* Type @{typ atom} is finitely-supported. *}
-
-lemma supp_atom:
-  shows "supp a = {a}"
-apply (rule finite_supp_unique)
-apply (clarsimp simp add: supports_def)
-apply simp
-apply simp
-done
-
-lemma fresh_atom: 
-  shows "a \<sharp> b \<longleftrightarrow> a \<noteq> b"
-  unfolding fresh_def supp_atom by simp
-
-instance atom :: fs
-by default (simp add: supp_atom)
-
-
-section {* Type @{typ perm} is finitely-supported. *}
-
-lemma perm_swap_eq:
-  shows "(a \<rightleftharpoons> b) \<bullet> p = p \<longleftrightarrow> (p \<bullet> (a \<rightleftharpoons> b)) = (a \<rightleftharpoons> b)"
-unfolding permute_perm_def
-by (metis add_diff_cancel minus_perm_def)
-
-lemma supports_perm: 
-  shows "{a. p \<bullet> a \<noteq> a} supports p"
-  unfolding supports_def
-  by (simp add: perm_swap_eq swap_eqvt)
-
-lemma finite_perm_lemma: 
-  shows "finite {a::atom. p \<bullet> a \<noteq> a}"
-  using finite_Rep_perm [of p]
-  unfolding permute_atom_def .
-
-lemma supp_perm:
-  shows "supp p = {a. p \<bullet> a \<noteq> a}"
-apply (rule finite_supp_unique)
-apply (rule supports_perm)
-apply (rule finite_perm_lemma)
-apply (simp add: perm_swap_eq swap_eqvt)
-apply (auto simp add: expand_perm_eq swap_atom)
-done
-
-lemma fresh_perm: 
-  shows "a \<sharp> p \<longleftrightarrow> p \<bullet> a = a"
-unfolding fresh_def by (simp add: supp_perm)
-
-lemma supp_swap:
-  shows "supp (a \<rightleftharpoons> b) = (if a = b \<or> sort_of a \<noteq> sort_of b then {} else {a, b})"
-  by (auto simp add: supp_perm swap_atom)
-
-lemma fresh_zero_perm: 
-  shows "a \<sharp> (0::perm)"
-  unfolding fresh_perm by simp
-
-lemma supp_zero_perm: 
-  shows "supp (0::perm) = {}"
-  unfolding supp_perm by simp
-
-lemma fresh_plus_perm:
-  fixes p q::perm
-  assumes "a \<sharp> p" "a \<sharp> q"
-  shows "a \<sharp> (p + q)"
-  using assms
-  unfolding fresh_def
-  by (auto simp add: supp_perm)
-
-lemma supp_plus_perm:
-  fixes p q::perm
-  shows "supp (p + q) \<subseteq> supp p \<union> supp q"
-  by (auto simp add: supp_perm)
-
-lemma fresh_minus_perm:
-  fixes p::perm
-  shows "a \<sharp> (- p) \<longleftrightarrow> a \<sharp> p"
-  unfolding fresh_def
-  apply(auto simp add: supp_perm)
-  apply(metis permute_minus_cancel)+
-  done
-
-lemma supp_minus_perm:
-  fixes p::perm
-  shows "supp (- p) = supp p"
-  unfolding supp_conv_fresh
-  by (simp add: fresh_minus_perm)
-
-instance perm :: fs
-by default (simp add: supp_perm finite_perm_lemma)
-
-
-section {* Finite Support instances for other types *}
-
-subsection {* Type @{typ "'a \<times> 'b"} is finitely-supported. *}
-
-lemma supp_Pair: 
-  shows "supp (x, y) = supp x \<union> supp y"
-  by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
-
-lemma fresh_Pair: 
-  shows "a \<sharp> (x, y) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> y"
-  by (simp add: fresh_def supp_Pair)
-
-instance "*" :: (fs, fs) fs
-apply default
-apply (induct_tac x)
-apply (simp add: supp_Pair finite_supp)
-done
-
-subsection {* Type @{typ "'a + 'b"} is finitely supported *}
-
-lemma supp_Inl: 
-  shows "supp (Inl x) = supp x"
-  by (simp add: supp_def)
-
-lemma supp_Inr: 
-  shows "supp (Inr x) = supp x"
-  by (simp add: supp_def)
-
-lemma fresh_Inl: 
-  shows "a \<sharp> Inl x \<longleftrightarrow> a \<sharp> x"
-  by (simp add: fresh_def supp_Inl)
-
-lemma fresh_Inr: 
-  shows "a \<sharp> Inr y \<longleftrightarrow> a \<sharp> y"
-  by (simp add: fresh_def supp_Inr)
-
-instance "+" :: (fs, fs) fs
-apply default
-apply (induct_tac x)
-apply (simp_all add: supp_Inl supp_Inr finite_supp)
-done
-
-subsection {* Type @{typ "'a option"} is finitely supported *}
-
-lemma supp_None: 
-  shows "supp None = {}"
-by (simp add: supp_def)
-
-lemma supp_Some: 
-  shows "supp (Some x) = supp x"
-  by (simp add: supp_def)
-
-lemma fresh_None: 
-  shows "a \<sharp> None"
-  by (simp add: fresh_def supp_None)
-
-lemma fresh_Some: 
-  shows "a \<sharp> Some x \<longleftrightarrow> a \<sharp> x"
-  by (simp add: fresh_def supp_Some)
-
-instance option :: (fs) fs
-apply default
-apply (induct_tac x)
-apply (simp_all add: supp_None supp_Some finite_supp)
-done
-
-subsubsection {* Type @{typ "'a list"} is finitely supported *}
-
-lemma supp_Nil: 
-  shows "supp [] = {}"
-  by (simp add: supp_def)
-
-lemma supp_Cons: 
-  shows "supp (x # xs) = supp x \<union> supp xs"
-by (simp add: supp_def Collect_imp_eq Collect_neg_eq)
-
-lemma fresh_Nil: 
-  shows "a \<sharp> []"
-  by (simp add: fresh_def supp_Nil)
-
-lemma fresh_Cons: 
-  shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
-  by (simp add: fresh_def supp_Cons)
-
-instance list :: (fs) fs
-apply default
-apply (induct_tac x)
-apply (simp_all add: supp_Nil supp_Cons finite_supp)
-done
-
-section {* Support and freshness for applications *}
-
-lemma supp_fun_app:
-  shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)"
-proof (rule subsetCI)
-  fix a::"atom"
-  assume a: "a \<in> supp (f x)"
-  assume b: "a \<notin> supp f \<union> supp x"
-  then have "finite {b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f}" "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}" 
-    unfolding supp_def by auto
-  then have "finite ({b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f} \<union> {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x})" by simp
-  moreover 
-  have "{b. ((a \<rightleftharpoons> b) \<bullet> f) ((a \<rightleftharpoons> b) \<bullet> x) \<noteq> f x} \<subseteq> ({b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f} \<union> {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x})"
-    by auto
-  ultimately have "finite {b. ((a \<rightleftharpoons> b) \<bullet> f) ((a \<rightleftharpoons> b) \<bullet> x) \<noteq> f x}"
-    using finite_subset by auto
-  then have "a \<notin> supp (f x)" unfolding supp_def
-    by (simp add: permute_fun_app_eq)
-  with a show "False" by simp
-qed
-    
-lemma fresh_fun_app:
-  shows "a \<sharp> (f, x) \<Longrightarrow> a \<sharp> f x"
-unfolding fresh_def
-using supp_fun_app
-by (auto simp add: supp_Pair)
-
-lemma fresh_fun_eqvt_app:
-  assumes a: "\<forall>p. p \<bullet> f = f"
-  shows "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
-proof -
-  from a have b: "supp f = {}"
-    unfolding supp_def by simp
-  show "a \<sharp> x \<Longrightarrow> a \<sharp> f x"
-    unfolding fresh_def
-    using supp_fun_app b
-    by auto
-qed
-
-end
--- a/Quot/Nominal/Nominal2_Eqvt.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,304 +0,0 @@
-(*  Title:      Nominal2_Eqvt
-    Authors:    Brian Huffman, Christian Urban
-
-    Equivariance, Supp and Fresh Lemmas for Operators. 
-    (Contains most, but not all such lemmas.)
-*)
-theory Nominal2_Eqvt
-imports Nominal2_Base
-uses ("nominal_thmdecls.ML")
-     ("nominal_permeq.ML")
-begin
-
-section {* Logical Operators *}
-
-
-lemma eq_eqvt:
-  shows "p \<bullet> (x = y) \<longleftrightarrow> (p \<bullet> x) = (p \<bullet> y)"
-  unfolding permute_eq_iff permute_bool_def ..
-
-lemma if_eqvt:
-  shows "p \<bullet> (if b then x else y) = (if p \<bullet> b then p \<bullet> x else p \<bullet> y)"
-  by (simp add: permute_fun_def permute_bool_def)
-
-lemma True_eqvt:
-  shows "p \<bullet> True = True"
-  unfolding permute_bool_def ..
-
-lemma False_eqvt:
-  shows "p \<bullet> False = False"
-  unfolding permute_bool_def ..
-
-lemma imp_eqvt:
-  shows "p \<bullet> (A \<longrightarrow> B) = ((p \<bullet> A) \<longrightarrow> (p \<bullet> B))"
-  by (simp add: permute_bool_def)
-
-lemma conj_eqvt:
-  shows "p \<bullet> (A \<and> B) = ((p \<bullet> A) \<and> (p \<bullet> B))"
-  by (simp add: permute_bool_def)
-
-lemma disj_eqvt:
-  shows "p \<bullet> (A \<or> B) = ((p \<bullet> A) \<or> (p \<bullet> B))"
-  by (simp add: permute_bool_def)
-
-lemma Not_eqvt:
-  shows "p \<bullet> (\<not> A) = (\<not> (p \<bullet> A))"
-  by (simp add: permute_bool_def)
-
-lemma all_eqvt:
-  shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. (p \<bullet> P) x)"
-  unfolding permute_fun_def permute_bool_def
-  by (auto, drule_tac x="p \<bullet> x" in spec, simp)
-
-lemma all_eqvt2:
-  shows "p \<bullet> (\<forall>x. P x) = (\<forall>x. p \<bullet> P (- p \<bullet> x))"
-  unfolding permute_fun_def permute_bool_def
-  by (auto, drule_tac x="p \<bullet> x" in spec, simp)
-
-lemma ex_eqvt:
-  shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. (p \<bullet> P) x)"
-  unfolding permute_fun_def permute_bool_def
-  by (auto, rule_tac x="p \<bullet> x" in exI, simp)
-
-lemma ex_eqvt2:
-  shows "p \<bullet> (\<exists>x. P x) = (\<exists>x. p \<bullet> P (- p \<bullet> x))"
-  unfolding permute_fun_def permute_bool_def
-  by (auto, rule_tac x="p \<bullet> x" in exI, simp)
-
-lemma ex1_eqvt:
-  shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. (p \<bullet> P) x)"
-  unfolding Ex1_def 
-  by (simp add: ex_eqvt permute_fun_def conj_eqvt all_eqvt imp_eqvt eq_eqvt)
-
-lemma ex1_eqvt2:
-  shows "p \<bullet> (\<exists>!x. P x) = (\<exists>!x. p \<bullet> P (- p \<bullet> x))"
-  unfolding Ex1_def ex_eqvt2 conj_eqvt all_eqvt2 imp_eqvt eq_eqvt
-  by simp
-
-lemma the_eqvt:
-  assumes unique: "\<exists>!x. P x"
-  shows "(p \<bullet> (THE x. P x)) = (THE x. p \<bullet> P (- p \<bullet> x))"
-  apply(rule the1_equality [symmetric])
-  apply(simp add: ex1_eqvt2[symmetric])
-  apply(simp add: permute_bool_def unique)
-  apply(simp add: permute_bool_def)
-  apply(rule theI'[OF unique])
-  done
-
-section {* Set Operations *}
-
-lemma mem_permute_iff:
-  shows "(p \<bullet> x) \<in> (p \<bullet> X) \<longleftrightarrow> x \<in> X"
-unfolding mem_def permute_fun_def permute_bool_def
-by simp
-
-lemma mem_eqvt:
-  shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
-  unfolding mem_permute_iff permute_bool_def by simp
-
-lemma not_mem_eqvt:
-  shows "p \<bullet> (x \<notin> A) \<longleftrightarrow> (p \<bullet> x) \<notin> (p \<bullet> A)"
-  unfolding mem_def permute_fun_def by (simp add: Not_eqvt)
-
-lemma Collect_eqvt:
-  shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
-  unfolding Collect_def permute_fun_def ..
-
-lemma Collect_eqvt2:
-  shows "p \<bullet> {x. P x} = {x. p \<bullet> (P (-p \<bullet> x))}"
-  unfolding Collect_def permute_fun_def ..
-
-lemma empty_eqvt:
-  shows "p \<bullet> {} = {}"
-  unfolding empty_def Collect_eqvt2 False_eqvt ..
-
-lemma supp_set_empty:
-  shows "supp {} = {}"
-  by (simp add: supp_def empty_eqvt)
-
-lemma fresh_set_empty:
-  shows "a \<sharp> {}"
-  by (simp add: fresh_def supp_set_empty)
-
-lemma UNIV_eqvt:
-  shows "p \<bullet> UNIV = UNIV"
-  unfolding UNIV_def Collect_eqvt2 True_eqvt ..
-
-lemma union_eqvt:
-  shows "p \<bullet> (A \<union> B) = (p \<bullet> A) \<union> (p \<bullet> B)"
-  unfolding Un_def Collect_eqvt2 disj_eqvt mem_eqvt by simp
-
-lemma inter_eqvt:
-  shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)"
-  unfolding Int_def Collect_eqvt2 conj_eqvt mem_eqvt by simp
-
-lemma Diff_eqvt:
-  fixes A B :: "'a::pt set"
-  shows "p \<bullet> (A - B) = p \<bullet> A - p \<bullet> B"
-  unfolding set_diff_eq Collect_eqvt2 conj_eqvt Not_eqvt mem_eqvt by simp
-
-lemma Compl_eqvt:
-  fixes A :: "'a::pt set"
-  shows "p \<bullet> (- A) = - (p \<bullet> A)"
-  unfolding Compl_eq_Diff_UNIV Diff_eqvt UNIV_eqvt ..
-
-lemma insert_eqvt:
-  shows "p \<bullet> (insert x A) = insert (p \<bullet> x) (p \<bullet> A)"
-  unfolding permute_set_eq_image image_insert ..
-
-lemma vimage_eqvt:
-  shows "p \<bullet> (f -` A) = (p \<bullet> f) -` (p \<bullet> A)"
-  unfolding vimage_def permute_fun_def [where f=f]
-  unfolding Collect_eqvt2 mem_eqvt ..
-
-lemma image_eqvt:
-  shows "p \<bullet> (f ` A) = (p \<bullet> f) ` (p \<bullet> A)"
-  unfolding permute_set_eq_image
-  unfolding permute_fun_def [where f=f]
-  by (simp add: image_image)
-
-lemma finite_permute_iff:
-  shows "finite (p \<bullet> A) \<longleftrightarrow> finite A"
-  unfolding permute_set_eq_vimage
-  using bij_permute by (rule finite_vimage_iff)
-
-lemma finite_eqvt:
-  shows "p \<bullet> finite A = finite (p \<bullet> A)"
-  unfolding finite_permute_iff permute_bool_def ..
-
-
-section {* List Operations *}
-
-lemma append_eqvt:
-  shows "p \<bullet> (xs @ ys) = (p \<bullet> xs) @ (p \<bullet> ys)"
-  by (induct xs) auto
-
-lemma supp_append:
-  shows "supp (xs @ ys) = supp xs \<union> supp ys"
-  by (induct xs) (auto simp add: supp_Nil supp_Cons)
-
-lemma fresh_append:
-  shows "a \<sharp> (xs @ ys) \<longleftrightarrow> a \<sharp> xs \<and> a \<sharp> ys"
-  by (induct xs) (simp_all add: fresh_Nil fresh_Cons)
-
-lemma rev_eqvt:
-  shows "p \<bullet> (rev xs) = rev (p \<bullet> xs)"
-  by (induct xs) (simp_all add: append_eqvt)
-
-lemma supp_rev:
-  shows "supp (rev xs) = supp xs"
-  by (induct xs) (auto simp add: supp_append supp_Cons supp_Nil)
-
-lemma fresh_rev:
-  shows "a \<sharp> rev xs \<longleftrightarrow> a \<sharp> xs"
-  by (induct xs) (auto simp add: fresh_append fresh_Cons fresh_Nil)
-
-lemma set_eqvt:
-  shows "p \<bullet> (set xs) = set (p \<bullet> xs)"
-  by (induct xs) (simp_all add: empty_eqvt insert_eqvt)
-
-(* needs finite support premise
-lemma supp_set:
-  fixes x :: "'a::pt"
-  shows "supp (set xs) = supp xs"
-*)
-
-
-section {* Product Operations *}
-
-lemma fst_eqvt:
-  "p \<bullet> (fst x) = fst (p \<bullet> x)"
- by (cases x) simp
-
-lemma snd_eqvt:
-  "p \<bullet> (snd x) = snd (p \<bullet> x)"
- by (cases x) simp
-
-
-section {* Units *}
-
-lemma supp_unit:
-  shows "supp () = {}"
-  by (simp add: supp_def)
-
-lemma fresh_unit:
-  shows "a \<sharp> ()"
-  by (simp add: fresh_def supp_unit)
-
-section {* Equivariance automation *}
-
-text {* Setup of the theorem attributes @{text eqvt} and @{text eqvt_force} *}
-
-use "nominal_thmdecls.ML"
-setup "Nominal_ThmDecls.setup"
-
-lemmas [eqvt] = 
-  (* connectives *)
-  eq_eqvt if_eqvt imp_eqvt disj_eqvt conj_eqvt Not_eqvt 
-  True_eqvt False_eqvt ex_eqvt all_eqvt ex1_eqvt
-  imp_eqvt [folded induct_implies_def]
-
-  (* nominal *)
-  permute_eqvt supp_eqvt fresh_eqvt
-  permute_pure
-
-  (* datatypes *)
-  permute_prod.simps append_eqvt rev_eqvt set_eqvt
-  fst_eqvt snd_eqvt
-
-  (* sets *)
-  empty_eqvt UNIV_eqvt union_eqvt inter_eqvt mem_eqvt
-  Diff_eqvt Compl_eqvt insert_eqvt Collect_eqvt
-
-thm eqvts
-thm eqvts_raw
-
-text {* helper lemmas for the eqvt_tac *}
-
-definition
-  "unpermute p = permute (- p)"
-
-lemma eqvt_apply:
-  fixes f :: "'a::pt \<Rightarrow> 'b::pt" 
-  and x :: "'a::pt"
-  shows "p \<bullet> (f x) \<equiv> (p \<bullet> f) (p \<bullet> x)"
-  unfolding permute_fun_def by simp
-
-lemma eqvt_lambda:
-  fixes f :: "'a::pt \<Rightarrow> 'b::pt"
-  shows "p \<bullet> (\<lambda>x. f x) \<equiv> (\<lambda>x. p \<bullet> (f (unpermute p x)))"
-  unfolding permute_fun_def unpermute_def by simp
-
-lemma eqvt_bound:
-  shows "p \<bullet> unpermute p x \<equiv> x"
-  unfolding unpermute_def by simp
-
-use "nominal_permeq.ML"
-
-
-lemma "p \<bullet> (A \<longrightarrow> B = C)"
-apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *}) 
-oops
-
-lemma "p \<bullet> (\<lambda>(x::'a::pt). A \<longrightarrow> (B::'a \<Rightarrow> bool) x = C) = foo"
-apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
-oops
-
-lemma "p \<bullet> (\<lambda>x y. \<exists>z. x = z \<and> x = y \<longrightarrow> z \<noteq> x) = foo"
-apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
-oops
-
-lemma "p \<bullet> (\<lambda>f x. f (g (f x))) = foo"
-apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
-oops
-
-lemma "p \<bullet> (\<lambda>q. q \<bullet> (r \<bullet> x)) = foo"
-apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
-oops
-
-lemma "p \<bullet> (q \<bullet> r \<bullet> x) = foo"
-apply (tactic {* Nominal_Permeq.eqvt_tac @{context} 1 *})
-oops
-
-
-end
\ No newline at end of file
--- a/Quot/Nominal/Nominal2_Supp.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,375 +0,0 @@
-(*  Title:      Nominal2_Supp
-    Authors:    Brian Huffman, Christian Urban
-
-    Supplementary Lemmas and Definitions for 
-    Nominal Isabelle. 
-*)
-theory Nominal2_Supp
-imports Nominal2_Base Nominal2_Eqvt Nominal2_Atoms
-begin
-
-
-section {* Fresh-Star *}
-
-text {* The fresh-star generalisation of fresh is used in strong
-  induction principles. *}
-
-definition 
-  fresh_star :: "atom set \<Rightarrow> 'a::pt \<Rightarrow> bool" ("_ \<sharp>* _" [80,80] 80)
-where 
-  "xs \<sharp>* c \<equiv> \<forall>x \<in> xs. x \<sharp> c"
-
-lemma fresh_star_prod:
-  fixes xs::"atom set"
-  shows "xs \<sharp>* (a, b) = (xs \<sharp>* a \<and> xs \<sharp>* b)"
-  by (auto simp add: fresh_star_def fresh_Pair)
-
-lemma fresh_star_union:
-  shows "(xs \<union> ys) \<sharp>* c = (xs \<sharp>* c \<and> ys \<sharp>* c)"
-  by (auto simp add: fresh_star_def)
-
-lemma fresh_star_insert:
-  shows "(insert x ys) \<sharp>* c = (x \<sharp> c \<and> ys \<sharp>* c)"
-  by (auto simp add: fresh_star_def)
-
-lemma fresh_star_Un_elim:
-  "((S \<union> T) \<sharp>* c \<Longrightarrow> PROP C) \<equiv> (S \<sharp>* c \<Longrightarrow> T \<sharp>* c \<Longrightarrow> PROP C)"
-  unfolding fresh_star_def
-  apply(rule)
-  apply(erule meta_mp)
-  apply(auto)
-  done
-
-lemma fresh_star_insert_elim:
-  "(insert x S \<sharp>* c \<Longrightarrow> PROP C) \<equiv> (x \<sharp> c \<Longrightarrow> S \<sharp>* c \<Longrightarrow> PROP C)"
-  unfolding fresh_star_def
-  by rule (simp_all add: fresh_star_def)
-
-lemma fresh_star_empty_elim:
-  "({} \<sharp>* c \<Longrightarrow> PROP C) \<equiv> PROP C"
-  by (simp add: fresh_star_def)
-
-lemma fresh_star_unit_elim: 
-  shows "(a \<sharp>* () \<Longrightarrow> PROP C) \<equiv> PROP C"
-  by (simp add: fresh_star_def fresh_unit) 
-
-lemma fresh_star_prod_elim: 
-  shows "(a \<sharp>* (x, y) \<Longrightarrow> PROP C) \<equiv> (a \<sharp>* x \<Longrightarrow> a \<sharp>* y \<Longrightarrow> PROP C)"
-  by (rule, simp_all add: fresh_star_prod)
-
-
-section {* Avoiding of atom sets *}
-
-text {* 
-  For every set of atoms, there is another set of atoms
-  avoiding a finitely supported c and there is a permutation
-  which 'translates' between both sets.
-*}
-
-lemma at_set_avoiding_aux:
-  fixes Xs::"atom set"
-  and   As::"atom set"
-  assumes b: "Xs \<subseteq> As"
-  and     c: "finite As"
-  shows "\<exists>p. (p \<bullet> Xs) \<inter> As = {} \<and> (supp p) \<subseteq> (Xs \<union> (p \<bullet> Xs))"
-proof -
-  from b c have "finite Xs" by (rule finite_subset)
-  then show ?thesis using b
-  proof (induct rule: finite_subset_induct)
-    case empty
-    have "0 \<bullet> {} \<inter> As = {}" by simp
-    moreover
-    have "supp (0::perm) \<subseteq> {} \<union> 0 \<bullet> {}" by (simp add: supp_zero_perm)
-    ultimately show ?case by blast
-  next
-    case (insert x Xs)
-    then obtain p where
-      p1: "(p \<bullet> Xs) \<inter> As = {}" and 
-      p2: "supp p \<subseteq> (Xs \<union> (p \<bullet> Xs))" by blast
-    from `x \<in> As` p1 have "x \<notin> p \<bullet> Xs" by fast
-    with `x \<notin> Xs` p2 have "x \<notin> supp p" by fast
-    hence px: "p \<bullet> x = x" unfolding supp_perm by simp
-    have "finite (As \<union> p \<bullet> Xs)"
-      using `finite As` `finite Xs`
-      by (simp add: permute_set_eq_image)
-    then obtain y where "y \<notin> (As \<union> p \<bullet> Xs)" "sort_of y = sort_of x"
-      by (rule obtain_atom)
-    hence y: "y \<notin> As" "y \<notin> p \<bullet> Xs" "sort_of y = sort_of x"
-      by simp_all
-    let ?q = "(x \<rightleftharpoons> y) + p"
-    have q: "?q \<bullet> insert x Xs = insert y (p \<bullet> Xs)"
-      unfolding insert_eqvt
-      using `p \<bullet> x = x` `sort_of y = sort_of x`
-      using `x \<notin> p \<bullet> Xs` `y \<notin> p \<bullet> Xs`
-      by (simp add: swap_atom swap_set_not_in)
-    have "?q \<bullet> insert x Xs \<inter> As = {}"
-      using `y \<notin> As` `p \<bullet> Xs \<inter> As = {}`
-      unfolding q by simp
-    moreover
-    have "supp ?q \<subseteq> insert x Xs \<union> ?q \<bullet> insert x Xs"
-      using p2 unfolding q
-      apply (intro subset_trans [OF supp_plus_perm])
-      apply (auto simp add: supp_swap)
-      done
-    ultimately show ?case by blast
-  qed
-qed
-
-lemma at_set_avoiding:
-  assumes a: "finite Xs"
-  and     b: "finite (supp c)"
-  obtains p::"perm" where "(p \<bullet> Xs)\<sharp>*c" and "(supp p) \<subseteq> (Xs \<union> (p \<bullet> Xs))"
-  using a b at_set_avoiding_aux [where Xs="Xs" and As="Xs \<union> supp c"]
-  unfolding fresh_star_def fresh_def by blast
-
-
-section {* The freshness lemma according to Andrew Pitts *}
-
-lemma fresh_conv_MOST: 
-  shows "a \<sharp> x \<longleftrightarrow> (MOST b. (a \<rightleftharpoons> b) \<bullet> x = x)"
-  unfolding fresh_def supp_def MOST_iff_cofinite by simp
-
-lemma fresh_apply:
-  assumes "a \<sharp> f" and "a \<sharp> x" 
-  shows "a \<sharp> f x"
-  using assms unfolding fresh_conv_MOST
-  unfolding permute_fun_app_eq [where f=f]
-  by (elim MOST_rev_mp, simp)
-
-lemma freshness_lemma:
-  fixes h :: "'a::at \<Rightarrow> 'b::pt"
-  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
-  shows  "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
-proof -
-  from a obtain b where a1: "atom b \<sharp> h" and a2: "atom b \<sharp> h b"
-    by (auto simp add: fresh_Pair)
-  show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
-  proof (intro exI allI impI)
-    fix a :: 'a
-    assume a3: "atom a \<sharp> h"
-    show "h a = h b"
-    proof (cases "a = b")
-      assume "a = b"
-      thus "h a = h b" by simp
-    next
-      assume "a \<noteq> b"
-      hence "atom a \<sharp> b" by (simp add: fresh_at_base)
-      with a3 have "atom a \<sharp> h b" by (rule fresh_apply)
-      with a2 have d1: "(atom b \<rightleftharpoons> atom a) \<bullet> (h b) = (h b)"
-        by (rule swap_fresh_fresh)
-      from a1 a3 have d2: "(atom b \<rightleftharpoons> atom a) \<bullet> h = h"
-        by (rule swap_fresh_fresh)
-      from d1 have "h b = (atom b \<rightleftharpoons> atom a) \<bullet> (h b)" by simp
-      also have "\<dots> = ((atom b \<rightleftharpoons> atom a) \<bullet> h) ((atom b \<rightleftharpoons> atom a) \<bullet> b)"
-        by (rule permute_fun_app_eq)
-      also have "\<dots> = h a"
-        using d2 by simp
-      finally show "h a = h b"  by simp
-    qed
-  qed
-qed
-
-lemma freshness_lemma_unique:
-  fixes h :: "'a::at \<Rightarrow> 'b::pt"
-  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
-  shows "\<exists>!x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
-proof (rule ex_ex1I)
-  from a show "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
-    by (rule freshness_lemma)
-next
-  fix x y
-  assume x: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = x"
-  assume y: "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = y"
-  from a x y show "x = y"
-    by (auto simp add: fresh_Pair)
-qed
-
-text {* packaging the freshness lemma into a function *}
-
-definition
-  fresh_fun :: "('a::at \<Rightarrow> 'b::pt) \<Rightarrow> 'b"
-where
-  "fresh_fun h = (THE x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x)"
-
-lemma fresh_fun_app:
-  fixes h :: "'a::at \<Rightarrow> 'b::pt"
-  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
-  assumes b: "atom a \<sharp> h"
-  shows "fresh_fun h = h a"
-unfolding fresh_fun_def
-proof (rule the_equality)
-  show "\<forall>a'. atom a' \<sharp> h \<longrightarrow> h a' = h a"
-  proof (intro strip)
-    fix a':: 'a
-    assume c: "atom a' \<sharp> h"
-    from a have "\<exists>x. \<forall>a. atom a \<sharp> h \<longrightarrow> h a = x" by (rule freshness_lemma)
-    with b c show "h a' = h a" by auto
-  qed
-next
-  fix fr :: 'b
-  assume "\<forall>a. atom a \<sharp> h \<longrightarrow> h a = fr"
-  with b show "fr = h a" by auto
-qed
-
-lemma fresh_fun_app':
-  fixes h :: "'a::at \<Rightarrow> 'b::pt"
-  assumes a: "atom a \<sharp> h" "atom a \<sharp> h a"
-  shows "fresh_fun h = h a"
-  apply (rule fresh_fun_app)
-  apply (auto simp add: fresh_Pair intro: a)
-  done
-
-lemma fresh_fun_eqvt:
-  fixes h :: "'a::at \<Rightarrow> 'b::pt"
-  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
-  shows "p \<bullet> (fresh_fun h) = fresh_fun (p \<bullet> h)"
-  using a
-  apply (clarsimp simp add: fresh_Pair)
-  apply (subst fresh_fun_app', assumption+)
-  apply (drule fresh_permute_iff [where p=p, THEN iffD2])
-  apply (drule fresh_permute_iff [where p=p, THEN iffD2])
-  apply (simp add: atom_eqvt permute_fun_app_eq [where f=h])
-  apply (erule (1) fresh_fun_app' [symmetric])
-  done
-
-lemma fresh_fun_supports:
-  fixes h :: "'a::at \<Rightarrow> 'b::pt"
-  assumes a: "\<exists>a. atom a \<sharp> (h, h a)"
-  shows "(supp h) supports (fresh_fun h)"
-  apply (simp add: supports_def fresh_def [symmetric])
-  apply (simp add: fresh_fun_eqvt [OF a] swap_fresh_fresh)
-  done
-
-notation fresh_fun (binder "FRESH " 10)
-
-lemma FRESH_f_iff:
-  fixes P :: "'a::at \<Rightarrow> 'b::pure"
-  fixes f :: "'b \<Rightarrow> 'c::pure"
-  assumes P: "finite (supp P)"
-  shows "(FRESH x. f (P x)) = f (FRESH x. P x)"
-proof -
-  obtain a::'a where "atom a \<notin> supp P"
-    using P by (rule obtain_at_base)
-  hence "atom a \<sharp> P"
-    by (simp add: fresh_def)
-  show "(FRESH x. f (P x)) = f (FRESH x. P x)"
-    apply (subst fresh_fun_app' [where a=a, OF _ pure_fresh])
-    apply (cut_tac `atom a \<sharp> P`)
-    apply (simp add: fresh_conv_MOST)
-    apply (elim MOST_rev_mp, rule MOST_I, clarify)
-    apply (simp add: permute_fun_def permute_pure expand_fun_eq)
-    apply (subst fresh_fun_app' [where a=a, OF `atom a \<sharp> P` pure_fresh])
-    apply (rule refl)
-    done
-qed
-
-lemma FRESH_binop_iff:
-  fixes P :: "'a::at \<Rightarrow> 'b::pure"
-  fixes Q :: "'a::at \<Rightarrow> 'c::pure"
-  fixes binop :: "'b \<Rightarrow> 'c \<Rightarrow> 'd::pure"
-  assumes P: "finite (supp P)" 
-  and     Q: "finite (supp Q)"
-  shows "(FRESH x. binop (P x) (Q x)) = binop (FRESH x. P x) (FRESH x. Q x)"
-proof -
-  from assms have "finite (supp P \<union> supp Q)" by simp
-  then obtain a::'a where "atom a \<notin> (supp P \<union> supp Q)"
-    by (rule obtain_at_base)
-  hence "atom a \<sharp> P" and "atom a \<sharp> Q"
-    by (simp_all add: fresh_def)
-  show ?thesis
-    apply (subst fresh_fun_app' [where a=a, OF _ pure_fresh])
-    apply (cut_tac `atom a \<sharp> P` `atom a \<sharp> Q`)
-    apply (simp add: fresh_conv_MOST)
-    apply (elim MOST_rev_mp, rule MOST_I, clarify)
-    apply (simp add: permute_fun_def permute_pure expand_fun_eq)
-    apply (subst fresh_fun_app' [where a=a, OF `atom a \<sharp> P` pure_fresh])
-    apply (subst fresh_fun_app' [where a=a, OF `atom a \<sharp> Q` pure_fresh])
-    apply (rule refl)
-    done
-qed
-
-lemma FRESH_conj_iff:
-  fixes P Q :: "'a::at \<Rightarrow> bool"
-  assumes P: "finite (supp P)" and Q: "finite (supp Q)"
-  shows "(FRESH x. P x \<and> Q x) \<longleftrightarrow> (FRESH x. P x) \<and> (FRESH x. Q x)"
-using P Q by (rule FRESH_binop_iff)
-
-lemma FRESH_disj_iff:
-  fixes P Q :: "'a::at \<Rightarrow> bool"
-  assumes P: "finite (supp P)" and Q: "finite (supp Q)"
-  shows "(FRESH x. P x \<or> Q x) \<longleftrightarrow> (FRESH x. P x) \<or> (FRESH x. Q x)"
-using P Q by (rule FRESH_binop_iff)
-
-
-section {* An example of a function without finite support *}
-
-primrec
-  nat_of :: "atom \<Rightarrow> nat"
-where
-  "nat_of (Atom s n) = n"
-
-lemma atom_eq_iff:
-  fixes a b :: atom
-  shows "a = b \<longleftrightarrow> sort_of a = sort_of b \<and> nat_of a = nat_of b"
-  by (induct a, induct b, simp)
-
-lemma not_fresh_nat_of:
-  shows "\<not> a \<sharp> nat_of"
-unfolding fresh_def supp_def
-proof (clarsimp)
-  assume "finite {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"
-  hence "finite ({a} \<union> {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of})"
-    by simp
-  then obtain b where
-    b1: "b \<noteq> a" and
-    b2: "sort_of b = sort_of a" and
-    b3: "(a \<rightleftharpoons> b) \<bullet> nat_of = nat_of"
-    by (rule obtain_atom) auto
-  have "nat_of a = (a \<rightleftharpoons> b) \<bullet> (nat_of a)" by (simp add: permute_nat_def)
-  also have "\<dots> = ((a \<rightleftharpoons> b) \<bullet> nat_of) ((a \<rightleftharpoons> b) \<bullet> a)" by (simp add: permute_fun_app_eq)
-  also have "\<dots> = nat_of ((a \<rightleftharpoons> b) \<bullet> a)" using b3 by simp
-  also have "\<dots> = nat_of b" using b2 by simp
-  finally have "nat_of a = nat_of b" by simp
-  with b2 have "a = b" by (simp add: atom_eq_iff)
-  with b1 show "False" by simp
-qed
-
-lemma supp_nat_of:
-  shows "supp nat_of = UNIV"
-  using not_fresh_nat_of [unfolded fresh_def] by auto
-
-
-section {* Support for sets of atoms *}
-
-lemma supp_finite_atom_set:
-  fixes S::"atom set"
-  assumes "finite S"
-  shows "supp S = S"
-  apply(rule finite_supp_unique)
-  apply(simp add: supports_def)
-  apply(simp add: swap_set_not_in)
-  apply(rule assms)
-  apply(simp add: swap_set_in)
-done
-
-
-(*
-lemma supp_infinite:
-  fixes S::"atom set"
-  assumes asm: "finite (UNIV - S)"
-  shows "(supp S) = (UNIV - S)"
-apply(rule finite_supp_unique)
-apply(auto simp add: supports_def permute_set_eq swap_atom)[1]
-apply(rule asm)
-apply(auto simp add: permute_set_eq swap_atom)[1]
-done
-
-lemma supp_infinite_coinfinite:
-  fixes S::"atom set"
-  assumes asm1: "infinite S"
-  and     asm2: "infinite (UNIV-S)"
-  shows "(supp S) = (UNIV::atom set)"
-*)
-
-
-end
\ No newline at end of file
--- a/Quot/Nominal/Perm.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,147 +0,0 @@
-theory Perm
-imports "Nominal2_Atoms"
-begin
-
-ML {*
-  open Datatype_Aux; (* typ_of_dtyp, DtRec, ... *)
-  fun permute ty = Const (@{const_name permute}, @{typ perm} --> ty --> ty);
-  val minus_perm = Const (@{const_name minus}, @{typ perm} --> @{typ perm});
-*}
-
-ML {*
-fun prove_perm_empty lthy induct perm_def perm_frees =
-let
-  val perm_types = map fastype_of perm_frees;
-  val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types);
-  val gl =
-    HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
-      (map (fn ((perm, T), x) => HOLogic.mk_eq
-          (perm $ @{term "0 :: perm"} $ Free (x, T),
-           Free (x, T)))
-       (perm_frees ~~
-        map body_type perm_types ~~ perm_indnames)));
-  fun tac _ =
-    EVERY [
-      indtac induct perm_indnames 1,
-      ALLGOALS (asm_full_simp_tac (HOL_ss addsimps (@{thm permute_zero} :: perm_def)))
-    ];
-in
-  split_conj_thm (Goal.prove lthy perm_indnames [] gl tac)
-end;
-*}
-
-ML {*
-fun prove_perm_append lthy induct perm_def perm_frees =
-let
-  val add_perm = @{term "op + :: (perm \<Rightarrow> perm \<Rightarrow> perm)"}
-  val pi1 = Free ("pi1", @{typ perm});
-  val pi2 = Free ("pi2", @{typ perm});
-  val perm_types = map fastype_of perm_frees
-  val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types);
-  val gl =
-    (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj
-      (map (fn ((perm, T), x) =>
-          let
-            val lhs = perm $ (add_perm $ pi1 $ pi2) $ Free (x, T)
-            val rhs = perm $ pi1 $ (perm $ pi2 $ Free (x, T))
-          in HOLogic.mk_eq (lhs, rhs)
-          end)
-        (perm_frees ~~ map body_type perm_types ~~ perm_indnames))))
-  fun tac _ =
-    EVERY [
-      indtac induct perm_indnames 1,
-      ALLGOALS (asm_full_simp_tac (HOL_ss addsimps (@{thm permute_plus} :: perm_def)))
-    ]
-in
-  split_conj_thm (Goal.prove lthy ("pi1" :: "pi2" :: perm_indnames) [] gl tac)
-end;
-*}
-
-ML {*
-(* TODO: full_name can be obtained from new_type_names with Datatype *)
-fun define_raw_perms new_type_names full_tnames thy =
-let
-  val {descr, induct, ...} = Datatype.the_info thy (hd full_tnames);
-  (* TODO: [] should be the sorts that we'll take from the specification *)
-  val sorts = [];
-  fun nth_dtyp i = typ_of_dtyp descr sorts (DtRec i);
-  val perm_names' = Datatype_Prop.indexify_names (map (fn (i, _) =>
-    "permute_" ^ name_of_typ (nth_dtyp i)) descr);
-  val perm_types = map (fn (i, _) =>
-    let val T = nth_dtyp i
-    in @{typ perm} --> T --> T end) descr;
-  val perm_names_types' = perm_names' ~~ perm_types;
-  val pi = Free ("pi", @{typ perm});
-  fun perm_eq_constr i (cname, dts) =
-    let
-      val Ts = map (typ_of_dtyp descr sorts) dts;
-      val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts);
-      val args = map Free (names ~~ Ts);
-      val c = Const (cname, Ts ---> (nth_dtyp i));
-      fun perm_arg (dt, x) =
-        let val T = type_of x
-        in
-          if is_rec_type dt then
-            let val (Us, _) = strip_type T
-            in list_abs (map (pair "x") Us,
-              Free (nth perm_names_types' (body_index dt)) $ pi $
-                list_comb (x, map (fn (i, U) =>
-                  (permute U) $ (minus_perm $ pi) $ Bound i)
-                  ((length Us - 1 downto 0) ~~ Us)))
-            end
-          else (permute T) $ pi $ x
-        end;
-    in
-      (Attrib.empty_binding, HOLogic.mk_Trueprop (HOLogic.mk_eq
-        (Free (nth perm_names_types' i) $
-           Free ("pi", @{typ perm}) $ list_comb (c, args),
-         list_comb (c, map perm_arg (dts ~~ args)))))
-    end;
-    fun perm_eq (i, (_, _, constrs)) = map (perm_eq_constr i) constrs;
-    val perm_eqs = maps perm_eq descr;
-    val lthy =
-      Theory_Target.instantiation (full_tnames, [], @{sort pt}) thy;
-    (* TODO: Use the version of prmrec that gives the names explicitely. *)
-    val ((_, perm_ldef), lthy') =
-      Primrec.add_primrec
-        (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs lthy;
-    val perm_frees =
-      (distinct (op =)) (map (fst o strip_comb o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o prop_of) perm_ldef);
-    val perm_empty_thms = List.take (prove_perm_empty lthy' induct perm_ldef perm_frees, length new_type_names);
-    val perm_append_thms = List.take (prove_perm_append lthy' induct perm_ldef perm_frees, length new_type_names)
-    val perms_name = space_implode "_" perm_names'
-    val perms_zero_bind = Binding.name (perms_name ^ "_zero")
-    val perms_append_bind = Binding.name (perms_name ^ "_append")
-    fun tac _ perm_thms =
-      (Class.intro_classes_tac []) THEN (ALLGOALS (
-        simp_tac (HOL_ss addsimps perm_thms
-      )));
-    fun morphism phi = map (Morphism.thm phi);
-  in
-  lthy'
-  |> snd o (Local_Theory.note ((perms_zero_bind, []), perm_empty_thms))
-  |> snd o (Local_Theory.note ((perms_append_bind, []), perm_append_thms))
-  |> Class_Target.prove_instantiation_exit_result morphism tac (perm_empty_thms @ perm_append_thms)
-  end
-
-*}
-
-(* Test
-atom_decl name
-
-datatype rtrm1 =
-  rVr1 "name"
-| rAp1 "rtrm1" "rtrm1 list"
-| rLm1 "name" "rtrm1"
-| rLt1 "bp" "rtrm1" "rtrm1"
-and bp =
-  BUnit
-| BVr "name"
-| BPr "bp" "bp"
-
-
-setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Perm.rtrm1", "Perm.bp"] *}
-print_theorems
-*)
-
-end
--- a/Quot/Nominal/Rsp.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,118 +0,0 @@
-theory Rsp
-imports Abs
-begin
-
-ML {*
-fun define_quotient_type args tac ctxt =
-let
-  val mthd = Method.SIMPLE_METHOD tac
-  val mthdt = Method.Basic (fn _ => mthd)
-  val bymt = Proof.global_terminal_proof (mthdt, NONE)
-in
-  bymt (Quotient_Type.quotient_type args ctxt)
-end
-*}
-
-ML {*
-fun const_rsp lthy const =
-let
-  val nty = fastype_of (Quotient_Term.quotient_lift_const ("", const) lthy)
-  val rel = Quotient_Term.equiv_relation_chk lthy (fastype_of const, nty);
-in
-  HOLogic.mk_Trueprop (rel $ const $ const)
-end
-*}
-
-(* Replaces bounds by frees and meta implications by implications *)
-ML {*
-fun prepare_goal trm =
-let
-  val vars = strip_all_vars trm
-  val fs = rev (map Free vars)
-  val (fixes, no_alls) = ((map fst vars), subst_bounds (fs, (strip_all_body trm)))
-  val prems = map HOLogic.dest_Trueprop (Logic.strip_imp_prems no_alls)
-  val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl no_alls)
-in
-  (fixes, fold (curry HOLogic.mk_imp) prems concl)
-end
-*}
-
-ML {*
-fun get_rsp_goal thy trm =
-let
-  val goalstate = Goal.init (cterm_of thy trm);
-  val tac = REPEAT o rtac @{thm fun_rel_id};
-in
-  case (SINGLE (tac 1) goalstate) of
-    NONE => error "rsp_goal failed"
-  | SOME th => prepare_goal (term_of (cprem_of th 1))
-end
-*}
-
-ML {*
-fun repeat_mp thm = repeat_mp (mp OF [thm]) handle THM _ => thm
-*}
-
-ML {*
-fun prove_const_rsp bind consts tac ctxt =
-let
-  val rsp_goals = map (const_rsp ctxt) consts
-  val thy = ProofContext.theory_of ctxt
-  val (fixed, user_goals) = split_list (map (get_rsp_goal thy) rsp_goals)
-  val fixed' = distinct (op =) (flat fixed)
-  val user_goal = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj user_goals)
-  val user_thm = Goal.prove ctxt fixed' [] user_goal tac
-  val user_thms = map repeat_mp (HOLogic.conj_elims user_thm)
-  fun tac _ = (REPEAT o rtac @{thm fun_rel_id} THEN' resolve_tac user_thms THEN_ALL_NEW atac) 1
-  val rsp_thms = map (fn gl => Goal.prove ctxt [] [] gl tac) rsp_goals
-in
-   ctxt
-|> snd o Local_Theory.note 
-  ((Binding.empty, [Attrib.internal (fn _ => Quotient_Info.rsp_rules_add)]), rsp_thms)
-|> snd o Local_Theory.note ((bind, []), user_thms)
-end
-*}
-
-
-
-ML {*
-fun fvbv_rsp_tac induct fvbv_simps =
-  ((((rtac impI THEN' etac induct) ORELSE' rtac induct) THEN_ALL_NEW
-  (TRY o rtac @{thm TrueI})) THEN_ALL_NEW asm_full_simp_tac
-  (HOL_ss addsimps (@{thm alpha_gen} :: fvbv_simps)))
-*}
-
-ML {*
-fun constr_rsp_tac inj rsp equivps =
-let
-  val reflps = map (fn x => @{thm equivp_reflp} OF [x]) equivps
-in
-  REPEAT o rtac impI THEN'
-  simp_tac (HOL_ss addsimps inj) THEN'
-  (TRY o REPEAT_ALL_NEW (CHANGED o rtac conjI)) THEN_ALL_NEW
-  (asm_simp_tac HOL_ss THEN_ALL_NEW (
-   rtac @{thm exI[of _ "0 :: perm"]} THEN'
-   asm_full_simp_tac (HOL_ss addsimps (rsp @ reflps @
-     @{thms alpha_gen fresh_star_def fresh_zero_perm permute_zero ball_triv}))
-  ))
-end
-*}
-
-(* Testing code
-local_setup {* prove_const_rsp @{binding fv_rtrm2_rsp} [@{term rbv2}]
-  (fn _ => fv_rsp_tac @{thm alpha_rtrm2_alpha_rassign.inducts(2)} @{thms fv_rtrm2_fv_rassign.simps} 1) *}*)
-
-(*ML {*
-  val rsp_goals = map (const_rsp @{context}) [@{term rbv2}]
-  val (fixed, user_goals) = split_list (map (get_rsp_goal @{theory}) rsp_goals)
-  val fixed' = distinct (op =) (flat fixed)
-  val user_goal = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj user_goals)
-*}
-prove ug: {* user_goal *}
-ML_prf {*
-val induct = @{thm alpha_rtrm2_alpha_rassign.inducts(2)}
-val fv_simps = @{thms rbv2.simps}
-*} 
-*)
-
-end
--- a/Quot/Nominal/Terms.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1043 +0,0 @@
-theory Terms
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Abs" "Perm" "Fv" "Rsp" "../../Attic/Prove"
-begin
-
-atom_decl name
-
-text {* primrec seems to be genarally faster than fun *}
-
-section {*** lets with binding patterns ***}
-
-datatype rtrm1 =
-  rVr1 "name"
-| rAp1 "rtrm1" "rtrm1"
-| rLm1 "name" "rtrm1"        --"name is bound in trm1"
-| rLt1 "bp" "rtrm1" "rtrm1"   --"all variables in bp are bound in the 2nd trm1"
-and bp =
-  BUnit
-| BVr "name"
-| BPr "bp" "bp"
-
-print_theorems
-
-(* to be given by the user *)
-
-primrec 
-  bv1
-where
-  "bv1 (BUnit) = {}"
-| "bv1 (BVr x) = {atom x}"
-| "bv1 (BPr bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)"
-
-setup {* snd o define_raw_perms ["rtrm1", "bp"] ["Terms.rtrm1", "Terms.bp"] *}
-thm permute_rtrm1_permute_bp.simps
-
-local_setup {*
-  snd o define_fv_alpha "Terms.rtrm1"
-  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[(SOME @{term bv1}, 0)], [], [(SOME @{term bv1}, 0)]]],
-  [[], [[]], [[], []]]] *}
-
-notation
-  alpha_rtrm1 ("_ \<approx>1 _" [100, 100] 100) and
-  alpha_bp ("_ \<approx>1b _" [100, 100] 100)
-thm alpha_rtrm1_alpha_bp.intros
-thm fv_rtrm1_fv_bp.simps
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_inj}, []), (build_alpha_inj @{thms alpha_rtrm1_alpha_bp.intros} @{thms rtrm1.distinct rtrm1.inject bp.distinct bp.inject} @{thms alpha_rtrm1.cases alpha_bp.cases} ctxt)) ctxt)) *}
-thm alpha1_inj
-
-lemma alpha_bp_refl: "alpha_bp a a"
-apply induct
-apply (simp_all  add: alpha1_inj)
-done
-
-lemma alpha_bp_eq_eq: "alpha_bp a b = (a = b)"
-apply rule
-apply (induct a b rule: alpha_rtrm1_alpha_bp.inducts(2))
-apply (simp_all add: alpha_bp_refl)
-done
-
-lemma alpha_bp_eq: "alpha_bp = (op =)"
-apply (rule ext)+
-apply (rule alpha_bp_eq_eq)
-done
-
-lemma bv1_eqvt[eqvt]:
-  shows "(pi \<bullet> bv1 x) = bv1 (pi \<bullet> x)"
-  apply (induct x)
-  apply (simp_all add: atom_eqvt eqvts)
-  done
-
-lemma fv_rtrm1_eqvt[eqvt]:
-    "(pi\<bullet>fv_rtrm1 t) = fv_rtrm1 (pi\<bullet>t)"
-    "(pi\<bullet>fv_bp b) = fv_bp (pi\<bullet>b)"
-  apply (induct t and b)
-  apply (simp_all add: insert_eqvt atom_eqvt empty_eqvt union_eqvt Diff_eqvt bv1_eqvt)
-  done
-
-lemma alpha1_eqvt:
-  "t \<approx>1 s \<Longrightarrow> (pi \<bullet> t) \<approx>1 (pi \<bullet> s)"
-  "alpha_bp a b \<Longrightarrow> alpha_bp (pi \<bullet> a) (pi \<bullet> b)"
-  apply (induct t s and a b rule: alpha_rtrm1_alpha_bp.inducts)
-  apply (simp_all add:eqvts alpha1_inj)
-  apply (erule exE)
-  apply (rule_tac x="pi \<bullet> pia" in exI)
-  apply (simp add: alpha_gen)
-  apply(erule conjE)+
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
-  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt empty_eqvt fv_rtrm1_eqvt)
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
-  apply(simp add: atom_eqvt Diff_eqvt fv_rtrm1_eqvt insert_eqvt empty_eqvt)
-  apply(simp add: permute_eqvt[symmetric])
-  apply (erule exE)
-  apply (erule exE)
-  apply (rule conjI)
-  apply (rule_tac x="pi \<bullet> pia" in exI)
-  apply (simp add: alpha_gen)
-  apply(erule conjE)+
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
-  apply(simp add: fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
-  apply(simp add: atom_eqvt fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
-  apply(simp add: permute_eqvt[symmetric])
-  apply (rule_tac x="pi \<bullet> piaa" in exI)
-  apply (simp add: alpha_gen)
-  apply(erule conjE)+
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1])
-  apply(simp add: fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
-  apply(rule conjI)
-  apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1])
-  apply(simp add: atom_eqvt fv_rtrm1_eqvt Diff_eqvt bv1_eqvt)
-  apply(simp add: permute_eqvt[symmetric])
-  done
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha1_equivp}, []),
-  (build_equivps [@{term alpha_rtrm1}, @{term alpha_bp}] @{thm rtrm1_bp.induct} @{thm alpha_rtrm1_alpha_bp.induct} @{thms rtrm1.inject bp.inject} @{thms alpha1_inj} @{thms rtrm1.distinct bp.distinct} @{thms alpha_rtrm1.cases alpha_bp.cases} @{thms alpha1_eqvt} ctxt)) ctxt)) *}
-thm alpha1_equivp
-
-local_setup  {* define_quotient_type [(([], @{binding trm1}, NoSyn), (@{typ rtrm1}, @{term alpha_rtrm1}))]
-  (rtac @{thm alpha1_equivp(1)} 1) *}
-
-local_setup {*
-(fn ctxt => ctxt
- |> snd o (Quotient_Def.quotient_lift_const ("Vr1", @{term rVr1}))
- |> snd o (Quotient_Def.quotient_lift_const ("Ap1", @{term rAp1}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lm1", @{term rLm1}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lt1", @{term rLt1}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_trm1", @{term fv_rtrm1})))
-*}
-print_theorems
-
-thm alpha_rtrm1_alpha_bp.induct
-local_setup {* prove_const_rsp @{binding fv_rtrm1_rsp} [@{term fv_rtrm1}]
-  (fn _ => fvbv_rsp_tac @{thm alpha_rtrm1_alpha_bp.inducts(1)} @{thms fv_rtrm1_fv_bp.simps} 1) *}
-local_setup {* prove_const_rsp @{binding rVr1_rsp} [@{term rVr1}]
-  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding rAp1_rsp} [@{term rAp1}]
-  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding rLm1_rsp} [@{term rLm1}]
-  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding rLt1_rsp} [@{term rLt1}]
-  (fn _ => constr_rsp_tac @{thms alpha1_inj} @{thms fv_rtrm1_rsp} @{thms alpha1_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding permute_rtrm1_rsp} [@{term "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"}]
-  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha1_eqvt}) 1) *}
-
-lemmas trm1_bp_induct = rtrm1_bp.induct[quot_lifted]
-lemmas trm1_bp_inducts = rtrm1_bp.inducts[quot_lifted]
-
-instantiation trm1 and bp :: pt
-begin
-
-quotient_definition
-  "permute_trm1 :: perm \<Rightarrow> trm1 \<Rightarrow> trm1"
-is
-  "permute :: perm \<Rightarrow> rtrm1 \<Rightarrow> rtrm1"
-
-instance by default 
-  (simp_all add: permute_rtrm1_permute_bp_zero[quot_lifted] permute_rtrm1_permute_bp_append[quot_lifted])
-
-end
-
-lemmas
-    permute_trm1 = permute_rtrm1_permute_bp.simps[quot_lifted]
-and fv_trm1 = fv_rtrm1_fv_bp.simps[quot_lifted]
-and fv_trm1_eqvt = fv_rtrm1_eqvt[quot_lifted]
-and alpha1_INJ = alpha1_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
-
-lemma supports:
-  "(supp (atom x)) supports (Vr1 x)"
-  "(supp t \<union> supp s) supports (Ap1 t s)"
-  "(supp (atom x) \<union> supp t) supports (Lm1 x t)"
-  "(supp b \<union> supp t \<union> supp s) supports (Lt1 b t s)"
-  "{} supports BUnit"
-  "(supp (atom x)) supports (BVr x)"
-  "(supp a \<union> supp b) supports (BPr a b)"
-apply(simp_all add: supports_def fresh_def[symmetric] swap_fresh_fresh permute_trm1)
-apply(rule_tac [!] allI)+
-apply(rule_tac [!] impI)
-apply(tactic {* ALLGOALS (REPEAT o etac conjE) *})
-apply(simp_all add: fresh_atom)
-done
-
-lemma rtrm1_bp_fs:
-  "finite (supp (x :: trm1))"
-  "finite (supp (b :: bp))"
-  apply (induct x and b rule: trm1_bp_inducts)
-  apply(tactic {* ALLGOALS (rtac @{thm supports_finite} THEN' resolve_tac @{thms supports}) *})
-  apply(simp_all add: supp_atom)
-  done
-
-instance trm1 :: fs
-apply default
-apply (rule rtrm1_bp_fs(1))
-done
-
-lemma fv_eq_bv: "fv_bp bp = bv1 bp"
-apply(induct bp rule: trm1_bp_inducts(2))
-apply(simp_all)
-done
-
-lemma helper2: "{b. \<forall>pi. pi \<bullet> (a \<rightleftharpoons> b) \<bullet> bp \<noteq> bp} = {}"
-apply auto
-apply (rule_tac x="(x \<rightleftharpoons> a)" in exI)
-apply auto
-done
-
-lemma supp_fv:
-  "supp t = fv_trm1 t"
-  "supp b = fv_bp b"
-apply(induct t and b rule: trm1_bp_inducts)
-apply(simp_all)
-apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1)
-apply(simp only: supp_at_base[simplified supp_def])
-apply(simp add: supp_def permute_trm1 alpha1_INJ fv_trm1)
-apply(simp add: Collect_imp_eq Collect_neg_eq)
-apply(subgoal_tac "supp (Lm1 name rtrm1) = supp (Abs {atom name} rtrm1)")
-apply(simp add: supp_Abs fv_trm1)
-apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt permute_trm1)
-apply(simp add: alpha1_INJ)
-apply(simp add: Abs_eq_iff)
-apply(simp add: alpha_gen.simps)
-apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric])
-apply(subgoal_tac "supp (Lt1 bp rtrm11 rtrm12) = supp(rtrm11) \<union> supp (Abs (bv1 bp) rtrm12)")
-apply(simp add: supp_Abs fv_trm1 fv_eq_bv)
-apply(simp (no_asm) add: supp_def permute_trm1)
-apply(simp add: alpha1_INJ alpha_bp_eq)
-apply(simp add: Abs_eq_iff)
-apply(simp add: alpha_gen)
-apply(simp add: supp_eqvt[symmetric] fv_trm1_eqvt[symmetric] bv1_eqvt fv_eq_bv)
-apply(simp add: Collect_imp_eq Collect_neg_eq fresh_star_def helper2)
-apply(simp (no_asm) add: supp_def permute_set_eq atom_eqvt)
-apply(simp (no_asm) add: supp_def eqvts)
-apply(fold supp_def)
-apply(simp add: supp_at_base)
-apply(simp (no_asm) add: supp_def Collect_imp_eq Collect_neg_eq)
-apply(simp add: Collect_imp_eq[symmetric] Collect_neg_eq[symmetric] supp_def[symmetric])
-done
-
-lemma trm1_supp:
-  "supp (Vr1 x) = {atom x}"
-  "supp (Ap1 t1 t2) = supp t1 \<union> supp t2"
-  "supp (Lm1 x t) = (supp t) - {atom x}"
-  "supp (Lt1 b t s) = supp t \<union> (supp s - bv1 b)"
-by (simp_all add: supp_fv fv_trm1 fv_eq_bv)
-
-lemma trm1_induct_strong:
-  assumes "\<And>name b. P b (Vr1 name)"
-  and     "\<And>rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12\<rbrakk> \<Longrightarrow> P b (Ap1 rtrm11 rtrm12)"
-  and     "\<And>name rtrm1 b. \<lbrakk>\<And>c. P c rtrm1; (atom name) \<sharp> b\<rbrakk> \<Longrightarrow> P b (Lm1 name rtrm1)"
-  and     "\<And>bp rtrm11 rtrm12 b. \<lbrakk>\<And>c. P c rtrm11; \<And>c. P c rtrm12; bv1 bp \<sharp>* b\<rbrakk> \<Longrightarrow> P b (Lt1 bp rtrm11 rtrm12)"
-  shows   "P a rtrma"
-sorry
-
-section {*** lets with single assignments ***}
-
-datatype rtrm2 =
-  rVr2 "name"
-| rAp2 "rtrm2" "rtrm2"
-| rLm2 "name" "rtrm2" --"bind (name) in (rtrm2)"
-| rLt2 "rassign" "rtrm2" --"bind (bv2 rassign) in (rtrm2)"
-and rassign =
-  rAs "name" "rtrm2"
-
-(* to be given by the user *)
-primrec 
-  rbv2
-where
-  "rbv2 (rAs x t) = {atom x}"
-
-setup {* snd o define_raw_perms ["rtrm2", "rassign"] ["Terms.rtrm2", "Terms.rassign"] *}
-
-local_setup {* snd o define_fv_alpha "Terms.rtrm2"
-  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv2}, 0)]]],
-   [[[], []]]] *}
-
-notation
-  alpha_rtrm2 ("_ \<approx>2 _" [100, 100] 100) and
-  alpha_rassign ("_ \<approx>2b _" [100, 100] 100)
-thm alpha_rtrm2_alpha_rassign.intros
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha2_inj}, []), (build_alpha_inj @{thms alpha_rtrm2_alpha_rassign.intros} @{thms rtrm2.distinct rtrm2.inject rassign.distinct rassign.inject} @{thms alpha_rtrm2.cases alpha_rassign.cases} ctxt)) ctxt)) *}
-thm alpha2_inj
-
-lemma alpha2_eqvt:
-  "t \<approx>2 s \<Longrightarrow> (pi \<bullet> t) \<approx>2 (pi \<bullet> s)"
-  "a \<approx>2b b \<Longrightarrow> (pi \<bullet> a) \<approx>2b (pi \<bullet> b)"
-sorry
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha2_equivp}, []),
-  (build_equivps [@{term alpha_rtrm2}, @{term alpha_rassign}] @{thm rtrm2_rassign.induct} @{thm alpha_rtrm2_alpha_rassign.induct} @{thms rtrm2.inject rassign.inject} @{thms alpha2_inj} @{thms rtrm2.distinct rassign.distinct} @{thms alpha_rtrm2.cases alpha_rassign.cases} @{thms alpha2_eqvt} ctxt)) ctxt)) *}
-thm alpha2_equivp
-
-local_setup  {* define_quotient_type 
-  [(([], @{binding trm2}, NoSyn), (@{typ rtrm2}, @{term alpha_rtrm2})),
-   (([], @{binding assign}, NoSyn), (@{typ rassign}, @{term alpha_rassign}))]
-  ((rtac @{thm alpha2_equivp(1)} 1) THEN (rtac @{thm alpha2_equivp(2)}) 1) *}
-
-local_setup {*
-(fn ctxt => ctxt
- |> snd o (Quotient_Def.quotient_lift_const ("Vr2", @{term rVr2}))
- |> snd o (Quotient_Def.quotient_lift_const ("Ap2", @{term rAp2}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lm2", @{term rLm2}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lt2", @{term rLt2}))
- |> snd o (Quotient_Def.quotient_lift_const ("As", @{term rAs}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_trm2", @{term fv_rtrm2}))
- |> snd o (Quotient_Def.quotient_lift_const ("bv2", @{term rbv2})))
-*}
-print_theorems
-
-local_setup {* prove_const_rsp @{binding fv_rtrm2_rsp} [@{term fv_rtrm2}, @{term fv_rassign}]
-  (fn _ => fvbv_rsp_tac @{thm alpha_rtrm2_alpha_rassign.induct} @{thms fv_rtrm2_fv_rassign.simps} 1) *}
-local_setup {* prove_const_rsp @{binding rbv2_rsp} [@{term rbv2}]
-  (fn _ => fvbv_rsp_tac @{thm alpha_rtrm2_alpha_rassign.inducts(2)} @{thms rbv2.simps} 1) *}
-local_setup {* prove_const_rsp @{binding rVr2_rsp} [@{term rVr2}]
-  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp} @{thms alpha2_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding rAp2_rsp} [@{term rAp2}]
-  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp} @{thms alpha2_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding rLm2_rsp} [@{term rLm2}]
-  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp} @{thms alpha2_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding rLt2_rsp} [@{term rLt2}]
-  (fn _ => constr_rsp_tac @{thms alpha2_inj} @{thms fv_rtrm2_rsp rbv2_rsp} @{thms alpha2_equivp} 1) *}
-local_setup {* prove_const_rsp @{binding permute_rtrm2_rsp} [@{term "permute :: perm \<Rightarrow> rtrm2 \<Rightarrow> rtrm2"}]
-  (fn _ => asm_simp_tac (HOL_ss addsimps @{thms alpha2_eqvt}) 1) *}
-
-
-section {*** lets with many assignments ***}
-
-datatype rtrm3 =
-  rVr3 "name"
-| rAp3 "rtrm3" "rtrm3"
-| rLm3 "name" "rtrm3" --"bind (name) in (trm3)"
-| rLt3 "rassigns" "rtrm3" --"bind (bv3 assigns) in (trm3)"
-and rassigns =
-  rANil
-| rACons "name" "rtrm3" "rassigns"
-
-(* to be given by the user *)
-primrec 
-  bv3
-where
-  "bv3 rANil = {}"
-| "bv3 (rACons x t as) = {atom x} \<union> (bv3 as)"
-
-setup {* snd o define_raw_perms ["rtrm3", "rassigns"] ["Terms.rtrm3", "Terms.rassigns"] *}
-
-local_setup {* snd o define_fv_alpha "Terms.rtrm3"
-  [[[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term bv3}, 0)]]],
-   [[], [[], [], []]]] *}
-print_theorems
-
-notation
-  alpha_rtrm3 ("_ \<approx>3 _" [100, 100] 100) and
-  alpha_rassigns ("_ \<approx>3a _" [100, 100] 100)
-thm alpha_rtrm3_alpha_rassigns.intros
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha3_inj}, []), (build_alpha_inj @{thms alpha_rtrm3_alpha_rassigns.intros} @{thms rtrm3.distinct rtrm3.inject rassigns.distinct rassigns.inject} @{thms alpha_rtrm3.cases alpha_rassigns.cases} ctxt)) ctxt)) *}
-thm alpha3_inj
-
-lemma alpha3_eqvt:
-  "t \<approx>3 s \<Longrightarrow> (pi \<bullet> t) \<approx>3 (pi \<bullet> s)"
-  "a \<approx>3a b \<Longrightarrow> (pi \<bullet> a) \<approx>3a (pi \<bullet> b)"
-sorry
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha3_equivp}, []),
-  (build_equivps [@{term alpha_rtrm3}, @{term alpha_rassigns}] @{thm rtrm3_rassigns.induct} @{thm alpha_rtrm3_alpha_rassigns.induct} @{thms rtrm3.inject rassigns.inject} @{thms alpha3_inj} @{thms rtrm3.distinct rassigns.distinct} @{thms alpha_rtrm3.cases alpha_rassigns.cases} @{thms alpha3_eqvt} ctxt)) ctxt)) *}
-thm alpha3_equivp
-
-quotient_type
-  trm3 = rtrm3 / alpha_rtrm3
-and
-  assigns = rassigns / alpha_rassigns
-  by (rule alpha3_equivp(1)) (rule alpha3_equivp(2))
-
-
-section {*** lam with indirect list recursion ***}
-
-datatype rtrm4 =
-  rVr4 "name"
-| rAp4 "rtrm4" "rtrm4 list"
-| rLm4 "name" "rtrm4"  --"bind (name) in (trm)"
-print_theorems
-
-thm rtrm4.recs
-
-(* there cannot be a clause for lists, as *)
-(* permutations are  already defined in Nominal (also functions, options, and so on) *)
-setup {* snd o define_raw_perms ["rtrm4"] ["Terms.rtrm4"] *}
-
-(* "repairing" of the permute function *)
-lemma repaired:
-  fixes ts::"rtrm4 list"
-  shows "permute_rtrm4_list p ts = p \<bullet> ts"
-  apply(induct ts)
-  apply(simp_all)
-  done
-
-thm permute_rtrm4_permute_rtrm4_list.simps
-thm permute_rtrm4_permute_rtrm4_list.simps[simplified repaired]
-
-local_setup {* snd o define_fv_alpha "Terms.rtrm4" [
-  [[[]], [[], []], [[(NONE, 0)], [(NONE, 0)]]], [[], [[], []]]  ] *}
-print_theorems
-
-notation
-  alpha_rtrm4 ("_ \<approx>4 _" [100, 100] 100) and
-  alpha_rtrm4_list ("_ \<approx>4l _" [100, 100] 100)
-thm alpha_rtrm4_alpha_rtrm4_list.intros
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha4_inj}, []), (build_alpha_inj @{thms alpha_rtrm4_alpha_rtrm4_list.intros} @{thms rtrm4.distinct rtrm4.inject list.distinct list.inject} @{thms alpha_rtrm4.cases alpha_rtrm4_list.cases} ctxt)) ctxt)) *}
-thm alpha4_inj
-
-lemma alpha4_eqvt:
-  "t \<approx>4 s \<Longrightarrow> (pi \<bullet> t) \<approx>4 (pi \<bullet> s)"
-  "a \<approx>4l b \<Longrightarrow> (pi \<bullet> a) \<approx>4l (pi \<bullet> b)"
-sorry
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha4_equivp}, []),
-  (build_equivps [@{term alpha_rtrm4}, @{term alpha_rtrm4_list}] @{thm rtrm4.induct} @{thm alpha_rtrm4_alpha_rtrm4_list.induct} @{thms rtrm4.inject list.inject} @{thms alpha4_inj} @{thms rtrm4.distinct list.distinct} @{thms alpha_rtrm4_list.cases alpha_rtrm4.cases} @{thms alpha4_eqvt} ctxt)) ctxt)) *}
-thm alpha4_equivp
-
-quotient_type 
-  qrtrm4 = rtrm4 / alpha_rtrm4 and
-  qrtrm4list = "rtrm4 list" / alpha_rtrm4_list
-  by (simp_all add: alpha4_equivp)
-
-
-datatype rtrm5 =
-  rVr5 "name"
-| rAp5 "rtrm5" "rtrm5"
-| rLt5 "rlts" "rtrm5" --"bind (bv5 lts) in (rtrm5)"
-and rlts =
-  rLnil
-| rLcons "name" "rtrm5" "rlts"
-
-primrec
-  rbv5
-where
-  "rbv5 rLnil = {}"
-| "rbv5 (rLcons n t ltl) = {atom n} \<union> (rbv5 ltl)"
-
-
-setup {* snd o define_raw_perms ["rtrm5", "rlts"] ["Terms.rtrm5", "Terms.rlts"] *}
-print_theorems
-
-local_setup {* snd o define_fv_alpha "Terms.rtrm5" [
-  [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[], [], []]]  ] *}
-print_theorems
-
-(* Alternate version with additional binding of name in rlts in rLcons *)
-(*local_setup {* snd o define_fv_alpha "Terms.rtrm5" [
-  [[[]], [[], []], [[(SOME @{term rbv5}, 0)], [(SOME @{term rbv5}, 0)]]], [[], [[(NONE,0)], [], [(NONE,0)]]]  ] *}
-print_theorems*)
-
-notation
-  alpha_rtrm5 ("_ \<approx>5 _" [100, 100] 100) and
-  alpha_rlts ("_ \<approx>l _" [100, 100] 100)
-thm alpha_rtrm5_alpha_rlts.intros
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_inj}, []), (build_alpha_inj @{thms alpha_rtrm5_alpha_rlts.intros} @{thms rtrm5.distinct rtrm5.inject rlts.distinct rlts.inject} @{thms alpha_rtrm5.cases alpha_rlts.cases} ctxt)) ctxt)) *}
-thm alpha5_inj
-
-lemma rbv5_eqvt:
-  "pi \<bullet> (rbv5 x) = rbv5 (pi \<bullet> x)"
-sorry
-
-lemma fv_rtrm5_eqvt:
-  "pi \<bullet> (fv_rtrm5 x) = fv_rtrm5 (pi \<bullet> x)"
-sorry
-
-lemma fv_rlts_eqvt:
-  "pi \<bullet> (fv_rlts x) = fv_rlts (pi \<bullet> x)"
-sorry
-
-lemma alpha5_eqvt:
-  "xa \<approx>5 y \<Longrightarrow> (x \<bullet> xa) \<approx>5 (x \<bullet> y)"
-  "xb \<approx>l ya \<Longrightarrow> (x \<bullet> xb) \<approx>l (x \<bullet> ya)"
-  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts)
-  apply (simp_all add: alpha5_inj)
-  apply (erule exE)+
-  apply(unfold alpha_gen)
-  apply (erule conjE)+
-  apply (rule conjI)
-  apply (rule_tac x="x \<bullet> pi" in exI)
-  apply (rule conjI)
-  apply(rule_tac ?p1="- x" in permute_eq_iff[THEN iffD1])
-  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rlts_eqvt)
-  apply(rule conjI)
-  apply(rule_tac ?p1="- x" in fresh_star_permute_iff[THEN iffD1])
-  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rlts_eqvt)
-  apply (subst permute_eqvt[symmetric])
-  apply (simp)
-  apply (rule_tac x="x \<bullet> pia" in exI)
-  apply (rule conjI)
-  apply(rule_tac ?p1="- x" in permute_eq_iff[THEN iffD1])
-  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rtrm5_eqvt)
-  apply(rule conjI)
-  apply(rule_tac ?p1="- x" in fresh_star_permute_iff[THEN iffD1])
-  apply(simp add: atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt rbv5_eqvt fv_rtrm5_eqvt)
-  apply (subst permute_eqvt[symmetric])
-  apply (simp)
-  done
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha5_equivp}, []),
-  (build_equivps [@{term alpha_rtrm5}, @{term alpha_rlts}] @{thm rtrm5_rlts.induct} @{thm alpha_rtrm5_alpha_rlts.induct} @{thms rtrm5.inject rlts.inject} @{thms alpha5_inj} @{thms rtrm5.distinct rlts.distinct} @{thms alpha_rtrm5.cases alpha_rlts.cases} @{thms alpha5_eqvt} ctxt)) ctxt)) *}
-thm alpha5_equivp
-
-quotient_type
-  trm5 = rtrm5 / alpha_rtrm5
-and
-  lts = rlts / alpha_rlts
-  by (auto intro: alpha5_equivp)
-
-local_setup {*
-(fn ctxt => ctxt
- |> snd o (Quotient_Def.quotient_lift_const ("Vr5", @{term rVr5}))
- |> snd o (Quotient_Def.quotient_lift_const ("Ap5", @{term rAp5}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lt5", @{term rLt5}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lnil", @{term rLnil}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lcons", @{term rLcons}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_trm5", @{term fv_rtrm5}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_lts", @{term fv_rlts}))
- |> snd o (Quotient_Def.quotient_lift_const ("bv5", @{term rbv5})))
-*}
-print_theorems
-
-lemma alpha5_rfv:
-  "(t \<approx>5 s \<Longrightarrow> fv_rtrm5 t = fv_rtrm5 s)"
-  "(l \<approx>l m \<Longrightarrow> fv_rlts l = fv_rlts m)"
-  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts)
-  apply(simp_all add: alpha_gen)
-  done
-
-lemma bv_list_rsp:
-  shows "x \<approx>l y \<Longrightarrow> rbv5 x = rbv5 y"
-  apply(induct rule: alpha_rtrm5_alpha_rlts.inducts(2))
-  apply(simp_all)
-  done
-
-lemma [quot_respect]:
-  "(alpha_rlts ===> op =) fv_rlts fv_rlts"
-  "(alpha_rtrm5 ===> op =) fv_rtrm5 fv_rtrm5"
-  "(alpha_rlts ===> op =) rbv5 rbv5"
-  "(op = ===> alpha_rtrm5) rVr5 rVr5"
-  "(alpha_rtrm5 ===> alpha_rtrm5 ===> alpha_rtrm5) rAp5 rAp5"
-  "(alpha_rlts ===> alpha_rtrm5 ===> alpha_rtrm5) rLt5 rLt5"
-  "(op = ===> alpha_rtrm5 ===> alpha_rlts ===> alpha_rlts) rLcons rLcons"
-  "(op = ===> alpha_rtrm5 ===> alpha_rtrm5) permute permute"
-  "(op = ===> alpha_rlts ===> alpha_rlts) permute permute"
-  apply (simp_all add: alpha5_inj alpha5_rfv alpha5_eqvt bv_list_rsp)
-  apply (clarify) apply (rule conjI)
-  apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
-  apply (rule_tac x="0" in exI) apply (simp add: fresh_star_def fresh_zero_perm alpha_gen alpha5_rfv)
-  done
-
-lemma
-  shows "(alpha_rlts ===> op =) rbv5 rbv5"
-  by (simp add: bv_list_rsp)
-
-lemmas trm5_lts_inducts = rtrm5_rlts.inducts[quot_lifted]
-
-instantiation trm5 and lts :: pt
-begin
-
-quotient_definition
-  "permute_trm5 :: perm \<Rightarrow> trm5 \<Rightarrow> trm5"
-is
-  "permute :: perm \<Rightarrow> rtrm5 \<Rightarrow> rtrm5"
-
-quotient_definition
-  "permute_lts :: perm \<Rightarrow> lts \<Rightarrow> lts"
-is
-  "permute :: perm \<Rightarrow> rlts \<Rightarrow> rlts"
-
-instance by default
-  (simp_all add: permute_rtrm5_permute_rlts_zero[quot_lifted] permute_rtrm5_permute_rlts_append[quot_lifted])
-
-end
-
-lemmas
-    permute_trm5_lts = permute_rtrm5_permute_rlts.simps[quot_lifted]
-and alpha5_INJ = alpha5_inj[unfolded alpha_gen, quot_lifted, folded alpha_gen]
-and bv5[simp] = rbv5.simps[quot_lifted]
-and fv_trm5_lts[simp] = fv_rtrm5_fv_rlts.simps[quot_lifted]
-
-lemma lets_ok:
-  "(Lt5 (Lcons x (Vr5 x) Lnil) (Vr5 x)) = (Lt5 (Lcons y (Vr5 y) Lnil) (Vr5 y))"
-apply (subst alpha5_INJ)
-apply (rule conjI)
-apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
-apply (simp only: alpha_gen)
-apply (simp add: permute_trm5_lts fresh_star_def)
-apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
-apply (simp only: alpha_gen)
-apply (simp add: permute_trm5_lts fresh_star_def)
-done
-
-lemma lets_ok2:
-  "(Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) =
-   (Lt5 (Lcons y (Vr5 y) (Lcons x (Vr5 x) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
-apply (subst alpha5_INJ)
-apply (rule conjI)
-apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
-apply (simp only: alpha_gen)
-apply (simp add: permute_trm5_lts fresh_star_def)
-apply (rule_tac x="0 :: perm" in exI)
-apply (simp only: alpha_gen)
-apply (simp add: permute_trm5_lts fresh_star_def)
-done
-
-
-lemma lets_not_ok1:
-  "x \<noteq> y \<Longrightarrow> (Lt5 (Lcons x (Vr5 x) (Lcons y (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
-             (Lt5 (Lcons y (Vr5 x) (Lcons x (Vr5 y) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
-apply (simp add: alpha5_INJ(3) alpha_gen)
-apply (simp add: permute_trm5_lts fresh_star_def alpha5_INJ(5) alpha5_INJ(2) alpha5_INJ(1))
-done
-
-lemma distinct_helper:
-  shows "\<not>(rVr5 x \<approx>5 rAp5 y z)"
-  apply auto
-  apply (erule alpha_rtrm5.cases)
-  apply (simp_all only: rtrm5.distinct)
-  done
-
-lemma distinct_helper2:
-  shows "(Vr5 x) \<noteq> (Ap5 y z)"
-  by (lifting distinct_helper)
-
-lemma lets_nok:
-  "x \<noteq> y \<Longrightarrow> x \<noteq> z \<Longrightarrow> z \<noteq> y \<Longrightarrow>
-   (Lt5 (Lcons x (Ap5 (Vr5 z) (Vr5 z)) (Lcons y (Vr5 z) Lnil)) (Ap5 (Vr5 x) (Vr5 y))) \<noteq>
-   (Lt5 (Lcons y (Vr5 z) (Lcons x (Ap5 (Vr5 z) (Vr5 z)) Lnil)) (Ap5 (Vr5 x) (Vr5 y)))"
-apply (simp only: alpha5_INJ(3) alpha5_INJ(5) alpha_gen permute_trm5_lts fresh_star_def)
-apply (simp add: distinct_helper2)
-done
-
-
-(* example with a bn function defined over the type itself *)
-datatype rtrm6 =
-  rVr6 "name"
-| rLm6 "name" "rtrm6" --"bind name in rtrm6"
-| rLt6 "rtrm6" "rtrm6" --"bind (bv6 left) in (right)"
-
-primrec
-  rbv6
-where
-  "rbv6 (rVr6 n) = {}"
-| "rbv6 (rLm6 n t) = {atom n} \<union> rbv6 t"
-| "rbv6 (rLt6 l r) = rbv6 l \<union> rbv6 r"
-
-setup {* snd o define_raw_perms ["rtrm6"] ["Terms.rtrm6"] *}
-print_theorems
-
-local_setup {* snd o define_fv_alpha "Terms.rtrm6" [
-  [[[]], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv6}, 0)]]]] *}
-notation alpha_rtrm6 ("_ \<approx>6 _" [100, 100] 100)
-(* HERE THE RULES DIFFER *)
-thm alpha_rtrm6.intros
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha6_inj}, []), (build_alpha_inj @{thms alpha_rtrm6.intros} @{thms rtrm6.distinct rtrm6.inject} @{thms alpha_rtrm6.cases} ctxt)) ctxt)) *}
-thm alpha6_inj
-
-lemma alpha6_eqvt:
-  "xa \<approx>6 y \<Longrightarrow> (x \<bullet> xa) \<approx>6 (x \<bullet> y)"
-sorry
-
-local_setup {* (fn ctxt => snd (Local_Theory.note ((@{binding alpha6_equivp}, []),
-  (build_equivps [@{term alpha_rtrm6}] @{thm rtrm6.induct} @{thm alpha_rtrm6.induct} @{thms rtrm6.inject} @{thms alpha6_inj} @{thms rtrm6.distinct} @{thms alpha_rtrm6.cases} @{thms alpha6_eqvt} ctxt)) ctxt)) *}
-thm alpha6_equivp
-
-quotient_type
-  trm6 = rtrm6 / alpha_rtrm6
-  by (auto intro: alpha6_equivp)
-
-local_setup {*
-(fn ctxt => ctxt
- |> snd o (Quotient_Def.quotient_lift_const ("Vr6", @{term rVr6}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lm6", @{term rLm6}))
- |> snd o (Quotient_Def.quotient_lift_const ("Lt6", @{term rLt6}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_trm6", @{term fv_rtrm6}))
- |> snd o (Quotient_Def.quotient_lift_const ("bv6", @{term rbv6})))
-*}
-print_theorems
-
-lemma [quot_respect]:
-  "(op = ===> alpha_rtrm6 ===> alpha_rtrm6) permute permute"
-by (auto simp add: alpha6_eqvt)
-
-(* Definitely not true , see lemma below *)
-lemma [quot_respect]:"(alpha_rtrm6 ===> op =) rbv6 rbv6"
-apply simp apply clarify
-apply (erule alpha_rtrm6.induct)
-oops
-
-lemma "(a :: name) \<noteq> b \<Longrightarrow> \<not> (alpha_rtrm6 ===> op =) rbv6 rbv6"
-apply simp
-apply (rule_tac x="rLm6 (a::name) (rVr6 (a :: name))" in  exI)
-apply (rule_tac x="rLm6 (b::name) (rVr6 (b :: name))" in  exI)
-apply simp
-apply (simp add: alpha6_inj)
-apply (rule_tac x="(a \<leftrightarrow> b)" in  exI)
-apply (simp add: alpha_gen fresh_star_def)
-apply (simp add: alpha6_inj)
-done
-
-lemma fv6_rsp: "x \<approx>6 y \<Longrightarrow> fv_rtrm6 x = fv_rtrm6 y"
-apply (induct_tac x y rule: alpha_rtrm6.induct)
-apply simp_all
-apply (erule exE)
-apply (simp_all add: alpha_gen)
-done
-
-lemma [quot_respect]:"(alpha_rtrm6 ===> op =) fv_rtrm6 fv_rtrm6"
-by (simp add: fv6_rsp)
-
-lemma [quot_respect]:
- "(op = ===> alpha_rtrm6) rVr6 rVr6"
- "(op = ===> alpha_rtrm6 ===> alpha_rtrm6) rLm6 rLm6"
-apply auto
-apply (simp_all add: alpha6_inj)
-apply (rule_tac x="0::perm" in exI)
-apply (simp add: alpha_gen fv6_rsp fresh_star_def fresh_zero_perm)
-done
-
-lemma [quot_respect]:
- "(alpha_rtrm6 ===> alpha_rtrm6 ===> alpha_rtrm6) rLt6 rLt6"
-apply auto
-apply (simp_all add: alpha6_inj)
-apply (rule_tac [!] x="0::perm" in exI)
-apply (simp_all add: alpha_gen fresh_star_def fresh_zero_perm)
-(* needs rbv6_rsp *)
-oops
-
-instantiation trm6 :: pt begin
-
-quotient_definition
-  "permute_trm6 :: perm \<Rightarrow> trm6 \<Rightarrow> trm6"
-is
-  "permute :: perm \<Rightarrow> rtrm6 \<Rightarrow> rtrm6"
-
-instance
-apply default
-sorry
-end
-
-lemma lifted_induct:
-"\<lbrakk>x1 = x2; \<And>name namea. name = namea \<Longrightarrow> P (Vr6 name) (Vr6 namea);
- \<And>name rtrm6 namea rtrm6a.
-    \<lbrakk>True;
-     \<exists>pi. fv_trm6 rtrm6 - {atom name} = fv_trm6 rtrm6a - {atom namea} \<and>
-          (fv_trm6 rtrm6 - {atom name}) \<sharp>* pi \<and> pi \<bullet> rtrm6 = rtrm6a \<and> P (pi \<bullet> rtrm6) rtrm6a\<rbrakk>
-    \<Longrightarrow> P (Lm6 name rtrm6) (Lm6 namea rtrm6a);
- \<And>rtrm61 rtrm61a rtrm62 rtrm62a.
-    \<lbrakk>rtrm61 = rtrm61a; P rtrm61 rtrm61a;
-     \<exists>pi. fv_trm6 rtrm62 - bv6 rtrm61 = fv_trm6 rtrm62a - bv6 rtrm61a \<and>
-          (fv_trm6 rtrm62 - bv6 rtrm61) \<sharp>* pi \<and> pi \<bullet> rtrm62 = rtrm62a \<and> P (pi \<bullet> rtrm62) rtrm62a\<rbrakk>
-    \<Longrightarrow> P (Lt6 rtrm61 rtrm62) (Lt6 rtrm61a rtrm62a)\<rbrakk>
-\<Longrightarrow> P x1 x2"
-apply (lifting alpha_rtrm6.induct[unfolded alpha_gen])
-apply injection
-(* notice unsolvable goals: (alpha_rtrm6 ===> op =) rbv6 rbv6 *)
-oops
-
-lemma lifted_inject_a3:
-"(Lt6 rtrm61 rtrm62 = Lt6 rtrm61a rtrm62a) =
-(rtrm61 = rtrm61a \<and>
- (\<exists>pi. fv_trm6 rtrm62 - bv6 rtrm61 = fv_trm6 rtrm62a - bv6 rtrm61a \<and>
-       (fv_trm6 rtrm62 - bv6 rtrm61) \<sharp>* pi \<and> pi \<bullet> rtrm62 = rtrm62a))"
-apply(lifting alpha6_inj(3)[unfolded alpha_gen])
-apply injection
-(* notice unsolvable goals: (alpha_rtrm6 ===> op =) rbv6 rbv6 *)
-oops
-
-
-
-
-(* example with a respectful bn function defined over the type itself *)
-
-datatype rtrm7 =
-  rVr7 "name"
-| rLm7 "name" "rtrm7" --"bind left in right"
-| rLt7 "rtrm7" "rtrm7" --"bind (bv7 left) in (right)"
-
-primrec
-  rbv7
-where
-  "rbv7 (rVr7 n) = {atom n}"
-| "rbv7 (rLm7 n t) = rbv7 t - {atom n}"
-| "rbv7 (rLt7 l r) = rbv7 l \<union> rbv7 r"
-
-setup {* snd o define_raw_perms ["rtrm7"] ["Terms.rtrm7"] *}
-thm permute_rtrm7.simps
-
-local_setup {* snd o define_fv_alpha "Terms.rtrm7" [
-  [[[]], [[(NONE, 0)], [(NONE, 0)]], [[], [(SOME @{term rbv7}, 0)]]]] *}
-print_theorems
-notation
-  alpha_rtrm7 ("_ \<approx>7a _" [100, 100] 100)
-(* HERE THE RULES DIFFER *)
-thm alpha_rtrm7.intros
-thm fv_rtrm7.simps
-inductive
-  alpha7 :: "rtrm7 \<Rightarrow> rtrm7 \<Rightarrow> bool" ("_ \<approx>7 _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (rVr7 a) \<approx>7 (rVr7 b)"
-| a2: "(\<exists>pi. (({atom a}, t) \<approx>gen alpha7 fv_rtrm7 pi ({atom b}, s))) \<Longrightarrow> rLm7 a t \<approx>7 rLm7 b s"
-| a3: "(\<exists>pi. (((rbv7 t1), s1) \<approx>gen alpha7 fv_rtrm7 pi ((rbv7 t2), s2))) \<Longrightarrow> rLt7 t1 s1 \<approx>7 rLt7 t2 s2"
-
-lemma "(x::name) \<noteq> y \<Longrightarrow> \<not> (alpha7 ===> op =) rbv7 rbv7"
-  apply simp
-  apply (rule_tac x="rLt7 (rVr7 x) (rVr7 x)" in exI)
-  apply (rule_tac x="rLt7 (rVr7 y) (rVr7 y)" in exI)
-  apply simp
-  apply (rule a3)
-  apply (rule_tac x="(x \<leftrightarrow> y)" in exI)
-  apply (simp_all add: alpha_gen fresh_star_def)
-  apply (rule a1)
-  apply (rule refl)
-done
-
-
-
-
-
-datatype rfoo8 =
-  Foo0 "name"
-| Foo1 "rbar8" "rfoo8" --"bind bv(bar) in foo"
-and rbar8 =
-  Bar0 "name"
-| Bar1 "name" "name" "rbar8" --"bind second name in b"
-
-primrec
-  rbv8
-where
-  "rbv8 (Bar0 x) = {}"
-| "rbv8 (Bar1 v x b) = {atom v}"
-
-setup {* snd o define_raw_perms ["rfoo8", "rbar8"] ["Terms.rfoo8", "Terms.rbar8"] *}
-print_theorems
-
-local_setup {* snd o define_fv_alpha "Terms.rfoo8" [
-  [[[]], [[], [(SOME @{term rbv8}, 0)]]], [[[]], [[], [(NONE, 1)], [(NONE, 1)]]]] *}
-notation
-  alpha_rfoo8 ("_ \<approx>f' _" [100, 100] 100) and
-  alpha_rbar8 ("_ \<approx>b' _" [100, 100] 100)
-(* HERE THE RULE DIFFERS *)
-thm alpha_rfoo8_alpha_rbar8.intros
-
-
-inductive
-  alpha8f :: "rfoo8 \<Rightarrow> rfoo8 \<Rightarrow> bool" ("_ \<approx>f _" [100, 100] 100)
-and
-  alpha8b :: "rbar8 \<Rightarrow> rbar8 \<Rightarrow> bool" ("_ \<approx>b _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (Foo0 a) \<approx>f (Foo0 b)"
-| a2: "a = b \<Longrightarrow> (Bar0 a) \<approx>b (Bar0 b)"
-| a3: "b1 \<approx>b b2 \<Longrightarrow> (\<exists>pi. (((rbv8 b1), t1) \<approx>gen alpha8f fv_rfoo8 pi ((rbv8 b2), t2))) \<Longrightarrow> Foo1 b1 t1 \<approx>f Foo1 b2 t2"
-| a4: "v1 = v2 \<Longrightarrow> (\<exists>pi. (({atom x1}, t1) \<approx>gen alpha8b fv_rbar8 pi ({atom x2}, t2))) \<Longrightarrow> Bar1 v1 x1 t1 \<approx>b Bar1 v2 x2 t2"
-
-lemma "(alpha8b ===> op =) rbv8 rbv8"
-  apply simp apply clarify
-  apply (erule alpha8f_alpha8b.inducts(2))
-  apply (simp_all)
-done
-
-lemma fv_rbar8_rsp_hlp: "x \<approx>b y \<Longrightarrow> fv_rbar8 x = fv_rbar8 y"
-  apply (erule alpha8f_alpha8b.inducts(2))
-  apply (simp_all add: alpha_gen)
-done
-lemma "(alpha8b ===> op =) fv_rbar8 fv_rbar8"
-  apply simp apply clarify apply (simp add: fv_rbar8_rsp_hlp)
-done
-
-lemma "(alpha8f ===> op =) fv_rfoo8 fv_rfoo8"
-  apply simp apply clarify
-  apply (erule alpha8f_alpha8b.inducts(1))
-  apply (simp_all add: alpha_gen fv_rbar8_rsp_hlp)
-done
-
-
-
-
-
-
-datatype rlam9 =
-  Var9 "name"
-| Lam9 "name" "rlam9" --"bind name in rlam"
-and rbla9 =
-  Bla9 "rlam9" "rlam9" --"bind bv(first) in second"
-
-primrec
-  rbv9
-where
-  "rbv9 (Var9 x) = {}"
-| "rbv9 (Lam9 x b) = {atom x}"
-
-setup {* snd o define_raw_perms ["rlam9", "rbla9"] ["Terms.rlam9", "Terms.rbla9"] *}
-print_theorems
-
-local_setup {* snd o define_fv_alpha "Terms.rlam9" [
-  [[[]], [[(NONE, 0)], [(NONE, 0)]]], [[[], [(SOME @{term rbv9}, 0)]]]] *}
-notation
-  alpha_rlam9 ("_ \<approx>9l' _" [100, 100] 100) and
-  alpha_rbla9 ("_ \<approx>9b' _" [100, 100] 100)
-(* HERE THE RULES DIFFER *)
-thm alpha_rlam9_alpha_rbla9.intros
-
-
-inductive
-  alpha9l :: "rlam9 \<Rightarrow> rlam9 \<Rightarrow> bool" ("_ \<approx>9l _" [100, 100] 100)
-and
-  alpha9b :: "rbla9 \<Rightarrow> rbla9 \<Rightarrow> bool" ("_ \<approx>9b _" [100, 100] 100)
-where
-  a1: "a = b \<Longrightarrow> (Var9 a) \<approx>9l (Var9 b)"
-| a4: "(\<exists>pi. (({atom x1}, t1) \<approx>gen alpha9l fv_rlam9 pi ({atom x2}, t2))) \<Longrightarrow> Lam9 x1 t1 \<approx>9l Lam9 x2 t2"
-| a3: "b1 \<approx>9l b2 \<Longrightarrow> (\<exists>pi. (((rbv9 b1), t1) \<approx>gen alpha9l fv_rlam9 pi ((rbv9 b2), t2))) \<Longrightarrow> Bla9 b1 t1 \<approx>9b Bla9 b2 t2"
-
-quotient_type
-  lam9 = rlam9 / alpha9l and bla9 = rbla9 / alpha9b
-sorry
-
-local_setup {*
-(fn ctxt => ctxt
- |> snd o (Quotient_Def.quotient_lift_const ("qVar9", @{term Var9}))
- |> snd o (Quotient_Def.quotient_lift_const ("qLam9", @{term Lam9}))
- |> snd o (Quotient_Def.quotient_lift_const ("qBla9", @{term Bla9}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_lam9", @{term fv_rlam9}))
- |> snd o (Quotient_Def.quotient_lift_const ("fv_bla9", @{term fv_rbla9}))
- |> snd o (Quotient_Def.quotient_lift_const ("bv9", @{term rbv9})))
-*}
-print_theorems
-
-instantiation lam9 and bla9 :: pt
-begin
-
-quotient_definition
-  "permute_lam9 :: perm \<Rightarrow> lam9 \<Rightarrow> lam9"
-is
-  "permute :: perm \<Rightarrow> rlam9 \<Rightarrow> rlam9"
-
-quotient_definition
-  "permute_bla9 :: perm \<Rightarrow> bla9 \<Rightarrow> bla9"
-is
-  "permute :: perm \<Rightarrow> rbla9 \<Rightarrow> rbla9"
-
-instance
-sorry
-
-end
-
-lemma "\<lbrakk>b1 = b2; \<exists>pi. fv_lam9 t1 - bv9 b1 = fv_lam9 t2 - bv9 b2 \<and> (fv_lam9 t1 - bv9 b1) \<sharp>* pi \<and> pi \<bullet> t1 = t2\<rbrakk>
- \<Longrightarrow> qBla9 b1 t1 = qBla9 b2 t2"
-apply (lifting a3[unfolded alpha_gen])
-apply injection
-sorry
-
-
-
-
-
-
-
-
-text {* type schemes *} 
-datatype ty = 
-  Var "name" 
-| Fun "ty" "ty"
-
-setup {* snd o define_raw_perms ["ty"] ["Terms.ty"] *}
-print_theorems
-
-datatype tyS = 
-  All "name set" "ty" 
-
-setup {* snd o define_raw_perms ["tyS"] ["Terms.tyS"] *}
-print_theorems
-
-local_setup {* snd o define_fv_alpha "Terms.ty" [[[[]], [[], []]]] *}
-print_theorems 
-
-(*
-Doesnot work yet since we do not refer to fv_ty
-local_setup {* define_raw_fv "Terms.tyS" [[[[], []]]] *}
-print_theorems
-*)
-
-primrec
-  fv_tyS
-where 
-  "fv_tyS (All xs T) = (fv_ty T - atom ` xs)"
-
-inductive
-  alpha_tyS :: "tyS \<Rightarrow> tyS \<Rightarrow> bool" ("_ \<approx>tyS _" [100, 100] 100)
-where
-  a1: "\<exists>pi. ((atom ` xs1, T1) \<approx>gen (op =) fv_ty pi (atom ` xs2, T2)) 
-        \<Longrightarrow> All xs1 T1 \<approx>tyS All xs2 T2"
-
-lemma
-  shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {b, a} (Fun (Var a) (Var b))"
-  apply(rule a1)
-  apply(simp add: alpha_gen)
-  apply(rule_tac x="0::perm" in exI)
-  apply(simp add: fresh_star_def)
-  done
-
-lemma
-  shows "All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var b) (Var a))"
-  apply(rule a1)
-  apply(simp add: alpha_gen)
-  apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI)
-  apply(simp add: fresh_star_def)
-  done
-
-lemma
-  shows "All {a, b, c} (Fun (Var a) (Var b)) \<approx>tyS All {a, b} (Fun (Var a) (Var b))"
-  apply(rule a1)
-  apply(simp add: alpha_gen)
-  apply(rule_tac x="0::perm" in exI)
-  apply(simp add: fresh_star_def)
-  done
-
-lemma
-  assumes a: "a \<noteq> b"
-  shows "\<not>(All {a, b} (Fun (Var a) (Var b)) \<approx>tyS All {c} (Fun (Var c) (Var c)))"
-  using a
-  apply(clarify)
-  apply(erule alpha_tyS.cases)
-  apply(simp add: alpha_gen)
-  apply(erule conjE)+
-  apply(erule exE)
-  apply(erule conjE)+
-  apply(clarify)
-  apply(simp)
-  apply(simp add: fresh_star_def)
-  apply(auto)
-  done
-
-
-end
--- a/Quot/Nominal/Test.thy	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,444 +0,0 @@
-theory Test
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp"
-begin
-
-atom_decl name
-
-
-(* tests *)
-ML {*
-Datatype.datatype_cmd;
-Datatype.add_datatype Datatype.default_config;
-
-Primrec.add_primrec_cmd;
-Primrec.add_primrec;
-Primrec.add_primrec_simple;
-*}
-
-section {* test for setting up a primrec on the ML-level *}
-
-section{* Interface for nominal_datatype *}
-
-text {*
-
-Nominal-Datatype-part:
-
-1st Arg: string list  
-         ^^^^^^^^^^^
-         strings of the types to be defined
-
-2nd Arg: (string list * binding * mixfix * (binding * typ list * mixfix) list) list
-         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^   ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^
-               type(s) to be defined             constructors list
-               (ty args, name, syn)              (name, typs, syn)
-
-Binder-Function-part:
-
-3rd Arg: (binding * typ option * mixfix) list 
-         ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^    
-            binding function(s)           
-              to be defined               
-            (name, type, syn)             
-
-4th Arg:  term list 
-          ^^^^^^^^^
-          the equations of the binding functions
-          (Trueprop equations)
-*}
-
-text {*****************************************************}
-ML {* 
-(* nominal datatype parser *)
-local
-  structure P = OuterParse
-in
-
-val _ = OuterKeyword.keyword "bind"
-val anno_typ = Scan.option (P.name --| P.$$$ "::") -- P.typ
-
-(* binding specification *)
-(* should use and_list *)
-val bind_parser = 
-  P.enum "," ((P.$$$ "bind" |-- P.term) -- (P.$$$ "in" |-- P.name))
-
-val constr_parser =
-  P.binding -- Scan.repeat anno_typ
-
-(* datatype parser *)
-val dt_parser =
-  ((P.type_args -- P.binding -- P.opt_infix) >> P.triple1) -- 
-    (P.$$$ "=" |-- P.enum1 "|" ((constr_parser -- bind_parser -- P.opt_mixfix) >> P.triple_swap))
-
-(* function equation parser *)
-val fun_parser = 
-  Scan.optional (P.$$$ "binder" |-- P.fixes -- SpecParse.where_alt_specs) ([],[])
-
-(* main parser *)
-val main_parser =
-  (P.and_list1 dt_parser) -- fun_parser
-
-end
-*}
-
-(* adds "_raw" to the end of constants and types *)
-ML {*
-fun add_raw s = s ^ "_raw"
-fun add_raws ss = map add_raw ss
-fun raw_bind bn = Binding.suffix_name "_raw" bn
-
-fun replace_str ss s = 
-  case (AList.lookup (op=) ss s) of 
-     SOME s' => s'
-   | NONE => s
-
-fun replace_typ ty_ss (Type (a, Ts)) = Type (replace_str ty_ss a, map (replace_typ ty_ss) Ts)
-  | replace_typ ty_ss T = T  
-
-fun raw_dts ty_ss dts =
-let
-  val ty_ss' = ty_ss ~~ (add_raws ty_ss)
-
-  fun raw_dts_aux1 (bind, tys, mx) =
-    (raw_bind bind, map (replace_typ ty_ss') tys, mx)
-
-  fun raw_dts_aux2 (ty_args, bind, mx, constrs) =
-    (ty_args, raw_bind bind, mx, map raw_dts_aux1 constrs)
-in
-  map raw_dts_aux2 dts
-end
-
-fun replace_aterm trm_ss (Const (a, T)) = Const (replace_str trm_ss a, T)
-  | replace_aterm trm_ss (Free (a, T)) = Free (replace_str trm_ss a, T)
-  | replace_aterm trm_ss trm = trm
-
-fun replace_term trm_ss ty_ss trm =
-  trm |> Term.map_aterms (replace_aterm trm_ss) |> map_types (replace_typ ty_ss) 
-*}
-
-ML {*
-fun get_constrs dts =
-  flat (map (fn (_, _, _, constrs) => constrs) dts)
-
-fun get_typed_constrs dts =
-  flat (map (fn (_, bn, _, constrs) => 
-   (map (fn (bn', _, _) => (Binding.name_of bn, Binding.name_of bn')) constrs)) dts)
-
-fun get_constr_strs dts =
-  map (fn (bn, _, _) => Binding.name_of bn) (get_constrs dts)
-
-fun get_bn_fun_strs bn_funs =
-  map (fn (bn_fun, _, _) => Binding.name_of bn_fun) bn_funs
-*}
-
-ML {*
-fun raw_dts_decl dt_names dts lthy =
-let
-  val thy = ProofContext.theory_of lthy
-  val conf = Datatype.default_config
-
-  val dt_names' = add_raws dt_names
-  val dt_full_names = map (Sign.full_bname thy) dt_names 
-  val dts' = raw_dts dt_full_names dts
-in
-  lthy
-  |> Local_Theory.theory_result (Datatype.add_datatype conf dt_names' dts')
-end 
-*}
-
-ML {*
-fun raw_bn_fun_decl dt_names dts bn_funs bn_eqs lthy =
-let
-  val thy = ProofContext.theory_of lthy
-
-  val dt_names' = add_raws dt_names
-  val dt_full_names = map (Sign.full_bname thy) dt_names 
-  val dt_full_names' = map (Sign.full_bname thy) dt_names' 
-  
-  val ctrs_names = map (Sign.full_bname thy) (get_constr_strs dts)
-  val ctrs_names' = map (fn (x, y) => (Sign.full_bname_path thy (add_raw x) (add_raw y))) 
-    (get_typed_constrs dts)
-
-  val bn_fun_strs = get_bn_fun_strs bn_funs
-  val bn_fun_strs' = add_raws bn_fun_strs
-  
-  val bn_funs' = map (fn (bn, opt_ty, mx) => 
-    (raw_bind bn, Option.map (replace_typ (dt_full_names ~~ dt_full_names')) opt_ty, mx)) bn_funs
-  
-  val bn_eqs' = map (fn trm => 
-    (Attrib.empty_binding, 
-      (replace_term ((ctrs_names ~~ ctrs_names') @ (bn_fun_strs ~~ bn_fun_strs')) (dt_full_names ~~ dt_full_names') trm))) bn_eqs
-in
-  if null bn_eqs 
-  then (([], []), lthy)
-  else Primrec.add_primrec bn_funs' bn_eqs' lthy 
-end 
-*}
-
-ML {* 
-fun nominal_datatype2 dts bn_funs bn_eqs lthy =
-let
-  val dts_names = map (fn (_, s, _, _) => Binding.name_of s) dts
-in
-  lthy
-  |> raw_dts_decl dts_names dts
-  ||>> raw_bn_fun_decl dts_names dts bn_funs bn_eqs
-end
-*}
-
-ML {*
-(* makes a full named type out of a binding with tvars applied to it *)
-fun mk_type thy bind tvrs =
-  Type (Sign.full_name thy bind, map (fn s => TVar ((s, 0), [])) tvrs)
-
-fun get_constrs2 thy dts =
-let
-  val dts' = map (fn (tvrs, tname, _, constrs) => (mk_type thy tname tvrs, constrs)) dts
-in
-  flat (map (fn (ty, constrs) => map (fn (bn, tys, mx) =>  (bn, tys ---> ty, mx)) constrs) dts')
-end
-*}
-
-ML {*
-fun nominal_datatype2_cmd (dt_strs, (bn_fun_strs, bn_eq_strs)) lthy =
-let
-  val thy = ProofContext.theory_of lthy
-
-  fun prep_typ ((tvs, tname, mx), _) = (tname, length tvs, mx);
-
-  (* adding the types for parsing datatypes *)
-  val lthy_tmp = lthy
-    |> Local_Theory.theory (Sign.add_types (map prep_typ dt_strs))
-
-  fun prep_cnstr lthy (((cname, atys), mx), binders) =
-    (cname, map (Syntax.read_typ lthy o snd) atys, mx)
-  
-  fun prep_dt lthy ((tvs, tname, mx), cnstrs) = 
-    (tvs, tname, mx, map (prep_cnstr lthy) cnstrs)
-
-  (* parsing the datatypes *)
-  val dts_prep = map (prep_dt lthy_tmp) dt_strs
-
-  (* adding constructors for parsing functions *)
-  val lthy_tmp2 = lthy_tmp
-    |> Local_Theory.theory (Sign.add_consts_i (get_constrs2 thy dts_prep))
-
-  val (bn_fun_aux, bn_eq_aux) = fst (Specification.read_spec bn_fun_strs bn_eq_strs lthy_tmp2)
-
-  fun prep_bn_fun ((bn, T), mx) = (bn, SOME T, mx) 
-
-  fun prep_bn_eq (attr, t) = t
-
-  val bn_fun_prep = map prep_bn_fun bn_fun_aux
-  val bn_eq_prep = map prep_bn_eq bn_eq_aux 
-
-in
-  nominal_datatype2 dts_prep bn_fun_prep bn_eq_prep lthy |> snd
-end
-*}
-
-(* Command Keyword *)
-ML {*
-let
-   val kind = OuterKeyword.thy_decl
-in
-   OuterSyntax.local_theory "nominal_datatype" "test" kind 
-     (main_parser >> nominal_datatype2_cmd)
-end
-*}
-
-text {* example 1 *}
-
-nominal_datatype lam =
-  VAR "name"
-| APP "lam" "lam"
-| LET bp::"bp" t::"lam"   bind "bi bp" in t ("Let _ in _" [100,100] 100)
-and bp = 
-  BP "name" "lam"  ("_ ::= _" [100,100] 100)
-binder
-  bi::"bp \<Rightarrow> name set"
-where
-  "bi (BP x t) = {x}"
-
-typ lam_raw
-term VAR_raw
-term Test.BP_raw
-thm bi_raw.simps
-
-print_theorems
-
-text {* examples 2 *}
-nominal_datatype trm =
-  Var "name"
-| App "trm" "trm"
-| Lam x::"name" t::"trm"        bind x in t 
-| Let p::"pat" "trm" t::"trm"   bind "f p" in t
-and pat =
-  PN
-| PS "name"
-| PD "name" "name"
-binder
-  f::"pat \<Rightarrow> name set"
-where 
-  "f PN = {}"
-| "f (PS x) = {x}"
-| "f (PD x y) = {x,y}"
-
-thm f_raw.simps
-
-nominal_datatype trm0 =
-  Var0 "name"
-| App0 "trm0" "trm0"
-| Lam0 x::"name" t::"trm0"          bind x in t 
-| Let0 p::"pat0" "trm0" t::"trm0"   bind "f0 p" in t
-and pat0 =
-  PN0
-| PS0 "name"
-| PD0 "pat0" "pat0"
-binder
-  f0::"pat0 \<Rightarrow> name set"
-where 
-  "f0 PN0 = {}"
-| "f0 (PS0 x) = {x}"
-| "f0 (PD0 p1 p2) = (f0 p1) \<union> (f0 p2)"
-
-thm f0_raw.simps
-
-text {* example type schemes *}
-datatype ty = 
-  Var "name" 
-| Fun "ty" "ty"
-
-nominal_datatype tyS = 
-  All xs::"name list" ty::"ty" bind xs in ty
-
-
-
-(* example 1 from Terms.thy *)
-
-nominal_datatype trm1 =
-  Vr1 "name"
-| Ap1 "trm1" "trm1"
-| Lm1 x::"name" t::"trm1"      bind x in t 
-| Lt1 p::"bp1" "trm1" t::"trm1" bind "bv1 p" in t 
-and bp1 =
-  BUnit1
-| BV1 "name"
-| BP1 "bp1" "bp1"
-binder
-  bv1
-where
-  "bv1 (BUnit1) = {}"
-| "bv1 (BV1 x) = {atom x}"
-| "bv1 (BP1 bp1 bp2) = (bv1 bp1) \<union> (bv1 bp2)"
-
-thm bv1_raw.simps
-
-(* example 2 from Terms.thy *)
-
-nominal_datatype trm2 =
-  Vr2 "name"
-| Ap2 "trm2" "trm2"
-| Lm2 x::"name" t::"trm2"       bind x in t
-| Lt2 r::"rassign" t::"trm2"    bind "bv2 r" in t
-and rassign = 
-  As "name" "trm2"
-binder
-  bv2
-where
-  "bv2 (As x t) = {atom x}"
-
-(* example 3 from Terms.thy *)
-
-nominal_datatype trm3 =
-  Vr3 "name"
-| Ap3 "trm3" "trm3"
-| Lm3 x::"name" t::"trm3"        bind x in t
-| Lt3 r::"rassigns3" t::"trm3"   bind "bv3 r" in t
-and rassigns3 =
-  ANil
-| ACons "name" "trm3" "rassigns3"
-binder
-  bv3
-where
-  "bv3 ANil = {}"
-| "bv3 (ACons x t as) = {atom x} \<union> (bv3 as)"
-
-(* example 4 from Terms.thy *)
-
-nominal_datatype trm4 =
-  Vr4 "name"
-| Ap4 "trm4" "trm4 list"
-| Lm4 x::"name" t::"trm4"  bind x in t
-
-(* example 5 from Terms.thy *)
-
-nominal_datatype trm5 =
-  Vr5 "name"
-| Ap5 "trm5" "trm5"
-| Lt5 l::"lts" t::"trm5"  bind "bv5 l" in t
-and lts =
-  Lnil
-| Lcons "name" "trm5" "lts"
-binder
-  bv5
-where
-  "bv5 Lnil = {}"
-| "bv5 (Lcons n t ltl) = {atom n} \<union> (bv5 ltl)"
-
-(* example 6 from Terms.thy *)
-
-nominal_datatype trm6 =
-  Vr6 "name"
-| Lm6 x::"name" t::"trm6"         bind x in t
-| Lt6 left::"trm6" right::"trm6"  bind "bv6 left" in right
-binder
-  bv6
-where
-  "bv6 (Vr6 n) = {}"
-| "bv6 (Lm6 n t) = {atom n} \<union> bv6 t"
-| "bv6 (Lt6 l r) = bv6 l \<union> bv6 r"
-
-(* example 7 from Terms.thy *)
-
-nominal_datatype trm7 =
-  Vr7 "name"
-| Lm7 l::"name" r::"trm7"   bind l in r
-| Lt7 l::"trm7" r::"trm7"   bind "bv7 l" in r
-binder 
-  bv7 
-where
-  "bv7 (Vr7 n) = {atom n}"
-| "bv7 (Lm7 n t) = bv7 t - {atom n}"
-| "bv7 (Lt7 l r) = bv7 l \<union> bv7 r"
-
-(* example 8 from Terms.thy *)
-
-nominal_datatype foo8 =
-  Foo0 "name"
-| Foo1 b::"bar8" f::"foo8" bind "bv8 b" in foo
-and bar8 =
-  Bar0 "name"
-| Bar1 "name" s::"name" b::"bar8" bind s in b
-binder 
-  bv8
-where
-  "bv8 (Bar0 x) = {}"
-| "bv8 (Bar1 v x b) = {atom v}"
-
-(* example 9 from Terms.thy *)
-
-nominal_datatype lam9 =
-  Var9 "name"
-| Lam9 n::"name" l::"lam9" bind n in l
-and bla9 =
-  Bla9 f::"lam9" s::"lam9" bind "bv9 f" in s
-binder
-  bv9
-where
-  "bv9 (Var9 x) = {}"
-| "bv9 (Lam9 x b) = {atom x}"
-
-end
-
-
-
--- a/Quot/Nominal/nominal_atoms.ML	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,94 +0,0 @@
-(*  Title:      nominal_atoms/ML
-    Authors:    Brian Huffman, Christian Urban
-
-    Command for defining concrete atom types. 
-    
-    At the moment, only single-sorted atom types
-    are supported. 
-*)
-
-signature ATOM_DECL =
-sig
-  val add_atom_decl: (binding * (binding option)) -> theory -> theory
-end;
-
-structure Atom_Decl :> ATOM_DECL =
-struct
-
-val atomT = @{typ atom};
-val permT = @{typ perm};
-
-val sort_of_const = @{term sort_of};
-fun atom_const T = Const (@{const_name atom}, T --> atomT);
-fun permute_const T = Const (@{const_name permute}, permT --> T --> T);
-
-fun mk_sort_of t = sort_of_const $ t;
-fun mk_atom t = atom_const (fastype_of t) $ t;
-fun mk_permute (p, t) = permute_const (fastype_of t) $ p $ t;
-
-fun atom_decl_set (str : string) : term =
-  let
-    val a = Free ("a", atomT);
-    val s = Const (@{const_name "Sort"}, @{typ "string => atom_sort list => atom_sort"})
-              $ HOLogic.mk_string str $ HOLogic.nil_const @{typ "atom_sort"};
-  in
-    HOLogic.mk_Collect ("a", atomT, HOLogic.mk_eq (mk_sort_of a, s))
-  end
-
-fun add_atom_decl (name : binding, arg : binding option) (thy : theory) =
-  let
-    val _ = Theory.requires thy "Nominal2_Atoms" "nominal logic";
-    val str = Sign.full_name thy name;
-
-    (* typedef *)
-    val set = atom_decl_set str;
-    val tac = rtac @{thm exists_eq_simple_sort} 1;
-    val ((full_tname, info as {type_definition, Rep_name, Abs_name, ...}), thy) =
-      Typedef.add_typedef false NONE (name, [], NoSyn) set NONE tac thy;
-
-    (* definition of atom and permute *)
-    val newT = #abs_type info;
-    val RepC = Const (Rep_name, newT --> atomT);
-    val AbsC = Const (Abs_name, atomT --> newT);
-    val a = Free ("a", newT);
-    val p = Free ("p", permT);
-    val atom_eqn =
-      HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_atom a, RepC $ a));
-    val permute_eqn =
-      HOLogic.mk_Trueprop (HOLogic.mk_eq
-        (mk_permute (p, a), AbsC $ (mk_permute (p, RepC $ a))));
-    val atom_def_name =
-      Binding.prefix_name "atom_" (Binding.suffix_name "_def" name);
-    val permute_def_name =
-      Binding.prefix_name "permute_" (Binding.suffix_name "_def" name);
-
-    (* at class instance *)
-    val lthy =
-      Theory_Target.instantiation ([full_tname], [], @{sort at}) thy;
-    val ((_, (_, permute_ldef)), lthy) =
-      Specification.definition (NONE, ((permute_def_name, []), permute_eqn)) lthy;
-    val ((_, (_, atom_ldef)), lthy) =
-      Specification.definition (NONE, ((atom_def_name, []), atom_eqn)) lthy;
-    val ctxt_thy = ProofContext.init (ProofContext.theory_of lthy);
-    val permute_def = singleton (ProofContext.export lthy ctxt_thy) permute_ldef;
-    val atom_def = singleton (ProofContext.export lthy ctxt_thy) atom_ldef;
-    val class_thm = @{thm at_class} OF [type_definition, atom_def, permute_def];
-    val thy = lthy
-      |> Class.prove_instantiation_instance (K (Tactic.rtac class_thm 1))
-      |> Local_Theory.exit_global;
-  in
-    thy
-  end;
-
-(** outer syntax **)
-
-local structure P = OuterParse and K = OuterKeyword in
-
-val _ =
-  OuterSyntax.command "atom_decl" "declaration of a concrete atom type" K.thy_decl
-    ((P.binding -- Scan.option (Args.parens (P.binding))) >>
-      (Toplevel.print oo (Toplevel.theory o add_atom_decl)));
-
-end;
-
-end;
--- a/Quot/Nominal/nominal_permeq.ML	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,71 +0,0 @@
-(*  Title:      nominal_thmdecls.ML
-    Author:     Brian Huffman, Christian Urban
-*)
-
-signature NOMINAL_PERMEQ =
-sig
-  val eqvt_tac: Proof.context -> int -> tactic 
-
-end;
-
-(* TODO:
-
- - provide a method interface with the usual add and del options
-
- - print a warning if for a constant no eqvt lemma is stored
-
- - there seems to be too much simplified in case of multiple 
-   permutations, like
-
-      p o q o r o x 
-
-   we usually only want the outermost permutation to "float" in
-*)
-
-
-structure Nominal_Permeq: NOMINAL_PERMEQ =
-struct
-
-local
-
-fun eqvt_apply_conv ctxt ct =
-  case (term_of ct) of
-    (Const (@{const_name "permute"}, _) $ _ $ (_ $ _)) =>
-      let
-        val (perm, t) = Thm.dest_comb ct
-        val (_, p) = Thm.dest_comb perm
-        val (f, x) = Thm.dest_comb t
-        val a = ctyp_of_term x;
-        val b = ctyp_of_term t;
-        val ty_insts = map SOME [b, a]
-        val term_insts = map SOME [p, f, x]
-      in
-        Drule.instantiate' ty_insts term_insts @{thm eqvt_apply}
-      end
-  | _ => Conv.no_conv ct
-
-fun eqvt_lambda_conv ctxt ct =
-  case (term_of ct) of
-    (Const (@{const_name "permute"}, _) $ _ $ Abs _) =>
-      Conv.rewr_conv @{thm eqvt_lambda} ct
-  | _ => Conv.no_conv ct
-
-in
-
-fun eqvt_conv ctxt ct =
-  Conv.first_conv
-    [ Conv.rewr_conv @{thm eqvt_bound},
-      eqvt_apply_conv ctxt
-        then_conv Conv.comb_conv (eqvt_conv ctxt),
-      eqvt_lambda_conv ctxt
-        then_conv Conv.abs_conv (fn (v, ctxt) => eqvt_conv ctxt) ctxt,
-      More_Conv.rewrs_conv (Nominal_ThmDecls.get_eqvts_raw_thms ctxt),
-      Conv.all_conv
-    ] ct
-
-fun eqvt_tac ctxt = 
-  CONVERSION (More_Conv.bottom_conv (fn ctxt => eqvt_conv ctxt) ctxt)
-
-end
-
-end; (* structure *)
\ No newline at end of file
--- a/Quot/Nominal/nominal_thmdecls.ML	Wed Feb 24 17:32:43 2010 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,134 +0,0 @@
-(*  Title:      nominal_thmdecls.ML
-    Author:     Christian Urban
-
-  Infrastructure for the lemma collection "eqvts".
-
-  Provides the attributes [eqvt] and [eqvt_raw], and the theorem
-  lists eqvts and eqvts_raw. The first attribute will store the 
-  theorem in the eqvts list and also in the eqvts_raw list. For 
-  the latter the theorem is expected to be of the form
-
-    p o (c x1 x2 ...) = c (p o x1) (p o x2) ...
-
-  and it is stored in the form
-
-    p o c == c
-
-  The [eqvt_raw] attribute just adds the theorem to eqvts_raw.
-
-  TODO:
-
-   - deal with eqvt-lemmas of the form 
-
-       c x1 x2 ... ==> c (p o x1) (p o x2) ..
-*)
-
-signature NOMINAL_THMDECLS =
-sig
-  val eqvt_add: attribute
-  val eqvt_del: attribute
-  val eqvt_raw_add: attribute
-  val eqvt_raw_del: attribute
-  val setup: theory -> theory
-  val get_eqvts_thms: Proof.context -> thm list
-  val get_eqvts_raw_thms: Proof.context -> thm list
-
-end;
-
-structure Nominal_ThmDecls: NOMINAL_THMDECLS =
-struct
-
-
-structure EqvtData = Generic_Data
-( type T = thm Item_Net.T;
-  val empty = Thm.full_rules;
-  val extend = I;
-  val merge = Item_Net.merge );
-
-structure EqvtRawData = Generic_Data
-( type T = thm Item_Net.T;
-  val empty = Thm.full_rules;
-  val extend = I;
-  val merge = Item_Net.merge );
-
-val eqvts = Item_Net.content o EqvtData.get;
-val eqvts_raw = Item_Net.content o EqvtRawData.get;
-
-val get_eqvts_thms = eqvts o Context.Proof; 
-val get_eqvts_raw_thms = eqvts_raw o Context.Proof; 
-
-val add_thm = EqvtData.map o Item_Net.update;
-val del_thm = EqvtData.map o Item_Net.remove;
-
-val add_raw_thm = EqvtRawData.map o Item_Net.update;
-val del_raw_thm = EqvtRawData.map o Item_Net.remove;
-
-fun dest_perm (Const (@{const_name "permute"}, _) $ p $ t) = (p, t)
-  | dest_perm t = raise TERM("dest_perm", [t])
-
-fun mk_perm p trm =
-let
-  val ty = fastype_of trm
-in
-  Const (@{const_name "permute"}, @{typ "perm"} --> ty --> ty) $ p $ trm
-end
-
-fun eqvt_transform_tac thm = REPEAT o FIRST' 
-  [CHANGED o simp_tac (HOL_basic_ss addsimps @{thms permute_minus_cancel}),
-   rtac (thm RS @{thm trans}),
-   rtac @{thm trans[OF permute_fun_def]} THEN' rtac @{thm ext}]
-
-(* transform equations into the required form *)
-fun transform_eq ctxt thm lhs rhs = 
-let
-  val (p, t) = dest_perm lhs
-  val (c, args) = strip_comb t
-  val (c', args') = strip_comb rhs 
-  val eargs = map Envir.eta_contract args 
-  val eargs' = map Envir.eta_contract args'
-  val p_str = fst (fst (dest_Var p))
-  val goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_perm p c, c))
-in
-  if c <> c' 
-    then error "eqvt lemma is not of the right form (constants do not agree)"
-  else if eargs' <> map (mk_perm p) eargs 
-    then error "eqvt lemma is not of the right form (arguments do not agree)"
-  else if args = [] 
-    then thm
-  else Goal.prove ctxt [p_str] [] goal
-    (fn _ => eqvt_transform_tac thm 1)
-end
-
-fun transform addel_fun thm context = 
-let
-  val ctxt = Context.proof_of context
-  val trm = HOLogic.dest_Trueprop (prop_of thm)
-in
-  case trm of
-    Const (@{const_name "op ="}, _) $ lhs $ rhs =>
-      let 
-        val thm' = transform_eq ctxt thm lhs rhs RS @{thm eq_reflection}
-      in
-        addel_fun thm' context
-      end
-  | _ => raise (error "only (op=) case implemented yet")
-end 
-
-val eqvt_add = Thm.declaration_attribute (fn thm => (add_thm thm) o (transform add_raw_thm thm));
-val eqvt_del = Thm.declaration_attribute (fn thm => (del_thm thm) o (transform del_raw_thm thm));
-
-val eqvt_raw_add = Thm.declaration_attribute add_raw_thm;
-val eqvt_raw_del = Thm.declaration_attribute del_raw_thm;
-
-val setup =
-  Attrib.setup @{binding "eqvt"} (Attrib.add_del eqvt_add eqvt_del) 
-    (cat_lines ["declaration of equivariance lemmas - they will automtically be",  
-       "brought into the form p o c = c"]) #>
-  Attrib.setup @{binding "eqvt_raw"} (Attrib.add_del eqvt_raw_add eqvt_raw_del) 
-    (cat_lines ["declaration of equivariance lemmas - no",
-       "transformation is performed"]) #>
-  PureThy.add_thms_dynamic (@{binding "eqvts"}, eqvts) #>
-  PureThy.add_thms_dynamic (@{binding "eqvts_raw"}, eqvts_raw);
-
-
-end;