merged
authorChristian Urban <urbanc@in.tum.de>
Mon, 09 May 2011 04:49:58 +0100
changeset 2780 2c6851248b3f
parent 2779 3c769bf10e63 (current diff)
parent 2776 8e0f0b2b51dd (diff)
child 2781 542ff50555f5
merged
Nominal/Nominal2_Base.thy
--- a/Nominal/Nominal2_Base.thy	Tue May 03 15:39:30 2011 +0100
+++ b/Nominal/Nominal2_Base.thy	Mon May 09 04:49:58 2011 +0100
@@ -761,7 +761,7 @@
   unfolding permute_perm_def by simp
 
 (* the normal version of this lemma would cause loops *)
-lemma permute_eqvt_raw[eqvt_raw]:
+lemma permute_eqvt_raw [eqvt_raw]:
   shows "p \<bullet> permute \<equiv> permute"
 apply(simp add: fun_eq_iff permute_fun_def)
 apply(subst permute_eqvt)
@@ -826,26 +826,6 @@
   unfolding Ex1_def
   by (perm_simp) (rule refl)
 
-lemma mem_eqvt [eqvt]:
-  shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
-  unfolding mem_def
-  by (simp add: permute_fun_app_eq)
-
-lemma Collect_eqvt [eqvt]:
-  shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
-  unfolding Collect_def permute_fun_def ..
-
-lemma inter_eqvt [eqvt]:
-  shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)"
-  unfolding Int_def
-  by (perm_simp) (rule refl)
-
-lemma image_eqvt [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 if_eqvt [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)
@@ -896,7 +876,21 @@
   apply(rule theI'[OF unique])
   done
 
-subsubsection {* Equivariance set operations *}
+subsubsection {* Equivariance of Set operators *}
+
+lemma mem_eqvt [eqvt]:
+  shows "p \<bullet> (x \<in> A) \<longleftrightarrow> (p \<bullet> x) \<in> (p \<bullet> A)"
+  unfolding mem_def
+  by (rule permute_fun_app_eq)
+
+lemma Collect_eqvt [eqvt]:
+  shows "p \<bullet> {x. P x} = {x. (p \<bullet> P) x}"
+  unfolding Collect_def permute_fun_def ..
+
+lemma inter_eqvt [eqvt]:
+  shows "p \<bullet> (A \<inter> B) = (p \<bullet> A) \<inter> (p \<bullet> B)"
+  unfolding Int_def
+  by (perm_simp) (rule refl)
 
 lemma Bex_eqvt [eqvt]:
   shows "p \<bullet> (\<exists>x \<in> S. P x) = (\<exists>x \<in> (p \<bullet> S). (p \<bullet> P) x)"
@@ -908,6 +902,11 @@
   unfolding Ball_def
   by (perm_simp) (rule refl)
 
+lemma image_eqvt [eqvt]:
+  shows "p \<bullet> (f ` A) = (p \<bullet> f) ` (p \<bullet> A)"
+  unfolding image_def
+  by (perm_simp) (rule refl)
+
 lemma UNIV_eqvt [eqvt]:
   shows "p \<bullet> UNIV = UNIV"
   unfolding UNIV_def
@@ -1039,12 +1038,12 @@
 subsubsection {* Equivariance for product operations *}
 
 lemma fst_eqvt [eqvt]:
-  "p \<bullet> (fst x) = fst (p \<bullet> x)"
- by (cases x) simp
+  shows "p \<bullet> (fst x) = fst (p \<bullet> x)"
+  by (cases x) simp
 
 lemma snd_eqvt [eqvt]:
-  "p \<bullet> (snd x) = snd (p \<bullet> x)"
- by (cases x) simp
+  shows "p \<bullet> (snd x) = snd (p \<bullet> x)"
+  by (cases x) simp
 
 lemma split_eqvt [eqvt]: 
   shows "p \<bullet> (split P x) = split (p \<bullet> P) (p \<bullet> x)"
@@ -1099,7 +1098,7 @@
 
 lemma union_fset_eqvt [eqvt]:
   shows "(p \<bullet> (S |\<union>| T)) = ((p \<bullet> S) |\<union>| (p \<bullet> T))"
-by (induct S) (simp_all)
+  by (induct S) (simp_all)
 
 lemma map_fset_eqvt [eqvt]: 
   shows "p \<bullet> (map_fset f S) = map_fset (p \<bullet> f) (p \<bullet> S)"
@@ -1410,7 +1409,7 @@
 lemma plus_perm_eq:
   fixes p q::"perm"
   assumes asm: "supp p \<inter> supp q = {}"
-  shows "p + q  = q + p"
+  shows "p + q = q + p"
 unfolding perm_eq_iff
 proof
   fix a::"atom"
@@ -1493,7 +1492,7 @@
 
 instance prod :: (fs, fs) fs
 apply default
-apply (induct_tac x)
+apply (case_tac x)
 apply (simp add: supp_Pair finite_supp)
 done
 
@@ -1518,7 +1517,7 @@
 
 instance sum :: (fs, fs) fs
 apply default
-apply (induct_tac x)
+apply (case_tac x)
 apply (simp_all add: supp_Inl supp_Inr finite_supp)
 done
 
@@ -1554,22 +1553,22 @@
   shows "supp [] = {}"
   by (simp add: supp_def)
 
+lemma fresh_Nil: 
+  shows "a \<sharp> []"
+  by (simp add: fresh_def supp_Nil)
+
 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_Cons: 
+  shows "a \<sharp> (x # xs) \<longleftrightarrow> a \<sharp> x \<and> a \<sharp> xs"
+  by (simp add: fresh_def supp_Cons)
+
 lemma supp_append:
   shows "supp (xs @ ys) = supp xs \<union> supp ys"
   by (induct xs) (auto simp add: supp_Nil supp_Cons)
 
-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)
-
 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)
@@ -1614,7 +1613,7 @@
   using assms
   unfolding fresh_conv_MOST
   unfolding permute_fun_app_eq
-  by (elim MOST_rev_mp, simp)
+  by (elim MOST_rev_mp) (simp)
 
 lemma supp_fun_app:
   shows "supp (f x) \<subseteq> (supp f) \<union> (supp x)"
@@ -1628,13 +1627,11 @@
 definition
   "eqvt f \<equiv> \<forall>p. p \<bullet> f = f"
 
-
 text {* equivariance of a function at a given argument *}
 
 definition
  "eqvt_at f x \<equiv> \<forall>p. p \<bullet> (f x) = f (p \<bullet> x)"
 
-
 lemma eqvtI:
   shows "(\<And>p. p \<bullet> f \<equiv> f) \<Longrightarrow> eqvt f"
 unfolding eqvt_def
@@ -2105,6 +2102,19 @@
 
 section {* Induction principle for permutations *}
 
+lemma smaller_supp:
+  assumes a: "a \<in> supp p"
+  shows "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subset> supp p"
+proof -
+  have "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subseteq> supp p"
+    unfolding supp_perm by (auto simp add: swap_atom)
+  moreover
+  have "a \<notin> supp ((p \<bullet> a \<rightleftharpoons> a) + p)" by (simp add: supp_perm)
+  then have "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<noteq> supp p" using a by auto
+  ultimately 
+  show "supp ((p \<bullet> a \<rightleftharpoons> a) + p) \<subset> supp p" by auto
+qed
+  
 
 lemma perm_struct_induct[consumes 1, case_names zero swap]:
   assumes S: "supp p \<subseteq> S"
@@ -2128,11 +2138,7 @@
       then have a1: "p \<bullet> a \<in> S" "a \<in> S" "sort_of (p \<bullet> a) = sort_of a" "p \<bullet> a \<noteq> a"
         using as by (auto simp add: supp_atom supp_perm swap_atom)
       let ?q = "(p \<bullet> a \<rightleftharpoons> a) + p"
-      have a2: "supp ?q \<subseteq> supp p" unfolding supp_perm by (auto simp add: swap_atom)
-      moreover
-      have "a \<notin> supp ?q" by (simp add: supp_perm)
-      then have "supp ?q \<noteq> supp p" using a0 by auto
-      ultimately have "supp ?q \<subset> supp p" using a2 by auto
+      have a2: "supp ?q \<subset> supp p" using a0 smaller_supp by simp
       then have "P ?q" using ih by simp
       moreover
       have "supp ?q \<subseteq> S" using as a2 by simp
@@ -2213,6 +2219,7 @@
   qed
 qed
 
+text {* same lemma as above, but proved with a different induction principle *}
 lemma supp_perm_eq_test:
   assumes "(supp x) \<sharp>* p"
   shows "p \<bullet> x = x"
@@ -2237,8 +2244,24 @@
 lemma perm_supp_eq:
   assumes a: "(supp p) \<sharp>* x"
   shows "p \<bullet> x = x"
-by (rule supp_perm_eq)
-   (simp add: fresh_star_supp_conv a)
+proof -
+  from assms have "supp p \<subseteq> {a. a \<sharp> x}"
+    unfolding supp_perm fresh_star_def fresh_def by auto
+  then show "p \<bullet> x = x"
+  proof (induct p rule: perm_struct_induct2)
+    case zero
+    show "0 \<bullet> x = x" by simp
+  next
+    case (swap a b)
+    then have "a \<sharp> x" "b \<sharp> x" by simp_all
+    then show "(a \<rightleftharpoons> b) \<bullet> x = x" by (simp add: swap_fresh_fresh)
+  next
+    case (plus p1 p2)
+    have "p1 \<bullet> x = x" "p2 \<bullet> x = x" by fact+
+    then show "(p1 + p2) \<bullet> x = x" by simp
+  qed
+qed
+
 
 lemma supp_perm_perm_eq:
   assumes a: "\<forall>a \<in> supp x. p \<bullet> a = q \<bullet> a"
@@ -2636,7 +2659,7 @@
   apply(simp)
   done
 
-lemma fresh_at_base_permute_iff[simp]:
+lemma fresh_at_base_permute_iff [simp]:
   fixes a::"'a::at_base"
   shows "atom (p \<bullet> a) \<sharp> p \<bullet> x \<longleftrightarrow> atom a \<sharp> x"
   unfolding atom_eqvt[symmetric]
--- a/Pearl-jv/Paper.thy	Tue May 03 15:39:30 2011 +0100
+++ b/Pearl-jv/Paper.thy	Mon May 09 04:49:58 2011 +0100
@@ -156,7 +156,8 @@
   $\new$
 
 
-
+  Obstacles for Coq; no type-classes, difficulties with quotient types,
+  need for classical reasoning
 
   Two binders
 
@@ -188,7 +189,7 @@
   related work).  Therefore we use here a single unified atom type to
   represent atoms of different sorts. A basic requirement is that
   there must be a countably infinite number of atoms of each sort.
-  This can be implemented as the datatype
+  This can be implemented in Isabelle/HOL as the datatype
 
 *}
 
@@ -221,6 +222,9 @@
   sort a = s and a \<notin> S"}.
   \end{proposition}
 
+  \noindent
+  This property will be used later on whenever we have to chose a `fresh' atom.
+
   For implementing sort-respecting permutations, we use functions of type @{typ
   "atom => atom"} that are bijective; are the
   identity on all atoms, except a finite number of them; and map
@@ -290,7 +294,7 @@
   a function in @{typ perm}. 
 
   One advantage of using functions as a representation for
-  permutations is that it is unique. For example the swappings
+  permutations is that it is a unique representation. For example the swappings
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
@@ -352,10 +356,14 @@
   \end{isabelle} 
 
   \noindent
-  whereby @{text "\<beta>"} is a generic type for @{text x}. The definition of this operation will be 
-  given by in terms of `induction' over this generic type. The type-class mechanism
-  of Isabelle/HOL \cite{Wenzel04} allows us to give a definition for
-  `base' types, such as atoms, permutations, booleans and natural numbers:
+  whereby @{text "\<beta>"} is a generic type for the object @{text
+  x}.\footnote{We will use the standard notation @{text "((op \<bullet>) \<pi>)
+  x"} for this operation in the few cases where we need to indicate
+  that it is a function applied with two arguments.}  The definition
+  of this operation will be given by in terms of `induction' over this
+  generic type. The type-class mechanism of Isabelle/HOL
+  \cite{Wenzel04} allows us to give a definition for `base' types,
+  such as atoms, permutations, booleans and natural numbers:
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}l@ {\hspace{4mm}}l@ {}}
@@ -435,7 +443,7 @@
 
   \noindent
   Note that the permutation operation for functions is defined so that
-  we have for applications the property
+  we have for applications the equation
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   @{text "\<pi> \<bullet> (f x) ="}
@@ -444,21 +452,21 @@
   \end{isabelle}
 
   \noindent
-  whenever the permutation properties hold for @{text x}. This property can
+  provided the permutation properties hold for @{text x}. This equation can
   be easily shown by unfolding the permutation operation for functions on
-  the right-hand side, simplifying the beta-redex and eliminating the permutations
-  in front of @{text x} using \eqref{cancel}.
+  the right-hand side of the equation, simplifying the resulting beta-redex 
+  and eliminating the permutations in front of @{text x} using \eqref{cancel}.
 
   The main benefit of the use of type classes is that it allows us to delegate 
   much of the routine resoning involved in determining whether the permutation properties
   are satisfied to Isabelle/HOL's type system: we only have to
   establish that base types satisfy them and that type-constructors
-  preserve them. Isabelle/HOL will use this information and determine
-  whether an object @{text x} with a compound type satisfies the
+  preserve them. Then Isabelle/HOL will use this information and determine
+  whether an object @{text x} with a compound type, like @{typ "atom \<Rightarrow> (atom set * nat)"}, satisfies the
   permutation properties.  For this we define the notion of a
   \emph{permutation type}:
 
-  \begin{definition}[Permutation type]
+  \begin{definition}[Permutation Type]
   A type @{text "\<beta>"} is a \emph{permutation type} if the permutation
   properties in \eqref{newpermprops} are satisfied for every @{text
   "x"} of type @{text "\<beta>"}.
@@ -494,47 +502,29 @@
 section {* Equivariance *}
 
 text {*
-  An important notion in the nominal logic work is
-  \emph{equivariance}.  This notion allows us to characterise how
-  permutations act upon compound statements in HOL by analysing how
-  these statements are constructed.  To do so, let us first define
-  \emph{HOL-terms}. They are given by the grammar
-
-  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
-  @{text "t ::= c | x | t\<^isub>1 t\<^isub>2 | \<lambda>x. t"}
-  \hfill\numbered{holterms}
-  \end{isabelle} 
-
-  \noindent
-  where @{text c} stands for constants and @{text x} for
-  variables. 
-  We assume HOL-terms are fully typed, but for the sake of
-  greater legibility we leave the typing information implicit.  We
-  also assume the usual notions for free and bound variables of a
-  HOL-term.  Furthermore, it is custom in HOL to regard terms as equal
-  modulo alpha-, beta- and eta-equivalence.
-
-  An \emph{equivariant} HOL-term is one that is invariant under the
-  permutation operation. This can be defined in Isabelle/HOL 
-  as follows:
+  Two important notions in the nominal logic work are what Pitts calls
+  \emph{equivariance} and the \emph{equivariance principle}.  These
+  notions allows us to characterise how permutations act upon compound
+  statements in HOL by analysing how these statements are constructed.
+  The notion of equivariance can defined as follows:
 
   \begin{definition}[Equivariance]\label{equivariance}
-  A  HOL-term @{text t} is \emph{equivariant} provided 
-  @{term "\<pi> \<bullet> t = t"} holds for all permutations @{text "\<pi>"}.
+  An object @{text "x"} of permutation type is \emph{equivariant} provided 
+  for all permutations @{text "\<pi>"}, \mbox{@{term "\<pi> \<bullet> x = x"}} holds.
   \end{definition}
 
   \noindent
-  In what follows we will primarily be interested in the cases where @{text t} 
-  is a constant, but of course there is no way in Isabelle/HOL to restrict 
-  this definition to just these cases.
+  In what follows we will primarily be interested in the cases where
+  @{text x} is a constant, but of course there is no way in
+  Isabelle/HOL to restrict this definition to just these cases.
 
   There are a number of equivalent formulations for the equivariance
-  property.  For example, assuming @{text t} is of permutation type @{text "\<alpha> \<Rightarrow>
-  \<beta>"}, then equivariance can also be stated as
+  property.  For example, assuming @{text f} is a function of permutation 
+  type @{text "\<alpha> \<Rightarrow> \<beta>"}, then equivariance can also be stated as
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}l}
-  @{text "\<forall>\<pi> x.  \<pi> \<bullet> (t x) = t (\<pi> \<bullet> x)"}
+  @{text "\<forall>\<pi> x.  \<pi> \<bullet> (f x) = f (\<pi> \<bullet> x)"}
   \end{tabular}\hfill\numbered{altequivariance}
   \end{isabelle} 
 
@@ -544,13 +534,13 @@
   the definition of the permutation operation for functions and
   simplify with the equation and the cancellation property shown in
   \eqref{cancel}. To see the other direction, we use
-  \eqref{permutefunapp}. Similarly for HOL-terms that take more than
+  \eqref{permutefunapp}. Similarly for functions that take more than
   one argument. The point to note is that equivariance and equivariance in fully
-  applied form are (for permutation types) always interderivable.
+  applied form are always interderivable (for permutation types).
 
   Both formulations of equivariance have their advantages and
   disadvantages: \eqref{altequivariance} is usually more convenient to
-  establish, since statements in Isabelle/HOL are commonly given in a
+  establish, since statements in HOL are commonly given in a
   form where functions are fully applied. For example we can easily
   show that equality is equivariant
 
@@ -579,61 +569,137 @@
   \noindent
   for all @{text a}, @{text b} and @{text \<pi>}. Also the booleans
   @{const True} and @{const False} are equivariant by the definition
-  of the permutation operation for booleans. It is easy to see
-  that the boolean operators, like @{text "\<and>"}, @{text "\<or>"}, @{text
-  "\<not>"} and @{text "\<longrightarrow>"}, are equivariant too; for example we have
+  of the permutation operation for booleans. Given this definition, it 
+  is also easy to see that the boolean operators, like @{text "\<and>"}, 
+  @{text "\<or>"}, @{text "\<longrightarrow>"} and  @{text "\<not>"} are equivariant:
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}lcl}
   @{text "\<pi> \<bullet> (A \<and> B) = (\<pi> \<bullet> A) \<and> (\<pi> \<bullet> B)"}\\
+  @{text "\<pi> \<bullet> (A \<or> B) = (\<pi> \<bullet> A) \<or> (\<pi> \<bullet> B)"}\\
   @{text "\<pi> \<bullet> (A \<longrightarrow> B) = (\<pi> \<bullet> A) \<longrightarrow> (\<pi> \<bullet> B)"}\\
+  @{text "\<pi> \<bullet> (\<not>A) = \<not>(\<pi> \<bullet> A)"}\\
   \end{tabular}
   \end{isabelle}
-
-  \noindent
-  by the definition of the permutation operation acting on booleans.
   
   In contrast, the advantage of Definition \ref{equivariance} is that
-  it leads to a relatively simple rewrite system that allows us to `push' a permutation
-  towards the leaves of a HOL-term (i.e.~constants and
-  variables).  Then the permutation disappears in cases where the
-  constants are equivariant. We have implemented this rewrite system
-  as a simplification tactic on the ML-level of Isabelle/HOL.  Having this tactic 
-  at our disposal, together with a collection of constants for which 
-  equivariance is already established, we can automatically establish 
-  equivariance of a constant for which equivariance is not yet known. For this we only have to 
-  make sure that the definiens of this constant 
-  is a HOL-term whose constants are all equivariant.  In what follows 
-  we shall specify this tactic and argue that it terminates and 
-  is correct (in the sense of pushing a 
-  permutation @{text "\<pi>"} inside a term and the only remaining 
-  instances of @{text "\<pi>"} are in front of the term's free variables). 
+  it allows us to state a general principle how permutations act on
+  statements in HOL.  For this we will define a rewrite system that
+  `pushes' a permutation towards the leaves of statements (i.e.~constants
+  and variables).  Then the permutations disappear in cases where the
+  constants are equivariant.  To do so, let us first define
+  \emph{HOL-terms}, which are the building blocks of statements in HOL.
+  They are given by the grammar
+
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  @{text "t ::= c | x | t\<^isub>1 t\<^isub>2 | \<lambda>x. t"}
+  \hfill\numbered{holterms}
+  \end{isabelle} 
+
+  \noindent
+  where @{text c} stands for constants and @{text x} for variables.
+  We assume HOL-terms are fully typed, but for the sake of better
+  legibility we leave the typing information implicit.  We also assume
+  the usual notions for free and bound variables of a HOL-term.
+  Furthermore, HOL-terms are regarded as equal modulo alpha-, beta-
+  and eta-equivalence. The equivariance principle can now be stated
+  formally as follows:
+
+  \begin{theorem}[Equivariance Principle]\label{eqvtprin}
+  Suppose a HOL-term @{text t} whose constants are all equivariant. For any 
+  permutation @{text \<pi>}, let @{text t'} be @{text t} except every 
+  free variable @{text x} in @{term t} is replaced by @{text "\<pi> \<bullet> x"}, then 
+  @{text "\<pi> \<bullet> t = t'"}.
+  \end{theorem}
+  
+  \noindent
+  The significance of this principle is that we can automatically establish
+  the equivariance of a constant for which equivariance is not yet
+  known. For this we only have to make sure that the definiens of this
+  constant is a HOL-term whose constants are all equivariant. For example
+  the universal quantifier @{text "\<forall>"} is definied in HOL as 
+
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  @{text "\<forall>x. P x \<equiv> "}~@{thm (rhs) All_def[no_vars]}
+  \end{isabelle} 
+
+  \noindent
+  The constants in the definiens @{thm (rhs) All_def[no_vars]}, namely @{text "="}
+  and @{text "True"}, are equivariant (we shown this above). Therefore
+  the equivariance principle gives us
 
-  The simplifiaction tactic is a rewrite systems consisting of four `oriented' 
-  equations. We will first give a naive version of this tactic, which however 
-  is in some cornercases incorrect and does not terminate, and then modify 
-  it in order to obtain the desired properties. A permutation @{text \<pi>} can 
-  be pushed into applications and abstractions as follows
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  @{text "\<pi> \<bullet> (\<forall>x. P x)  \<equiv>  \<pi> \<bullet> (P = (\<lambda>x. True))  =  ((\<pi> \<bullet> P) = (\<lambda>x. True))  \<equiv>  \<forall>x. (\<pi> \<bullet> P) x"}
+  \end{isabelle} 
+
+  \noindent
+  and consequently, the constant @{text "\<forall>"} must be equivariant. From this
+  we can deduce that the existential quantifier @{text "\<exists>"} is equivariant. 
+  Its definition in HOL is
+
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  @{text "\<exists>x. P x \<equiv> "}~@{thm (rhs) Ex_def[no_vars]}
+  \end{isabelle} 
+
+  \noindent
+  where again the HOL-term on the right-hand side only contains equivariant constants 
+  (namely @{text "\<forall>"} and @{text "\<longrightarrow>"}). Taking both facts together, we can deduce that
+  the unique existential quantifier @{text "\<exists>!"} is equivariant. Its definition 
+  is
+
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  @{text "\<exists>!x. P x \<equiv> "}~@{thm (rhs) Ex1_def[no_vars]}
+  \end{isabelle} 
 
+  \noindent
+  with all constants on the right-hand side being equivariant. With this kind
+  of reasoning we can build up a database of equivariant constants.
+
+  Before we proceed, let us give a justification for the equivariance principle. 
+  For this we will use a rewrite system consisting of a series of equalities 
+  
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  @{text "\<pi> \<cdot> t  =  ...  =  t'"}
+  \end{isabelle}
+  
+  \noindent 
+  that establish the equality between @{term "\<pi> \<bullet> t"} and
+  @{text "t'"}. The idea of the rewrite system is to push the
+  permutation inside the term @{text t}. We have implemented this as a
+  conversion tactic on the ML-level of Isabelle/HOL.  In what follows,
+  we will show that this tactic produces only finitely many equations
+  and also show that is correct (in the sense of pushing a permutation
+  @{text "\<pi>"} inside a term and the only remaining instances of @{text
+  "\<pi>"} are in front of the term's free variables).  The tactic applies
+  four `oriented' equations. We will first give a naive version of
+  this tactic, which however in some cornercases produces incorrect
+  results or does not terminate.  We then give a modification in order
+  to obtain the desired properties.
+
+  Consider the following oriented equations
+  
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}lr@ {\hspace{3mm}}c@ {\hspace{3mm}}l}
   i) & @{text "\<pi> \<bullet> (t\<^isub>1 t\<^isub>2)"} & \rrh & @{term "(\<pi> \<bullet> t\<^isub>1) (\<pi> \<bullet> t\<^isub>2)"}\\
   ii) & @{text "\<pi> \<bullet> (\<lambda>x. t)"} & \rrh & @{text "\<lambda>x. \<pi> \<bullet> (t[x :=  (-\<pi>) \<bullet> x])"}\\
+  iii) & @{term "\<pi> \<bullet> (- \<pi>) \<bullet> x"} & \rrh & @{term "x"}\\
+  iv) &  @{term "\<pi> \<bullet> c"} & \rrh & 
+            {\rm @{term "c"}\hspace{6mm}provided @{text c} is equivariant}\\
   \end{tabular}\hfill\numbered{rewriteapplam}
   \end{isabelle}
 
   \noindent
+  A permutation @{text \<pi>} can be pushed into applications and abstractions as follows
+
   The first equation we established in \eqref{permutefunapp};
   the second follows from the definition of permutations acting on functions
   and the fact that HOL-terms are equal modulo beta-equivalence.
-  Once the permutations are pushed towards the leaves we need the
+  Once the permutations are pushed towards the leaves, we need the
   following two equations
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}lr@ {\hspace{3mm}}c@ {\hspace{3mm}}l}
-  iii) & @{term "\<pi> \<bullet> (- \<pi>) \<bullet> x"} & \rrh & @{term "x"}\\
-  iv) &  @{term "\<pi> \<bullet> c"} & \rrh & 
-            {\rm @{term "c"}\hspace{6mm}provided @{text c} is equivariant}\\
+  
   \end{tabular}\hfill\numbered{rewriteother}
   \end{isabelle}
 
@@ -641,10 +707,10 @@
   in order to remove permuations in front of bound variables and
   equivariant constants.  Unfortunately, we have to be careful with
   the rules {\it i)} and {\it iv}): they can lead to a loop whenever
-  \mbox{@{text "t\<^isub>1 t\<^isub>2"}} is of the form @{text "((op \<bullet>) \<pi>') t"}. Note
-  that we usually write this application using infix notation as
-  @{text "\<pi> \<bullet> t"} and recall that by Lemma \ref{permutecompose} the
-  constant @{text "(op \<bullet>)"} is equivariant. Now consider the infinite
+  \mbox{@{text "t\<^isub>1 t\<^isub>2"}} is of the form @{text "((op \<bullet>) \<pi>') t"}, where
+  we do not write the permutation operation infix, as usual, but
+  as an application. Recall that we established in Lemma \ref{permutecompose} that the
+  constant @{text "(op \<bullet>)"} is equivariant and consider the infinite
   reduction sequence
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
@@ -659,37 +725,36 @@
   \end{isabelle}
 
   \noindent
-  where the last step is again an instance of the first term, but it
+  The last term is again an instance of rewrite rule {\it i}), but the term
   is bigger.  To avoid this loop we need to apply our rewrite rule
   using an `outside to inside' strategy.  This strategy is sufficient
   since we are only interested of rewriting terms of the form @{term
   "\<pi> \<bullet> t"}, where an outermost permutation needs to pushed inside a term.
 
-  Another problem we have to avoid is that the rules {\it i)} and
-  {\it iii)} can `overlap'. For this note that
-  the term @{term "\<pi> \<bullet>(\<lambda>x. x)"} reduces by {\it ii)} to 
-  @{term "\<lambda>x. \<pi> \<bullet> (- \<pi>) \<bullet> x"}, to which we can apply rule {\it iii)} 
-  in order to obtain @{term "\<lambda>x. x"}, as is desired---there is no 
-  free variable in the original term and so the permutation should completely
-  vanish. However, the subterm @{text
-  "(- \<pi>) \<bullet> x"} is also an application. Consequently, the term 
-  @{term "\<lambda>x. \<pi> \<bullet> (- \<pi>) \<bullet>x"} can reduce to @{text "\<lambda>x. (- (\<pi> \<bullet> \<pi>)) \<bullet> (\<pi> \<bullet> x)"} using
-  {\it i)}.  Given our strategy we cannot apply rule {\it iii)} anymore and 
-  even worse the measure we will introduce shortly increased. On the
-  other hand, if we had started with the term @{text "\<pi> \<bullet> ((- \<pi>) \<bullet> x)"}
-  where @{text \<pi>} and @{text x} are free variables, then we \emph{do}
-  want to apply rule  {\it i)} and not rule {\it iii)}. The latter 
-  would eliminate @{text \<pi>} completely. The problem is that rule {\it iii)}
-  should only apply to instances where the variable is to bound; for free variables 
-  we want to use {\it ii)}. 
+  Another problem we have to avoid is that the rules {\it i)} and {\it
+  iii)} can `overlap'. For this note that the term @{term "\<pi>
+  \<bullet>(\<lambda>x. x)"} reduces by {\it ii)} to @{term "\<lambda>x. \<pi> \<bullet> (- \<pi>) \<bullet> x"}, to
+  which we can apply rule {\it iii)} in order to obtain @{term
+  "\<lambda>x. x"}, as is desired---there is no free variable in the original
+  term and so the permutation should completely vanish. However, the
+  subterm @{text "(- \<pi>) \<bullet> x"} is also an application. Consequently,
+  the term @{term "\<lambda>x. \<pi> \<bullet> (- \<pi>) \<bullet>x"} can reduce to @{text "\<lambda>x. (- (\<pi>
+  \<bullet> \<pi>)) \<bullet> (\<pi> \<bullet> x)"} using {\it i)}.  Given our strategy we cannot
+  apply rule {\it iii)} anymore and even worse the measure we will
+  introduce shortly increased. On the other hand, if we had started
+  with the term @{text "\<pi> \<bullet> ((- \<pi>) \<bullet> x)"} where @{text \<pi>} and @{text
+  x} are free variables, then we \emph{do} want to apply rule {\it i)}
+  and not rule {\it iii)}. The latter would eliminate @{text \<pi>}
+  completely. The problem is that rule {\it iii)} should only apply to
+  instances where the variable is to bound; for free variables we want
+  to use {\it ii)}.  In order to distinguish these cases we have to
+  maintain the information which variable is bound when inductively
+  taking a term `apart'. This, unfortunately, does not mesh well with
+  the way how conversion tactics are implemented in Isabelle/HOL.
 
-  The problem is that in order to distinguish both cases when
-  inductively taking a term `apart', we have to maintain the
-  information which variable is bound. This, unfortunately, does not
-  mesh well with the way how simplification tactics are implemented in
-  Isabelle/HOL. Our remedy is to use a standard trick in HOL: we
-  introduce a separate definition for terms of the form @{text "(- \<pi>)
-  \<bullet> x"}, namely as
+  Our remedy is to use a standard trick in HOL: we introduce a
+  separate definition for terms of the form @{text "(- \<pi>) \<bullet> x"},
+  namely as
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   @{term "unpermute \<pi> x \<equiv> (- \<pi>) \<bullet> x"}
@@ -721,43 +786,45 @@
   leaves) and the second is the number of occurences of @{text
   "unpermute \<pi> x"} and @{text "\<pi> \<bullet> c"}.
 
-  With the definition of the simplification tactic in place, we can
-  establish its correctness. The property we are after is that for for
-  a HOL-term @{text t} whose constants are all equivariant, the
-  HOL-term @{text "\<pi> \<bullet> t"} is equal to @{text "t'"} with @{text "t'"}
+  With the rules of the conversions tactic in place, we can
+  establish its correctness. The property we are after is that 
+  for a HOL-term @{text t} whose constants are all equivariant the 
+  term \mbox{@{text "\<pi> \<bullet> t"}} is equal to @{text "t'"} with @{text "t'"}
   being equal to @{text t} except that every free variable @{text x}
-  in @{text t} is replaced by @{text "\<pi> \<bullet> x"}.  Pitts calls this
-  property \emph{equivariance principle} (book ref ???). In our
-  setting the precise statement of this property is a slightly more
-  involved because of the fact that @{text unpermutes} needs to be
-  treated specially.
-  
-  \begin{theorem}[Equivariance Principle]
-  Suppose a HOL-term @{text t} does not contain any @{text unpermutes} and all
-  its constants are equivariant. For any permutation @{text \<pi>}, let @{text t'} 
-  be the HOL-term @{text t} except every free variable @{text x} in @{term t} is 
-  replaced by @{text "\<pi> \<bullet> x"}, then @{text "\<pi> \<bullet> t = t'"}.
-  \end{theorem}
-  
+  in @{text t} is replaced by \mbox{@{text "\<pi> \<bullet> x"}}. Let us call
+  a variable @{text x} \emph{really free}, if it is free and not occuring
+  in an unpermute, such as @{text "unpermute _ x"} and @{text "unpermute x _"}.
+  We need the following technical notion characterising a \emph{@{text "\<pi>"}-proper} HOL-term
 
-
-  With these definitions in place we can define the notion of an \emph{equivariant}
-  HOL-term
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  \begin{tabular}{@ {}ll}
+  $\bullet$ & variables and constants are @{text \<pi>}-proper,\\
+  $\bullet$ & @{term "unpermute \<pi> x"} is @{text \<pi>}-proper,\\
+  $\bullet$ & @{term "\<lambda>x. t"} is @{text \<pi>}-proper, if @{text t} is @{text \<pi>}-proper and @{text x} is 
+  really free in @{text t}, and\\
+  $\bullet$ & @{term "t\<^isub>1 t\<^isub>2"} is @{text \<pi>}-proper, if @{text "t\<^isub>1"} and @{text "t\<^isub>2"} are 
+  @{text \<pi>}-proper.
+  \end{tabular}
+  \end{isabelle}
 
-  \begin{definition}[Equivariant HOL-term]
-  A HOL-term is \emph{equivariant}, provided it is closed and composed of applications, 
-  abstractions and equivariant constants only.
-  \end{definition}
+  \begin{proof}[Theorem~\ref{eqvtprin}] We establish the property if @{text t} 
+  is @{text \<pi>}-proper and only contains equivaraint constants, then 
+  @{text "\<pi> \<bullet> t = t'"} where @{text "t'"} is equal to @{text "t"} except all really
+  free variables @{text x} are replaced by @{text "\<pi> \<bullet> x"}, and all semi-free variables 
+  @{text "unpermute \<pi> x"} by @{text "x"}. We establish this property by induction
+  on the size of HOL-terms, counting terms like @{text "unpermuting \<pi> x"} as leafes,
+  like variables and constants. The cases for variables, constants and @{text unpermutes}
+  are routine. In the case of abstractions we have by induction hypothesis that 
+  @{text "\<pi> \<bullet> (t[x := unpermute \<pi> x]) = t'"} with @{text "t'"} satisfying our
+  correctness property. This implies that @{text "\<lambda>x. \<pi> \<bullet> (t[x := unpermute \<pi> x]) = \<lambda>x. t'"}
+  and hence @{text "\<pi> \<bullet> (\<lambda>x. t) = \<lambda>x. t'"} as needed.\hfill\qed  
+  \end{proof}
+
+  Pitts calls this property \emph{equivariance principle} (book ref ???). 
   
-  \noindent
-  For equivariant terms we have
-
-  \begin{lemma}
-  For an equivariant HOL-term @{text "t"},  @{term "\<pi> \<bullet> t = t"} for all permutations @{term "\<pi>"}.
-  \end{lemma}
-
-  Let us now see how to use the equivariance principle. We have 
-
+  Problems with @{text undefined}
+  
+  Lines of code
 *}
 
 
@@ -770,7 +837,8 @@
   but to any type for which a permutation operation is defined 
   (like lists, sets, functions and so on). 
 
-  \begin{definition}[Support] Given @{text x} is of permutation type, then 
+  \begin{definition}[Support] \label{support} 
+  Given @{text x} is of permutation type, then 
   
   @{thm [display,indent=10] supp_def[no_vars, THEN eq_reflection]}
   \end{definition}
@@ -789,13 +857,49 @@
   
   @{thm [display,indent=10] fresh_star_def[no_vars]}
 
+  \noindent
+  Using the equivariance principle, it can be easily checked that all three notions
+  are equivariant. A simple consequence of the definition of support and equivariance 
+  is that if @{text x} is equivariant then we have 
+
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  \begin{tabular}{@ {}l}
+  @{thm  (concl) supp_fun_eqvt[where f="x", no_vars]}
+  \end{tabular}\hfill\numbered{suppeqvtfun}
+  \end{isabelle} 
 
   \noindent
-  A striking consequence of these definitions is that we can prove
-  without knowing anything about the structure of @{term x} that
-  swapping two fresh atoms, say @{text a} and @{text b}, leave 
-  @{text x} unchanged. For the proof we use the following lemma 
-  about swappings applied to an @{text x}:
+  For function applications we can establish the following two properties.
+
+  \begin{lemma}\label{suppfunapp} Let @{text f} and @{text x} be of permutation type, then
+  \begin{isabelle}
+  \begin{tabular}{r@ {\hspace{4mm}}p{10cm}}
+  {\it i)} & @{thm[mode=IfThen] fresh_fun_app[no_vars]}\\
+  {\it ii)} & @{thm supp_fun_app[no_vars]}\\
+  \end{tabular}
+  \end{isabelle}
+  \end{lemma}
+
+  \begin{proof}
+  For the first property, we know from the assumption that @{term
+  "finite {b. (a \<rightleftharpoons> b) \<bullet> f \<noteq> f}"} and @{term "finite {b . (a \<rightleftharpoons> b) \<bullet> x \<noteq>
+  x}"} hold. That means for all, but finitely many @{text b} we have
+  @{term "(a \<rightleftharpoons> b) \<bullet> f = f"} and @{term "(a \<rightleftharpoons> b) \<bullet> x = x"}. Similarly,
+  we have to show that for but, but finitely @{text b} the equation
+  @{term "(a \<rightleftharpoons> b) \<bullet> f x = f x"} holds. The left-hand side of this
+  equation is equal to @{term "((a \<rightleftharpoons> b) \<bullet> f) ((a \<rightleftharpoons> b) \<bullet> x)"} by
+  \eqref{permutefunapp}, which we know by the previous two facts for
+  @{text f} and @{text x} is equal to the right-hand side for all,
+  but finitely many @{text b}. This establishes the first
+  property. The second is a simple corollary of {\it i)} by
+  unfolding the definition of freshness.\qed
+  \end{proof}
+
+  A striking consequence of the definitions for support and freshness
+  is that we can prove without knowing anything about the structure of
+  @{term x} that swapping two fresh atoms, say @{text a} and @{text
+  b}, leave @{text x} unchanged. For the proof we use the following
+  lemma about swappings applied to an @{text x}:
 
   \begin{lemma}\label{swaptriple}
   Assuming @{text x} is of permutation type, and @{text a}, @{text b} and @{text c} 
@@ -812,7 +916,7 @@
   
   \noindent
   are equal. The lemma is then by application of the second permutation 
-  property shown in~\eqref{newpermprops}.\hfill\qed
+  property shown in~\eqref{newpermprops}.\qed
   \end{proof}
 
   \begin{theorem}\label{swapfreshfresh}
@@ -830,81 +934,10 @@
   Now the theorem follows from Lemma~\ref{swaptriple}.\hfill\qed
   \end{proof}
   
-  \noindent
-  Two important properties that need to be established for later calculations is 
-  that @{text "supp"} and freshness are equivariant. For this we first show that:
-
-  \begin{lemma}\label{half}
-  If @{term x} is a permutation type, then @{thm (lhs) fresh_permute_iff[where p="\<pi>",no_vars]} 
-  if and only if @{thm (rhs) fresh_permute_iff[where p="\<pi>",no_vars]}.
-  \end{lemma}
-
-  \begin{proof}
-  \begin{isabelle}
-  \begin{tabular}[t]{c@ {\hspace{2mm}}l@ {\hspace{5mm}}l}
-  & @{thm (lhs) fresh_permute_iff[where p="\<pi>",no_vars]}\\ 
-  @{text "\<equiv>"} & 
-    @{term "finite {b. (\<pi> \<bullet> a \<rightleftharpoons> b) \<bullet> \<pi> \<bullet> x \<noteq> \<pi> \<bullet> x}"}\\
-  @{text "\<Leftrightarrow>"}
-  & @{term "finite {b. (\<pi> \<bullet> a \<rightleftharpoons> \<pi> \<bullet> b) \<bullet> \<pi> \<bullet> x \<noteq> \<pi> \<bullet> x}"} 
-  & since @{text "\<pi> \<bullet> _"} is bijective\\ 
-  @{text "\<Leftrightarrow>"}
-  & @{term "finite {b. \<pi> \<bullet> (a \<rightleftharpoons> b) \<bullet> x \<noteq> \<pi> \<bullet> x}"}
-  & by Lemma~\ref{permutecompose} and \eqref{swapeqvt}\\
-  @{text "\<Leftrightarrow>"}
-  & @{term "finite {b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"}
-  & by \eqref{permuteequ}\\
-   @{text "\<equiv>"}
-  & @{thm (rhs) fresh_permute_iff[where p="\<pi>",no_vars]}
-  \end{tabular}
-  \end{isabelle}\hfill\qed
-  \end{proof}
-
-  \noindent
-  Together with the definition of the permutation operation on booleans,
-  we can immediately infer equivariance of freshness: 
-
-  @{thm [display,indent=10] fresh_eqvt[where p="\<pi>",no_vars]}
-
-  \noindent
-  Now equivariance of @{text "supp"}, namely
-  
-  @{thm [display,indent=10] supp_eqvt[where p="\<pi>",no_vars]}
-  
-  \noindent
-  is by noting that @{thm supp_conv_fresh[no_vars]} and that freshness and 
-  the logical connectives are equivariant. ??? Equivariance
-
-  A simple consequence of the definition of support and equivariance is that 
-  if a function @{text f} is equivariant then we have 
-
-  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
-  \begin{tabular}{@ {}l}
-  @{thm  (concl) supp_fun_eqvt[no_vars]}
-  \end{tabular}\hfill\numbered{suppeqvtfun}
-  \end{isabelle} 
-
-  \noindent 
-  For function applications we can establish the two following properties.
-
-  \begin{lemma} Let @{text f} and @{text x} be of permutation type, then
-  \begin{isabelle}
-  \begin{tabular}{r@ {\hspace{4mm}}p{10cm}}
-  @{text "i)"} & @{thm[mode=IfThen] fresh_fun_app[no_vars]}\\
-  @{text "ii)"} & @{thm supp_fun_app[no_vars]}\\
-  \end{tabular}
-  \end{isabelle}
-  \end{lemma}
-
-  \begin{proof}
-  ???
-  \end{proof}
-
-
   While the abstract properties of support and freshness, particularly 
   Theorem~\ref{swapfreshfresh}, are useful for developing Nominal Isabelle, 
-  one often has to calculate the support of some concrete object. This is 
-  straightforward for example for booleans, nats, products and lists:
+  one often has to calculate the support of concrete objects.
+  For booleans, nats, products and lists it is easy to check that
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}{@ {}l@ {\hspace{4mm}}l@ {}}
@@ -916,8 +949,8 @@
   \end{tabular}
   \end{isabelle}
 
-  \noindent
-  But establishing the support of atoms and permutations is a bit 
+  \noindent 
+  hold.  Establishing the support of atoms and permutations is a bit 
   trickier. To do so we will use the following notion about a \emph{supporting set}.
   
   \begin{definition}[Supporting Set]
@@ -945,21 +978,21 @@
   \end{lemma}
 
   \begin{proof}
-  For @{text "i)"} we derive a contradiction by assuming there is an atom @{text a}
+  For {\it i)} we derive a contradiction by assuming there is an atom @{text a}
   with @{term "a \<in> supp x"} and @{text "a \<notin> S"}. Using the second fact, the 
   assumption that @{term "S supports x"} gives us that @{text S} is a superset of 
   @{term "{b. (a \<rightleftharpoons> b) \<bullet> x \<noteq> x}"}, which is finite by the assumption of @{text S}
   being finite. But this means @{term "a \<notin> supp x"}, contradicting our assumption.
-  Property @{text "ii)"} is by a direct application of 
-  Theorem~\ref{swapfreshfresh}. For the last property, part @{text "i)"} proves
+  Property {\it ii)} is by a direct application of 
+  Theorem~\ref{swapfreshfresh}. For the last property, part {\it i)} proves
   one ``half'' of the claimed equation. The other ``half'' is by property 
-  @{text "ii)"} and the fact that @{term "supp x"} is finite by @{text "i)"}.\hfill\qed  
+  {\it ii)} and the fact that @{term "supp x"} is finite by {\it i)}.\hfill\qed  
   \end{proof}
 
   \noindent
   These are all relatively straightforward proofs adapted from the existing 
   nominal logic work. However for establishing the support of atoms and 
-  permutations we found  the following `optimised' variant of @{text "iii)"} 
+  permutations we found  the following `optimised' variant of {\it iii)} 
   more useful:
 
   \begin{lemma}\label{optimised} Let @{text x} be of permutation type.
@@ -1035,27 +1068,56 @@
 
   The main point about support is that whenever an object @{text x} has finite
   support, then Proposition~\ref{choosefresh} allows us to choose for @{text x} a 
-  fresh atom with arbitrary sort. This is an important operation in Nominal
+  fresh atom with arbitrary sort. This is a crucial operation in Nominal
   Isabelle in situations where, for example, a bound variable needs to be
   renamed. To allow such a choice, we only have to assume that 
   @{text "finite (supp x)"} holds. For more convenience we
-  can define a type class for types where every element has finite support, and
-  prove that the types @{term "atom"}, @{term "perm"}, lists, products and
-  booleans are instances of this type class. 
+  can define a type class in Isabelle/HOL corresponding to the 
+  property:
+  
+  \begin{definition}[Finitely Supported Type]
+  A type @{text "\<beta>"} is \emph{finitely supported} if @{term "finite (supp x)"}
+  holds for all @{text x} of type @{text "\<beta>"}.
+  \end{definition}
+ 
+  \noindent
+  By the calculations above we can easily establish 
+
+  \begin{theorem}\label{finsuptype} 
+  The types @{type atom}, @{type perm}, @{type bool} and @{type nat}
+  are fintitely supported, and assuming @{text \<beta>}, @{text "\<beta>\<^isub>1"} and
+  @{text "\<beta>\<^isub>2"} are finitely supported types, then @{text "\<beta> list"} and
+  @{text "\<beta>\<^isub>1 \<times> \<beta>\<^isub>2"} are finitely supported.
+  \end{theorem}
 
-  Unfortunately, this does not work for sets or Isabelle/HOL's function
-  type. There are functions and sets definable in Isabelle/HOL for which the
-  finite support property does not hold.  A simple example of a function with
-  infinite support is @{const nat_of} shown in \eqref{sortnatof}.  This
-  function's support is the set of \emph{all} atoms @{term "UNIV::atom set"}. 
-  To establish this we show
-  @{term "\<not> a \<sharp> nat_of"}. This is equivalent to assuming the set @{term
-  "{b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite and deriving a
-  contradiction. From the assumption we also know that @{term "{a} \<union> {b. (a \<rightleftharpoons>
-  b) \<bullet> nat_of \<noteq> nat_of}"} is finite. Then we can use
-  Proposition~\ref{choosefresh} to choose an atom @{text c} such that @{term
-  "c \<noteq> a"}, @{term "sort_of c = sort_of a"} and @{term "(a \<rightleftharpoons> c) \<bullet> nat_of =
-  nat_of"}.  Now we can reason as follows:
+  \noindent
+  The main benefit of using the finite support property for choosing a
+  fresh atom is that the reasoning is `compositional'. To see this,
+  assume we have a list of atoms and a method of choosing a fresh atom
+  that is not a member in this list---for example the maximum plus
+  one.  Then if we enlarge this list \emph{after} the choice, then
+  obviously the fresh atom might not be fresh anymore. In contrast, by
+  the classical reasoning of Proposition~\ref{choosefresh} we know a
+  fresh atom exists for every list of atoms and no matter how we
+  extend this list of atoms, we always preserve the property of being
+  finitely supported. Consequently the existence of a fresh atom is
+  still guarantied by Proposition~\ref{choosefresh}.  Using the method
+  of `maximum plus one' we might have to adapt the choice of a fresh
+  atom.
+
+  Unfortunately, Theorem~\ref{finsuptype} does not work in general for the
+  types of sets and functions.  There are functions definable in HOL
+  for which the finite support property does not hold.  A simple
+  example of a function with infinite support is @{const nat_of} shown
+  in \eqref{sortnatof}.  This function's support is the set of
+  \emph{all} atoms @{term "UNIV::atom set"}.  To establish this we
+  show @{term "\<not> a \<sharp> nat_of"}. This is equivalent to assuming the set
+  @{term "{b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite and deriving a
+  contradiction. From the assumption we also know that @{term "{a} \<union>
+  {b. (a \<rightleftharpoons> b) \<bullet> nat_of \<noteq> nat_of}"} is finite. Then we can use
+  Proposition~\ref{choosefresh} to choose an atom @{text c} such that
+  @{term "c \<noteq> a"}, @{term "sort_of c = sort_of a"} and @{term "(a \<rightleftharpoons> c)
+  \<bullet> nat_of = nat_of"}.  Now we can reason as follows:
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   \begin{tabular}[b]{@ {}rcl@ {\hspace{5mm}}l}
@@ -1075,15 +1137,19 @@
 section {* Support of Finite Sets *}
 
 text {*
-  Also the set type is one instance whose elements are not generally finitely 
+  Also the set type is an instance whose elements are not generally finitely 
   supported (we will give an example in Section~\ref{concrete}). 
   However, we can easily show that finite sets and co-finite sets of atoms are finitely
-  supported, because their support can be characterised as:
+  supported. Their support can be characterised as:
 
-  \begin{lemma}\label{finatomsets}\mbox{}\\
-  @{text "i)"} If @{text S} is a finite set of atoms, then @{thm (concl) supp_finite_atom_set[no_vars]}.\\
-  @{text "ii)"} If @{term "UNIV - (S::atom set)"} is a finite set of atoms, then 
+  \begin{lemma}\label{finatomsets}
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  \begin{tabular}[b]{@ {}rl}
+  {\it i)} & If @{text S} is a finite set of atoms, then @{thm (concl) supp_finite_atom_set[no_vars]}.\\
+  {\it ii)} & If @{term "UNIV - (S::atom set)"} is a finite set of atoms, then 
   @{thm (concl) supp_cofinite_atom_set[no_vars]}.
+  \end{tabular}
+  \end{isabelle}
   \end{lemma}
 
   \begin{proof}
@@ -1100,23 +1166,32 @@
   @{term "supp (UNIV::atom set) = {}"}.
   More difficult, however, is it to establish that finite sets of finitely 
   supported objects are finitely supported. For this we first show that
-  the union of the suports of finitely many and finitely supported  objects 
+  the union of the supports of finitely many and finitely supported  objects 
   is finite, namely
 
   \begin{lemma}\label{unionsupp}
-  If @{text S} is a finite set whose elements are all finitely supported, then\\
-  @{text "i)"} @{thm (concl) Union_of_finite_supp_sets[no_vars]} and\\
-  @{text "ii)"} @{thm (concl) Union_included_in_supp[no_vars]}.
+  If @{text S} is a finite set whose elements are all finitely supported, then
+  %
+  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+  \begin{tabular}[b]{@ {}rl}
+  {\it i)} & @{thm (concl) Union_of_finite_supp_sets[no_vars]} and\\
+  {\it ii)} & @{thm (concl) Union_included_in_supp[no_vars]}.
+  \end{tabular}
+  \end{isabelle}
   \end{lemma}
 
   \begin{proof}
-  The first part is by a straightforward induction on the finiteness of @{text S}. 
-  For the second part, we know that @{term "\<Union>x\<in>S. supp x"} is a set of atoms, which
-  by the first part is finite. Therefore we know by Lemma~\ref{finatomsets}.@{text "i)"}
-  that @{term "(\<Union>x\<in>S. supp x) = supp (\<Union>x\<in>S. supp x)"}. Taking @{text "f"} to be
-  \mbox{@{text "\<lambda>S. \<Union> (supp ` S)"}}, we can write the right hand side as @{text "supp (f S)"}.
-  Since @{text "f"} is an equivariant function, we have that 
-  @{text "supp (f S) \<subseteq> supp S"} by ??? This completes the second part.\hfill\qed
+  The first part is by a straightforward induction on the finiteness
+  of @{text S}.  For the second part, we know that @{term "\<Union>x\<in>S. supp
+  x"} is a set of atoms, which by the first part is finite. Therefore
+  we know by Lemma~\ref{finatomsets}.{\it i)} that @{term "(\<Union>x\<in>S. supp
+  x) = supp (\<Union>x\<in>S. supp x)"}. Taking @{text "f"} to be the function
+  \mbox{@{text "\<lambda>S. \<Union> (supp ` S)"}}, we can write the right-hand side
+  as @{text "supp (f S)"}.  Since @{text "f"} is an equivariant
+  function (can be easily checked by the equivariance principle), we
+  have that @{text "supp (f S) \<subseteq> supp S"} by
+  Lemma~\ref{suppfunapp}.{\it ii)}. This completes the second
+  part.\hfill\qed
   \end{proof}
 
   \noindent
@@ -1127,26 +1202,50 @@
   \end{lemma}
 
   \begin{proof}
-  The right-to-left inclusion is proved in Lemma~\ref{unionsupp}.@{text "ii)"}. To show the inclusion 
-  in the other direction we have to show Lemma~\ref{supports}.@{text "i)"}
+  The right-to-left inclusion is proved in Lemma~\ref{unionsupp}.{\it ii)}. To show the inclusion 
+  in the other direction we can use Lemma~\ref{supports}.{\it i)}. This means
+  for all @{text a} and @{text b} that are not in @{text S} we have to show that
+  @{term "(a \<rightleftharpoons> b) \<bullet> (\<Union>x \<in> S. supp x) = (\<Union>x \<in> S. supp x)"} holds. By the equivariance
+  principle, the left-hand side is equal to @{term "\<Union>x \<in> ((a \<rightleftharpoons> b) \<bullet> S). supp x"}. Now
+  the swapping in front of @{text S} disappears, since @{term "a \<sharp> S"} and @{term "b \<sharp> S"}
+  whenever @{text "a, b \<notin> S"}. Thus we are done.\hfill\qed
   \end{proof}
+
+  \noindent
+  To sum up, every finite set of finitely supported elements has
+  finite support.  Unfortunately, we cannot use
+  Theorem~\ref{finsuptype} to let Isabelle/HOL find this out
+  automatically. This would require to introduce a separate type of
+  finite sets, which however is not so convenient to reason about as
+  Isabelle's standard set type.
 *}
 
 
-section {* Induction Principles *}
+section {* Induction Principles for Permutations *}
 
 text {*
   While the use of functions as permutation provides us with a unique
-  representation for permutations (for example @{term "(a \<rightleftharpoons> b)"} and 
-  @{term "(b \<rightleftharpoons> a)"} are equal permutations), this representation has 
-  one draw back: it does not come readily with an induction principle. 
-  Such an induction principle is handy for deriving properties like
+  representation for permutations (for example @{term "(a \<rightleftharpoons> b)"} and
+  @{term "(b \<rightleftharpoons> a)"} are equal permutations), this representation does
+  not come automatically with an induction principle.  Such an
+  induction principle is however handy for generalising
+  Lemma~\ref{swapfreshfresh} from swappings to permutations
   
-  @{thm [display, indent=10] supp_perm_eq[no_vars]}
+  \begin{lemma}
+  @{thm [mode=IfThen] perm_supp_eq[no_vars]}
+  \end{lemma}
 
   \noindent
-  However, it is not too difficult to derive an induction principle, 
-  given the fact that we allow only permutations with a finite domain. 
+  In this section we will establish an induction principle for permutations
+  with which this lemma can be easily proved. It is not too difficult to derive 
+  an induction principle for permutations, given the fact that we allow only 
+  permutations having a finite support. 
+
+  Using a the property from \cite{???}
+
+  \begin{lemma}\label{smallersupp}
+  @{thm [mode=IfThen] smaller_supp[no_vars]}
+  \end{lemma}
 *}
 
 
@@ -1370,14 +1469,13 @@
   type.  This design gives us the flexibility to define operations and prove
   theorems that are generic with respect to atom sorts.  For example, as
   illustrated above the @{term supp} function returns a set that includes the
-  free atoms of \emph{all} sorts together; the flexibility offered by the new
-  atom type makes this possible.  
+  free atoms of \emph{all} sorts together.
 
   However, the single multi-sorted atom type does not make an ideal interface
   for end-users of Nominal Isabelle.  If sorts are not distinguished by
   Isabelle's type system, users must reason about atom sorts manually.  That
-  means subgoals involving sorts must be discharged explicitly within proof
-  scripts, instead of being inferred by Isabelle/HOL's type checker.  In other
+  means for example that subgoals involving sorts must be discharged explicitly within proof
+  scripts, instead of being inferred automatically.  In other
   cases, lemmas might require additional side conditions about sorts to be true.
   For example, swapping @{text a} and @{text b} in the pair \mbox{@{term "(a,
   b)"}} will only produce the expected result if we state the lemma in
@@ -1433,10 +1531,10 @@
   must all have the same sort, that is
 
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
-  \begin{tabular}{r@ {\hspace{3mm}}l}
-  i)   if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\\
-  ii)  @{thm atom_eqvt[where p="\<pi>", no_vars]}\\
-  iii) @{thm sort_of_atom_eq [no_vars]}
+  \begin{tabular}{@ {}r@ {\hspace{3mm}}l}
+  i)   & if @{thm (lhs) atom_eq_iff [no_vars]} then @{thm (rhs) atom_eq_iff [no_vars]}\\
+  ii)  & @{thm atom_eqvt[where p="\<pi>", no_vars]}\\
+  iii) & @{thm sort_of_atom_eq [no_vars]}
   \end{tabular}\hfill\numbered{atomprops}
   \end{isabelle}
 
@@ -1445,7 +1543,7 @@
   show that @{typ name} satisfies all the above requirements of a concrete atom
   type.
 
-  The whole point of defining the concrete atom type class was to let users
+  The whole point of defining the concrete atom type class is to let users
   avoid explicit reasoning about sorts.  This benefit is realised by defining a
   special swapping operation of type @{text "\<alpha> \<Rightarrow> \<alpha>
   \<Rightarrow> perm"}, where @{text "\<alpha>"} is a concrete atom type:
@@ -1486,6 +1584,18 @@
 
   \noindent
   This command can be implemented using less than 100 lines of custom ML-code.
+  
+*}
+
+
+
+section {* Related Work\label{related} *}
+
+text {*
+  Coq-tries, but failed
+
+  Add here comparison with old work.
+
   In comparison, the old version of Nominal Isabelle included more than 1000
   lines of ML-code for creating concrete atom types, and for defining various
   type classes and instantiating generic lemmas for them.  In addition to
@@ -1494,14 +1604,6 @@
   theories that separately declare different atom types can be merged
   together---it is no longer required to collect all atom declarations in one
   place.
-*}
-
-
-
-section {* Related Work\label{related} *}
-
-text {*
-  Add here comparison with old work.
 
     Using a single atom type to represent atoms of different sorts and
   representing permutations as functions are not new ideas; see
--- a/Pearl-jv/document/root.tex	Tue May 03 15:39:30 2011 +0100
+++ b/Pearl-jv/document/root.tex	Mon May 09 04:49:58 2011 +0100
@@ -5,6 +5,7 @@
 \usepackage{amsmath}
 \usepackage{amssymb}
 \usepackage{mathabx}
+\usepackage{proof}
 \usepackage{longtable}
 \usepackage{graphics}
 \usepackage{pdfsetup}
@@ -47,7 +48,8 @@
 to have finite support. We present a formalisation that is based on a
 unified atom type and that represents permutations by bijective functions from
 atoms to atoms. Interestingly, we allow swappings, which are permutations
-build from two atoms, to be ill-sorted.  We also describe a formalisation of
+build from two atoms, to be ill-sorted.  We also describe a reasoning infrastructure
+that automates properties about equivariance, and present a formalisation of
 two abstraction operators that bind sets of names.
 \end{abstract}
 
--- a/Quotient-Paper/Paper.thy	Tue May 03 15:39:30 2011 +0100
+++ b/Quotient-Paper/Paper.thy	Mon May 09 04:49:58 2011 +0100
@@ -1,6 +1,7 @@
 (*<*)
 theory Paper
-imports "Quotient" "Quotient_Syntax"
+imports "Quotient" 
+        "~~/src/HOL/Library/Quotient_Syntax"
         "~~/src/HOL/Library/LaTeXsugar"
         "~~/src/HOL/Quotient_Examples/FSet"
 begin
@@ -25,7 +26,7 @@
   pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
   implies (infix "\<longrightarrow>" 100) and
   "==>" (infix "\<Longrightarrow>" 100) and
-  fun_map ("_ \<singlearr> _" 51) and
+  map_fun ("_ \<singlearr> _" 51) and
   fun_rel ("_ \<doublearr> _" 51) and
   list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
   empty_fset ("\<emptyset>") and
@@ -62,6 +63,15 @@
          (id ---> rep_fset ---> abs_fset) op #"
   by (simp add: fun_eq_iff abs_o_rep[OF Quotient_fset] map_id Quotient_abs_rep[OF Quotient_fset])
 
+lemma list_all2_symp:
+  assumes a: "equivp R"
+  and b: "list_all2 R xs ys"
+  shows "list_all2 R ys xs"
+using list_all2_lengthD[OF b] b
+apply(induct xs ys rule: list_induct2)
+apply(auto intro: equivp_symp[OF a])
+done
+
 lemma concat_rsp_unfolded:
   "\<lbrakk>list_all2 list_eq a ba; list_eq ba bb; list_all2 list_eq bb b\<rbrakk> \<Longrightarrow> list_eq (concat a) (concat b)"
 proof -
@@ -364,7 +374,7 @@
   also the following map for function types
 
   \begin{isabelle}\ \ \ \ \ %%%
-  @{thm fun_map_def[no_vars, THEN eq_reflection]}
+  @{thm map_fun_def[no_vars, THEN eq_reflection]}
   \end{isabelle}
 
   \noindent
--- a/Slides/Slides7.thy	Tue May 03 15:39:30 2011 +0100
+++ b/Slides/Slides7.thy	Mon May 09 04:49:58 2011 +0100
@@ -12,7 +12,7 @@
 (*>*)
 
 text_raw {*
-  \renewcommand{\slidecaption}{Hefei, 15.~April 2011}
+  \renewcommand{\slidecaption}{Beijing, 29.~April 2011}
 
   \newcommand{\abst}[2]{#1.#2}% atom-abstraction
   \newcommand{\pair}[2]{\langle #1,#2\rangle} % pairing
@@ -90,7 +90,7 @@
   \item Theorem provers can prevent mistakes, {\bf if} the problem
   is formulated so that it is suitable for theorem provers.\bigskip
   \item This re-formulation can be done, even in domains where
-  we do not expect it.
+  we least expect it.
   \end{itemize}
 
   \end{frame}}
@@ -237,7 +237,7 @@
   \end{center}
  
   \onslide<3->
-  {looks OK \ldots let's ship it to customers\hspace{5mm} 
+  {Looks OK \ldots let's ship it to customers\hspace{5mm} 
    \raisebox{-5mm}{\includegraphics[scale=0.05]{sun.png}}}
   
   \end{frame}}
@@ -329,7 +329,7 @@
   \bl{der c (d)}                   & \bl{$=$} & \bl{if c = d then [] else $\varnothing$} & \\
   \bl{der c (r$_1$ + r$_2$)}       & \bl{$=$} & \bl{(der c r$_1$) + (der c r$_2$)} & \\
   \bl{der c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{((der c r$_1$) $\cdot$ r$_2$)} & \\
-       &          & \bl{\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
+       &          & \bl{\;\;\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
   \bl{der c (r$^*$)}          & \bl{$=$} & \bl{(der c r) $\cdot$ r$^*$} &\smallskip\\
 
   \bl{derivative r []}     & \bl{$=$} & \bl{r} & \\
@@ -383,7 +383,7 @@
   \end{tabular}
   \end{center}
   \pause\pause\bigskip
-  ??? By \alert<4->{induction}, we can {\bf prove} these properties.\bigskip
+  By \alert<4->{induction}, we can {\bf prove} these properties.\bigskip
 
   \begin{tabular}{lrcl}
   Lemmas:  & \bl{nullable (r)}          & \bl{$\Longleftrightarrow$} & \bl{[] $\in$ \LL (r)}\\
@@ -524,7 +524,7 @@
   My point:\bigskip\\
 
   The theory about regular languages can be reformulated 
-  to be more suitable for theorem proving.
+  to be more\\ suitable for theorem proving.
   \end{tabular}
   \end{center}
   \end{frame}}
@@ -614,7 +614,7 @@
   \begin{center}
   \begin{tabular}{l}
   finite $\Rightarrow$ regular\\
-  \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_L) \Rightarrow \exists r. L = \mathbb{L}(r)}\\[3mm]
+  \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_L) \Rightarrow \exists r.\; L = \mathbb{L}(r)}\\[3mm]
   regular $\Rightarrow$ finite\\
   \;\;\;\smath{\text{finite}\, (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
   \end{tabular}
@@ -631,15 +631,16 @@
   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   \mode<presentation>{
   \begin{frame}[c]
-  \frametitle{\LARGE Final States}
+  \frametitle{\LARGE Final Equiv.~Classes}
 
   \mbox{}\\[3cm]
 
   \begin{itemize}
-  \item ??? \smath{\text{final}_L\,X \dn \{[|s|]_\approx\;|\; s \in X\}}\\
+  \item \smath{\text{finals}\,L \dn 
+     \{{\lbrack\mkern-2mu\lbrack{s}\rbrack\mkern-2mu\rbrack}_\approx\;|\; s \in L\}}\\
   \medskip
 
-  \item we can prove: \smath{L = \bigcup \{X\;|\;\text{final}_L\,X\}}
+  \item we can prove: \smath{L = \bigcup (\text{finals}\,L)}
 
   \end{itemize}
 
@@ -651,7 +652,7 @@
   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   \mode<presentation>{
   \begin{frame}[c]
-  \frametitle{\LARGE Transitions between\\[-3mm] Equivalence Classes}
+  \frametitle{\LARGE Transitions between ECs}
 
   \smath{L = \{[c]\}}
 
@@ -725,9 +726,9 @@
   & \smath{R_2} & \smath{\equiv} & \smath{R_1;a + R_2;a}\medskip\\
   \onslide<3->{we can prove} 
   & \onslide<3->{\smath{R_1}} & \onslide<3->{\smath{=}} 
-      & \onslide<3->{\smath{R_1; \mathbb{L}(b) \,\cup\, R_2;\mathbb{L}(b) \,\cup\, \{[]\};\{[]\}}}\\
+      & \onslide<3->{\smath{R_1;; \mathbb{L}(b) \,\cup\, R_2;;\mathbb{L}(b) \,\cup\, \{[]\}}}\\
   & \onslide<3->{\smath{R_2}} & \onslide<3->{\smath{=}}    
-      & \onslide<3->{\smath{R_1; \mathbb{L}(a) \,\cup\, R_2;\mathbb{L}(a)}}\\
+      & \onslide<3->{\smath{R_1;; \mathbb{L}(a) \,\cup\, R_2;;\mathbb{L}(a)}}\\
   \end{tabular}
   \end{center}
 
@@ -928,23 +929,45 @@
 *}
 
 
+
 text_raw {*
   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   \mode<presentation>{
   \begin{frame}[c]
-  \frametitle{\LARGE Other Direction}
-
+  \frametitle{\LARGE The Other Direction}
+  
   One has to prove
 
   \begin{center}
   \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}
   \end{center}
 
-  by induction on \smath{r}. Not trivial, but after a bit 
-  of thinking, one can prove that if
+  by induction on \smath{r}. This is straightforward for \\the base cases:\small
 
   \begin{center}
-  \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1)})}\hspace{5mm}
+  \begin{tabular}{l@ {\hspace{1mm}}l}
+  \smath{U\!N\!IV /\!/ \!\approx_{\emptyset}} & \smath{= \{U\!N\!IV\}}\smallskip\\
+  \smath{U\!N\!IV /\!/ \!\approx_{\{[]\}}} & \smath{\subseteq \{\{[]\}, U\!N\!IV - \{[]\}\}}\smallskip\\
+  \smath{U\!N\!IV /\!/ \!\approx_{\{[c]\}}} & \smath{\subseteq \{\{[]\}, \{[c]\}, U\!N\!IV - \{[], [c]\}\}}
+  \end{tabular}
+  \end{center}
+
+  
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[t]
+  \frametitle{\LARGE The Other Direction}
+
+  More complicated are the inductive cases:\\ one needs to prove that if
+
+  \begin{center}
+  \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1)})}\hspace{3mm}
   \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_2)})}
   \end{center}
 
@@ -954,12 +977,149 @@
   \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r_1) \,\cup\, \mathbb{L}(r_2)})}
   \end{center}
   
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[t]
+  \frametitle{\LARGE Helper Lemma}
+
+  \begin{center}
+  \begin{tabular}{p{10cm}}
+  %If \smath{\text{finite} (f\;' A)} and \smath{f} is injective 
+  %(on \smath{A}),\\ then \smath{\text{finite}\,A}.
+  Given two equivalence relations \smath{R_1} and \smath{R_2} with
+  \smath{R_1} refining \smath{R_2} (\smath{R_1 \subseteq R_2}).\\ 
+  Then\medskip\\
+  \smath{\;\;\text{finite} (U\!N\!IV /\!/ R_1) \Rightarrow \text{finite} (U\!N\!IV /\!/ R_2)}
+  \end{tabular}
+  \end{center}
+  
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{\Large Derivatives and Left-Quotients}
+  \small
+  Work by Brozowski ('64) and Antimirov ('96):\pause\smallskip
+
+
+  \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
+  \multicolumn{4}{@ {}l}{Left-Quotient:}\\
+  \multicolumn{4}{@ {}l}{\bl{$\text{Ders}\;\text{s}\,A \dn \{\text{s'} \;|\; \text{s @ s'} \in A\}$}}\bigskip\\
+
+  \multicolumn{4}{@ {}l}{Derivative:}\\
+  \bl{der c ($\varnothing$)}       & \bl{$=$} & \bl{$\varnothing$} & \\
+  \bl{der c ([])}                  & \bl{$=$} & \bl{$\varnothing$} & \\
+  \bl{der c (d)}                   & \bl{$=$} & \bl{if c = d then [] else $\varnothing$} & \\
+  \bl{der c (r$_1$ + r$_2$)}       & \bl{$=$} & \bl{(der c r$_1$) + (der c r$_2$)} & \\
+  \bl{der c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{((der c r$_1$) $\cdot$ r$_2$)} & \\
+       &          & \bl{\;\;\;\;+ (if nullable r$_1$ then der c r$_2$ else $\varnothing$)}\\
+  \bl{der c (r$^*$)}          & \bl{$=$} & \bl{(der c r) $\cdot$ r$^*$} &\smallskip\\
+
+  \bl{ders [] r}     & \bl{$=$} & \bl{r} & \\
+  \bl{ders (s @ [c]) r} & \bl{$=$} & \bl{der c (ders s r)} & \\
+  \end{tabular}\pause
+
+  \begin{center}
+  \alert{$\Rightarrow$}\smath{\;\;\text{Ders}\,\text{s}\,(\mathbb{L}(\text{r})) = \mathbb{L} (\text{ders s r})}
+  \end{center}
+  
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{\LARGE Left-Quotients and MN-Rels}
+
+  \begin{itemize}
+  \item \smath{x \approx_{A} y \,\dn\, \forall z.\; x @ z \in A \Leftrightarrow y @ z \in A}\medskip
+  \item \bl{$\text{Ders}\;s\,A \dn \{s' \;|\; s @ s' \in A\}$}
+  \end{itemize}\bigskip
+
+  \begin{center}
+  \smath{x \approx_A y  \Longleftrightarrow \text{Ders}\;x\;A = \text{Ders}\;y\;A}
+  \end{center}\bigskip\pause\small
+
+  which means
+
+  \begin{center}
+  \smath{x \approx_{\mathbb{L}(r)} y  \Longleftrightarrow 
+  \mathbb{L}(\text{ders}\;x\;r) = \mathbb{L}(\text{ders}\;y\;r)}
+  \end{center}\pause
+
+  \hspace{8.8mm}or
+  \smath{\;x \approx_{\mathbb{L}(r)} y  \Longleftarrow 
+  \text{ders}\;x\;r = \text{ders}\;y\;r}
+
   
 
   \end{frame}}
   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
 *}
 
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[c]
+  \frametitle{\LARGE Partial Derivatives}
+
+  Antimirov: \bl{pder : rexp $\Rightarrow$ rexp set}\bigskip
+
+  \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
+  \bl{pder c ($\varnothing$)}       & \bl{$=$} & \bl{\{$\varnothing$\}} & \\
+  \bl{pder c ([])}                  & \bl{$=$} & \bl{\{$\varnothing$\}} & \\
+  \bl{pder c (d)}                   & \bl{$=$} & \bl{if c = d then \{[]\} else \{$\varnothing$\}} & \\
+  \bl{pder c (r$_1$ + r$_2$)}       & \bl{$=$} & \bl{(pder c r$_1$) $\cup$ (pder c r$_2$)} & \\
+  \bl{pder c (r$_1$ $\cdot$ r$_2$)} & \bl{$=$} & \bl{(pder c r$_1$) $\odot$ r$_2$} & \\
+       &          & \bl{\hspace{-10mm}$\cup$ (if nullable r$_1$ then pder c r$_2$ else $\varnothing$)}\\
+  \bl{pder c (r$^*$)}          & \bl{$=$} & \bl{(pder c r) $\odot$ r$^*$} &\smallskip\\
+  \end{tabular}
+
+  \begin{tabular}{@ {}l@ {\hspace{2mm}}c@ {\hspace{2mm}}l@ {\hspace{-10mm}}l@ {}}
+  \bl{pders [] r}     & \bl{$=$} & \bl{r} & \\
+  \bl{pders (s @ [c]) r} & \bl{$=$} & \bl{pder c (pders s r)} & \\
+  \end{tabular}\pause
+
+  \begin{center}
+  \alert{$\Rightarrow$}\smath{\;\;\text{Ders}\,\text{s}\,(\mathbb{L}(\text{r})) = \bigcup (\mathbb{L}\;`\; (\text{pders s r}))}
+  \end{center}
+
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
+text_raw {*
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
+  \mode<presentation>{
+  \begin{frame}[t]
+  \frametitle{\LARGE Final Result}
+  
+  \mbox{}\\[7mm]
+
+  \begin{itemize}
+  \item \alt<1>{\smath{\text{pders x r \mbox{$=$} pders y r}}}
+            {\smath{\underbrace{\text{pders x r \mbox{$=$} pders y r}}_{R_1}}} 
+        refines \bl{x $\approx_{\mathbb{L}(\text{r})}$ y}\pause
+  \item \smath{\text{finite} (U\!N\!IV /\!/ R_1)} \bigskip\pause
+  \item Therefore \smath{\text{finite} (U\!N\!IV /\!/ \approx_{\mathbb{L}(r)})}. Qed.
+  \end{itemize}
+  
+  \end{frame}}
+  %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
+*}
+
 
 text_raw {*
   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%