major update of the nominal package; there is now an infrastructure
authorurbanc
Tue, 06 Mar 2007 15:28:22 +0100
changeset 22418 49e2d9744ae1
parent 22417 014e7696ac6b
child 22419 17441293ebc6
major update of the nominal package; there is now an infrastructure for equivariance lemmas which eases definitions of nominal functions
src/HOL/Nominal/Examples/Class.thy
src/HOL/Nominal/Examples/Compile.thy
src/HOL/Nominal/Examples/Crary.thy
src/HOL/Nominal/Examples/Fsub.thy
src/HOL/Nominal/Examples/Height.thy
src/HOL/Nominal/Examples/Lam_Funs.thy
src/HOL/Nominal/Examples/SN.thy
src/HOL/Nominal/Examples/Weakening.thy
src/HOL/Nominal/Nominal.thy
src/HOL/Nominal/nominal_atoms.ML
src/HOL/Nominal/nominal_permeq.ML
src/HOL/Nominal/nominal_thmdecls.ML
--- a/src/HOL/Nominal/Examples/Class.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Class.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -1,26 +1,37 @@
 (* $Id$ *)
 
 theory Class 
-imports "Nominal" 
+imports "../Nominal" 
 begin
 
 section {* Term-Calculus from Urban's PhD *}
 
 atom_decl name coname
 
+text {* types *}
+
+nominal_datatype ty =
+    PROP "string"
+  | NOT  "ty"
+  | AND  "ty" "ty"   ("_ AND _" [100,100] 100)
+  | OR   "ty" "ty"   ("_ OR _" [100,100] 100)
+  | IMP  "ty" "ty"   ("_ IMPL _" [100,100] 100)
+
+text {* terms *}
+
 nominal_datatype trm = 
     Ax   "name" "coname"
-  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            ("Cut <_>._ [_]._" [100,100,100,100] 100)
-  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 ("NotR [_]._ _" [100,100,100] 100)
+  | Cut  "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm"            ("Cut <_>._ (_)._" [100,100,100,100] 100)
+  | NotR "\<guillemotleft>name\<guillemotright>trm" "coname"                 ("NotR (_)._ _" [100,100,100] 100)
   | NotL "\<guillemotleft>coname\<guillemotright>trm" "name"                 ("NotL <_>._ _" [100,100,100] 100)
   | AndR "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>coname\<guillemotright>trm" "coname" ("AndR <_>._ <_>._ _" [100,100,100,100,100] 100)
-  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL1 [_]._ _" [100,100,100] 100)
-  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL2 [_]._ _" [100,100,100] 100)
+  | AndL1 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL1 (_)._ _" [100,100,100] 100)
+  | AndL2 "\<guillemotleft>name\<guillemotright>trm" "name"                  ("AndL2 (_)._ _" [100,100,100] 100)
   | OrR1 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR1 <_>._ _" [100,100,100] 100)
   | OrR2 "\<guillemotleft>coname\<guillemotright>trm" "coname"               ("OrR2 <_>._ _" [100,100,100] 100)
-  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        ("OrL [_]._ [_]._ _" [100,100,100,100,100] 100)
-  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       ("ImpR [_].<_>._ _" [100,100,100,100] 100)
-  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     ("ImpL <_>._ [_]._ _" [100,100,100,100,100] 100)
+  | OrL "\<guillemotleft>name\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"        ("OrL (_)._ (_)._ _" [100,100,100,100,100] 100)
+  | ImpR "\<guillemotleft>name\<guillemotright>(\<guillemotleft>coname\<guillemotright>trm)" "coname"       ("ImpR (_).<_>._ _" [100,100,100,100] 100)
+  | ImpL "\<guillemotleft>coname\<guillemotright>trm" "\<guillemotleft>name\<guillemotright>trm" "name"     ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100)
 
 text {* Induction principles *}
 
@@ -30,24 +41,12 @@
 
 text {* named terms *}
 
-nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("([_]:_)" [100,100] 100)
+nominal_datatype ntrm = Na "\<guillemotleft>name\<guillemotright>trm" ("((_):_)" [100,100] 100)
 
 text {* conamed terms *}
 
 nominal_datatype ctrm = Co "\<guillemotleft>coname\<guillemotright>trm" ("(<_>:_)" [100,100] 100)
 
-lemma eq_eqvt_name[eqvt]:
-  fixes pi::"name prm"
-  and   x::"'a::pt_name"
-  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
-  by (simp add: perm_bool perm_bij)
-
-lemma eq_eqvt_coname[eqvt]:
-  fixes pi::"coname prm"
-  and   x::"'a::pt_coname"
-  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
-  by (simp add: perm_bool perm_bij)
-
 text {* renaming functions *}
 
 consts
@@ -56,49 +55,232 @@
 
 nominal_primrec (freshness_context: "(d::coname,e::coname)") 
   "(Ax x a)[d\<turnstile>c>e] = (if a=d then Ax x e else Ax x a)" 
-  "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M [x].N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) [x].(N[d\<turnstile>c>e])" 
-  "(NotR [x].M a)[d\<turnstile>c>e] = (if a=d then NotR [x].(M[d\<turnstile>c>e]) e else NotR [x].(M[d\<turnstile>c>e]) a)" 
+  "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>M\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[d\<turnstile>c>e] = Cut <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e])" 
+  "(NotR (x).M a)[d\<turnstile>c>e] = (if a=d then NotR (x).(M[d\<turnstile>c>e]) e else NotR (x).(M[d\<turnstile>c>e]) a)" 
   "a\<sharp>(d,e) \<Longrightarrow> (NotL <a>.M x)[d\<turnstile>c>e] = (NotL <a>.(M[d\<turnstile>c>e]) x)" 
   "\<lbrakk>a\<sharp>(d,e,N,c);b\<sharp>(d,e,M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[d\<turnstile>c>e] = 
           (if c=d then AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d \<turnstile>c>e]) e else AndR <a>.(M[d\<turnstile>c>e]) <b>.(N[d\<turnstile>c>e]) c)" 
-  "x\<sharp>y \<Longrightarrow> (AndL1 [x].M y)[d\<turnstile>c>e] = AndL1 [x].(M[d\<turnstile>c>e]) y"
-  "x\<sharp>y \<Longrightarrow> (AndL2 [x].M y)[d\<turnstile>c>e] = AndL2 [x].(M[d\<turnstile>c>e]) y"
+  "x\<sharp>y \<Longrightarrow> (AndL1 (x).M y)[d\<turnstile>c>e] = AndL1 (x).(M[d\<turnstile>c>e]) y"
+  "x\<sharp>y \<Longrightarrow> (AndL2 (x).M y)[d\<turnstile>c>e] = AndL2 (x).(M[d\<turnstile>c>e]) y"
   "a\<sharp>(d,e,b) \<Longrightarrow> (OrR1 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR1 <a>.(M[d\<turnstile>c>e]) e else OrR1 <a>.(M[d\<turnstile>c>e]) b)"
   "a\<sharp>(d,e,b) \<Longrightarrow> (OrR2 <a>.M b)[d\<turnstile>c>e] = (if b=d then OrR2 <a>.(M[d\<turnstile>c>e]) e else OrR2 <a>.(M[d\<turnstile>c>e]) b)"
-  "\<lbrakk>x\<sharp>(N,z);y\<sharp>(M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL [x].M [y].N z)[d\<turnstile>c>e] = OrL [x].(M[d\<turnstile>c>e]) [y].(N[d\<turnstile>c>e]) z"
-  "a\<sharp>(d,e,b) \<Longrightarrow> (ImpR [x].<a>.M b)[d\<turnstile>c>e] = 
-       (if b=d then ImpR [x].<a>.(M[d\<turnstile>c>e]) e else ImpR [x].<a>.(M[d\<turnstile>c>e]) b)"
-  "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>(M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M [x].N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) [x].(N[d\<turnstile>c>e]) y"
-apply(finite_guess add: eqvt perm_if fs_coname1 fs_name1 | 
-      perm_simp add: abs_fresh abs_supp fs_name1 fs_coname1)+
-apply(fresh_guess add: eqvt perm_if fs_coname1 fs_name1 | perm_simp add: fresh_atm)+
+  "\<lbrakk>x\<sharp>(N,z);y\<sharp>(M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[d\<turnstile>c>e] = OrL (x).(M[d\<turnstile>c>e]) (y).(N[d\<turnstile>c>e]) z"
+  "a\<sharp>(d,e,b) \<Longrightarrow> (ImpR (x).<a>.M b)[d\<turnstile>c>e] = 
+       (if b=d then ImpR (x).<a>.(M[d\<turnstile>c>e]) e else ImpR (x).<a>.(M[d\<turnstile>c>e]) b)"
+  "\<lbrakk>a\<sharp>(d,e,N);x\<sharp>(M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[d\<turnstile>c>e] = ImpL <a>.(M[d\<turnstile>c>e]) (x).(N[d\<turnstile>c>e]) y"
+apply(finite_guess)+
+apply(rule TrueI)+
+apply(simp add: abs_fresh abs_supp fin_supp)+
+apply(fresh_guess)+
 done
 
 nominal_primrec (freshness_context: "(u::name,v::name)") 
   "(Ax x a)[u\<turnstile>n>v] = (if x=u then Ax v a else Ax x a)" 
-  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M [x].N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v])" 
-  "x\<sharp>(u,v) \<Longrightarrow> (NotR [x].M a)[u\<turnstile>n>v] = NotR [x].(M[u\<turnstile>n>v]) a" 
+  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M)\<rbrakk> \<Longrightarrow> (Cut <a>.M (x).N)[u\<turnstile>n>v] = Cut <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v])" 
+  "x\<sharp>(u,v) \<Longrightarrow> (NotR (x).M a)[u\<turnstile>n>v] = NotR (x).(M[u\<turnstile>n>v]) a" 
   "(NotL <a>.M x)[u\<turnstile>n>v] = (if x=u then NotL <a>.(M[u\<turnstile>n>v]) v else NotL <a>.(M[u\<turnstile>n>v]) x)" 
   "\<lbrakk>a\<sharp>(N,c);b\<sharp>(M,c);b\<noteq>a\<rbrakk> \<Longrightarrow> (AndR <a>.M <b>.N c)[u\<turnstile>n>v] = AndR <a>.(M[u\<turnstile>n>v]) <b>.(N[u\<turnstile>n>v]) c" 
-  "x\<sharp>(u,v,y) \<Longrightarrow> (AndL1 [x].M y)[u\<turnstile>n>v] = (if y=u then AndL1 [x].(M[u\<turnstile>n>v]) v else AndL1 [x].(M[u\<turnstile>n>v]) y)"
-  "x\<sharp>(u,v,y) \<Longrightarrow> (AndL2 [x].M y)[u\<turnstile>n>v] = (if y=u then AndL2 [x].(M[u\<turnstile>n>v]) v else AndL2 [x].(M[u\<turnstile>n>v]) y)"
+  "x\<sharp>(u,v,y) \<Longrightarrow> (AndL1 (x).M y)[u\<turnstile>n>v] = (if y=u then AndL1 (x).(M[u\<turnstile>n>v]) v else AndL1 (x).(M[u\<turnstile>n>v]) y)"
+  "x\<sharp>(u,v,y) \<Longrightarrow> (AndL2 (x).M y)[u\<turnstile>n>v] = (if y=u then AndL2 (x).(M[u\<turnstile>n>v]) v else AndL2 (x).(M[u\<turnstile>n>v]) y)"
   "a\<sharp>b \<Longrightarrow> (OrR1 <a>.M b)[u\<turnstile>n>v] = OrR1 <a>.(M[u\<turnstile>n>v]) b"
   "a\<sharp>b \<Longrightarrow> (OrR2 <a>.M b)[u\<turnstile>n>v] = OrR2 <a>.(M[u\<turnstile>n>v]) b"
-  "\<lbrakk>x\<sharp>(u,v,N,z);y\<sharp>(u,v,M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL [x].M [y].N z)[u\<turnstile>n>v] = 
-        (if z=u then OrL [x].(M[u\<turnstile>n>v]) [y].(N[u\<turnstile>n>v]) v else OrL [x].(M[u\<turnstile>n>v]) [y].(N[u\<turnstile>n>v]) z)"
-  "\<lbrakk>a\<sharp>b; x\<sharp>(u,v)\<rbrakk> \<Longrightarrow> (ImpR [x].<a>.M b)[u\<turnstile>n>v] = ImpR [x].<a>.(M[u\<turnstile>n>v]) b"
-  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M [x].N y)[u\<turnstile>n>v] = 
-        (if y=u then ImpL <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) [x].(N[u\<turnstile>n>v]) y)"
-apply(finite_guess add: eqvt perm_if fs_coname1 fs_name1 | 
-      perm_simp add: abs_fresh abs_supp fresh_prod fs_name1 fs_coname1)+
-apply(fresh_guess add: eqvt perm_if fs_coname1 fs_name1 | perm_simp add: fresh_atm)+
+  "\<lbrakk>x\<sharp>(u,v,N,z);y\<sharp>(u,v,M,z);y\<noteq>x\<rbrakk> \<Longrightarrow> (OrL (x).M (y).N z)[u\<turnstile>n>v] = 
+        (if z=u then OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) v else OrL (x).(M[u\<turnstile>n>v]) (y).(N[u\<turnstile>n>v]) z)"
+  "\<lbrakk>a\<sharp>b; x\<sharp>(u,v)\<rbrakk> \<Longrightarrow> (ImpR (x).<a>.M b)[u\<turnstile>n>v] = ImpR (x).<a>.(M[u\<turnstile>n>v]) b"
+  "\<lbrakk>a\<sharp>N;x\<sharp>(u,v,M,y)\<rbrakk> \<Longrightarrow> (ImpL <a>.M (x).N y)[u\<turnstile>n>v] = 
+        (if y=u then ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) v else ImpL <a>.(M[u\<turnstile>n>v]) (x).(N[u\<turnstile>n>v]) y)"
+apply(finite_guess)+
+apply(rule TrueI)+
+apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1)+
+apply(fresh_guess)+
+done
+
+lemmas eq_bij = pt_bij[OF pt_name_inst, OF at_name_inst] pt_bij[OF pt_coname_inst, OF at_coname_inst]
+
+
+lemma crename_name_eqvt[eqvt]:
+  fixes pi::"name prm"
+  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
+apply(nominal_induct M avoiding: d e rule: trm.induct)
+apply(auto simp add: fresh_bij eq_bij)
+done
+
+
+lemma crename_coname_eqvt[eqvt]:
+  fixes pi::"coname prm"
+  shows "pi\<bullet>(M[d\<turnstile>c>e]) = (pi\<bullet>M)[(pi\<bullet>d)\<turnstile>c>(pi\<bullet>e)]"
+apply(nominal_induct M avoiding: d e rule: trm.induct)
+apply(auto simp add: fresh_bij eq_bij)
+done
+
+lemma nrename_name_eqvt[eqvt]:
+  fixes pi::"name prm"
+  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
+apply(nominal_induct M avoiding: x y rule: trm.induct)
+apply(auto simp add: fresh_bij eq_bij)
+done
+
+lemma nrename_coname_eqvt[eqvt]:
+  fixes pi::"coname prm"
+  shows "pi\<bullet>(M[x\<turnstile>n>y]) = (pi\<bullet>M)[(pi\<bullet>x)\<turnstile>n>(pi\<bullet>y)]"
+apply(nominal_induct M avoiding: x y rule: trm.induct)
+apply(auto simp add: fresh_bij eq_bij)
+done
+
+text {* substitution functions *}
+
+consts
+  substn :: "trm \<Rightarrow> name   \<Rightarrow> coname \<Rightarrow> trm \<Rightarrow> trm" ("_[_:=<_>._]" [100,100,100,100] 100) 
+  substc :: "trm \<Rightarrow> coname \<Rightarrow> name   \<Rightarrow> trm \<Rightarrow> trm" ("_[_:=(_)._]" [100,100,100,100] 100)
+
+nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)")
+  "(Ax x a)[y:=<c>.P] = (if x=y then P[c\<turnstile>c>a] else Ax x a)" 
+  "\<lbrakk>a\<sharp>(c,P,N);x\<sharp>(y,c,P,M)\<rbrakk> \<Longrightarrow> 
+   (Cut <a>.M (x).N)[y:=<c>.P] = Cut <a>.(M[y:=<c>.P]) (x).(N[y:=<c>.P])" 
+  "x\<sharp>(y,c,P) \<Longrightarrow> 
+   (NotR (x).M a)[y:=<c>.P] = NotR (x).(M[y:=<c>.P]) a" 
+  "a\<sharp>(c,P)\<Longrightarrow>
+   (NotL <a>.M x)[y:=<c>.P] = (if x=y then Cut <c>.P (x).(NotL <a>. (M[y:=<c>.P]) x) 
+                                      else NotL <a>. (M[y:=<c>.P]) x)" 
+  "\<lbrakk>a\<sharp>(c,P,N,d);b\<sharp>(c,P,M,d);b\<noteq>a\<rbrakk> \<Longrightarrow> 
+   (AndR <a>.M <b>.N d)[y:=<c>.P] = AndR <a>.(M[y:=<c>.P]) <b>.(N[y:=<c>.P]) d" 
+  "x\<sharp>(y,c,P,z) \<Longrightarrow> 
+   (AndL1 (x).M z)[y:=<c>.P] = (if z=y then Cut <c>.P (z).AndL1 (x).(M[y:=<c>.P]) z 
+                                       else AndL1 (x).(M[y:=<c>.P]) z)"
+  "x\<sharp>(y,c,P,z) \<Longrightarrow> 
+   (AndL2 (x).M z)[y:=<c>.P] = (if z=y then Cut <c>.P (z).AndL2 (x).(M[y:=<c>.P]) z 
+                                       else AndL2 (x).(M[y:=<c>.P]) z)"
+  "a\<sharp>(c,P,b) \<Longrightarrow> 
+   (OrR1 <a>.M b)[y:=<c>.P] = OrR1 <a>.(M[y:=<c>.P]) b"
+  "a\<sharp>(c,P,b) \<Longrightarrow> 
+   (OrR2 <a>.M b)[y:=<c>.P] = OrR2 <a>.(M[y:=<c>.P]) b"
+  "\<lbrakk>x\<sharp>(y,N,c,P,z);u\<sharp>(y,M,c,P,z);x\<noteq>u\<rbrakk> \<Longrightarrow> 
+   (OrL (x).M (u).N z)[y:=<c>.P] = (if z=y then Cut <c>.P (z).(OrL (x).(M[y:=<c>.P]) (u).(N[y:=<c>.P]) z) 
+                                           else OrL (x).(M[y:=<c>.P]) (u).(N[y:=<c>.P]) z)"
+  "\<lbrakk>a\<sharp>(b,c,P); x\<sharp>(y,c,P)\<rbrakk> \<Longrightarrow> 
+   (ImpR (x).<a>.M b)[y:=<c>.P] = ImpR (x).<a>.(M[y:=<c>.P]) b"
+  "\<lbrakk>a\<sharp>(N,c,P);x\<sharp>(y,c,P,M,z)\<rbrakk> \<Longrightarrow> 
+   (ImpL <a>.M (x).N z)[y:=<c>.P] = (if y=z then Cut <c>.P (z).(ImpL <a>.(M[y:=<c>.P]) (x).(N[y:=<c>.P]) z) 
+                                            else ImpL <a>.(M[y:=<c>.P]) (x).(N[y:=<c>.P]) z)"
+apply(finite_guess)+
+apply(rule TrueI)+
+apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1)+
+apply(fresh_guess)+
 done
 
-text {* We should now define the two forms of substitition :o( *}
+text {* typing contexts *}
+
+types 
+  ctxtn = "(name\<times>ty) list"
+  ctxtc = "(coname\<times>ty) list"
+
+inductive2
+  validc :: "ctxtc \<Rightarrow> bool"
+where
+  vc1[intro]: "validc []"
+| vc2[intro]: "\<lbrakk>a\<sharp>\<Delta>; validc \<Delta>\<rbrakk> \<Longrightarrow> validc ((a,T)#\<Delta>)"  
+
+inductive2
+  validn :: "ctxtn \<Rightarrow> bool"
+where
+  vn1[intro]: "validn []"
+| vn2[intro]: "\<lbrakk>x\<sharp>\<Gamma>; validn \<Gamma>\<rbrakk> \<Longrightarrow> validn ((x,T)#\<Gamma>)"
+
+text {* typing relation *}
+
+inductive2
+   typing :: "ctxtn \<Rightarrow> trm \<Rightarrow> ctxtc \<Rightarrow> bool" ("_ \<turnstile> _ \<turnstile> _" [100,100,100] 100)
+where
+  "\<lbrakk>validn \<Gamma>;validc \<Delta>; (x,B)\<in>set \<Gamma>; (a,B)\<in>set \<Delta>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Ax x a \<turnstile> \<Delta>"
+| "\<lbrakk>x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> NotR (x).M a \<turnstile> ((a,NOT B)#\<Delta>)"
+| "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>)\<rbrakk> \<Longrightarrow>  ((x,NOT B)#\<Gamma>) \<turnstile> NotL <a>.M x \<turnstile> \<Delta>"
+| "\<lbrakk>x\<sharp>\<Gamma>; ((x,B1)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>\<rbrakk> \<Longrightarrow>  ((y,B1 AND B2)#\<Gamma>) \<turnstile> AndL1 (x).M y \<turnstile> \<Delta>"
+| "\<lbrakk>x\<sharp>\<Gamma>; ((x,B2)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>\<rbrakk> \<Longrightarrow>  ((y,B1 AND B2)#\<Gamma>) \<turnstile> AndL2 (x).M y \<turnstile> \<Delta>"
+| "\<lbrakk>a\<sharp>\<Delta>;b\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); \<Gamma> \<turnstile> N \<turnstile> ((b,C)#\<Delta>)\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> AndR <a>.M <b>.N c \<turnstile> ((c,B AND C)#\<Delta>)"
+| "\<lbrakk>x\<sharp>\<Gamma>;y\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> \<Delta>; ((y,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>\<rbrakk> \<Longrightarrow>  ((z,B OR C)#\<Gamma>) \<turnstile> OrL (x).M (y).N z \<turnstile> \<Delta>"
+| "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B1)#\<Delta>)\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> OrR1 <a>.M b \<turnstile> ((b,B1 OR B2)#\<Delta>)"
+| "\<lbrakk>a\<sharp>\<Delta>; \<Gamma> \<turnstile> M \<turnstile> ((a,B2)#\<Delta>)\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> OrR2 <a>.M b \<turnstile> ((b,B1 OR B2)#\<Delta>)"
+| "\<lbrakk>a\<sharp>\<Delta>;x\<sharp>\<Gamma>; \<Gamma> \<turnstile> M \<turnstile> ((a,B)#\<Delta>); ((x,C)#\<Gamma>) \<turnstile> N \<turnstile> \<Delta>\<rbrakk> \<Longrightarrow>  ((y,B IMPL C)#\<Gamma>) \<turnstile> ImpL <a>.M (x).N y \<turnstile> \<Delta>"
+| "\<lbrakk>a\<sharp>\<Delta>;x\<sharp>\<Gamma>; ((x,B)#\<Gamma>) \<turnstile> M \<turnstile> ((a,C)#\<Delta>)\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> ImpR (x).<a>.M b \<turnstile> ((b,B IMPL C)#\<Delta>)"
+| "\<lbrakk>a\<sharp>\<Delta>1;x\<sharp>\<Gamma>2; validn(\<Gamma>1@\<Gamma>2); validc (\<Delta>1@\<Delta>2); \<Gamma>1 \<turnstile> M \<turnstile> ((a,B)#\<Delta>1); ((x,B)#\<Gamma>2) \<turnstile> N \<turnstile> \<Delta>2\<rbrakk> 
+   \<Longrightarrow>  (\<Gamma>1@\<Gamma>2) \<turnstile> Cut <a>.M (x).N \<turnstile> (\<Delta>1@\<Delta>2)"
+
+text {* relations about freshly introducing a name or coname *} 
+
+inductive2
+  fin :: "trm \<times> name \<Rightarrow> bool"
+where
+  "fin (Ax x a,x)"
+| "x\<sharp>M \<Longrightarrow> fin (NotL <a>.M x,x)"
+| "y\<sharp>[x].M \<Longrightarrow> fin (AndL1 (x).M y,y)"
+| "y\<sharp>[x].M \<Longrightarrow> fin (AndL2 (x).M y,y)"
+| "\<lbrakk>z\<sharp>[x].M;z\<sharp>[y].N\<rbrakk> \<Longrightarrow> fin (OrL (x).M (y).N z,z)"
+| "\<lbrakk>y\<sharp>M;y\<sharp>[x].N\<rbrakk> \<Longrightarrow> fin (ImpL <a>.M (x).N y,y)"
 
-consts
-  substn :: "trm \<Rightarrow> name   \<Rightarrow> ctrm \<Rightarrow> trm" ("_[_::n=_]" [100,100,100] 100) 
-  substc :: "trm \<Rightarrow> coname \<Rightarrow> ntrm \<Rightarrow> trm" ("_[_::c=_]" [100,100,100] 100)
+inductive2
+  fic :: "trm \<times> coname \<Rightarrow> bool"
+where
+  "fic (Ax x a,a)"
+| "a\<sharp>M \<Longrightarrow> fic (NotR (x).M a,a)"
+| "\<lbrakk>c\<sharp>[a].M;c\<sharp>[b].N\<rbrakk> \<Longrightarrow> fic (AndR <a>.M <b>.N c,c)"
+| "b\<sharp>[a].M \<Longrightarrow> fic (OrR1 <a>.M b,b)"
+| "b\<sharp>[a].M \<Longrightarrow> fic (OrR2 <a>.M b,b)"
+| "\<lbrakk>b\<sharp>[a].M\<rbrakk> \<Longrightarrow> fic (ImpR (x).<a>.M b,b)"
+
+text {* cut-reductions *}
+
+inductive2
+  red :: "trm \<Rightarrow> trm \<Rightarrow> bool" ("_ \<longrightarrow>\<^isub>c _" [100,100] 100)
+where
+  "fic (M,a) \<Longrightarrow> Cut <a>.M (x).(Ax x b) \<longrightarrow>\<^isub>c M[a\<turnstile>c>b]"
+| "fin (M,x) \<Longrightarrow> Cut <a>.(Ax y a) (x).M \<longrightarrow>\<^isub>c M[x\<turnstile>n>y]"
+| "\<lbrakk>fic (NotR (x).M a,a); fin (NotL <b>.N y,y)\<rbrakk> \<Longrightarrow>
+   Cut <a>.(NotR (x).M a) (y).(NotL <b>.N y) \<longrightarrow>\<^isub>c Cut <b>.N (x).M" 
+| "\<lbrakk>fic (AndR <a1>.M1 <a2>.M2 b,b); fin (AndL1 (x).N y,y)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>c Cut <a1>.M1 (x).N"
+| "\<lbrakk>fic (AndR <a1>.M1 <a2>.M2 b,b); fin (AndL2 (x).N y,y)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL2 (x).N y) \<longrightarrow>\<^isub>c Cut <a2>.M2 (x).N"
+| "\<lbrakk>fic (AndR <a1>.M1 <a2>.M2 b,b); fin (AndL1 (x).N y,y)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(AndR <a1>.M1 <a2>.M2 b) (y).(AndL1 (x).N y) \<longrightarrow>\<^isub>c Cut <a1>.M1 (x).N"
+| "\<lbrakk>fic (OrR1 <a>.M b,b); fin (OrL (x1).N1 (x2).N2 y,y)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(OrR1 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>c Cut <a>.M (x1).N1"
+| "\<lbrakk>fic (OrR2 <a>.M b,b); fin (OrL (x1).N1 (x2).N2 y,y)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(OrR2 <a>.M b) (y).(OrL (x1).N1 (x2).N2 y) \<longrightarrow>\<^isub>c Cut <a>.M (x2).N2"
+| "\<lbrakk>fin (ImpL <c>.N (y).P z,z); fic (ImpR (x).<a>.M b,b)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>c Cut <a>.(Cut <c>.N (x).M) (y).P"
+| "\<lbrakk>fin (ImpL <c>.N (y).P z,z); fic (ImpR (x).<a>.M b,b)\<rbrakk> \<Longrightarrow>
+   Cut <b>.(ImpR (x).<a>.M b) (z).(ImpL <c>.N (y).P z) \<longrightarrow>\<^isub>c Cut <a>.N (x).(Cut <a>.M (y).P)"
+| "\<not>fic (M,a) \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c M[a:=(x).N]"
+| "\<not>fin (N,x) \<Longrightarrow> Cut <a>.M (x).N \<longrightarrow>\<^isub>c N[x:=<a>.M]"
+
+text {* PROOFS *}
+
+lemma validn_eqvt[eqvt]:
+  fixes   pi1:: "name prm"
+  and     pi2:: "coname prm"
+  assumes a: "validn \<Gamma>"
+  shows   "validn (pi1\<bullet>\<Gamma>)" and "validn (pi2\<bullet>\<Gamma>)"
+using a by (induct) (auto simp add: fresh_bij)
+
+lemma validc_eqvt[eqvt]:
+  fixes   pi1:: "name prm"
+  and     pi2:: "coname prm"
+  assumes a: "validc \<Delta>"
+  shows   "validc (pi1\<bullet>\<Delta>)" and "validc (pi2\<bullet>\<Delta>)"
+using a by (induct) (auto simp add: fresh_bij)
+
+text {* Weakening Lemma *}
+
+abbreviation
+  "subn" :: "ctxtn \<Rightarrow> ctxtn \<Rightarrow> bool" (" _ \<lless>n _ " [80,80] 80) 
+where
+  "\<Gamma>1 \<lless>n \<Gamma>2 \<equiv> \<forall>x B. (x,B)\<in>set \<Gamma>1 \<longrightarrow> (x,B)\<in>set \<Gamma>2"
+
+abbreviation
+  "subc" :: "ctxtc \<Rightarrow> ctxtc \<Rightarrow> bool" (" _ \<lless>c _ " [80,80] 80) 
+where
+  "\<Delta>1 \<lless>c \<Delta>2 \<equiv> \<forall>a B. (a,B)\<in>set \<Delta>1 \<longrightarrow> (a,B)\<in>set \<Delta>2"
+
 
 end
 
--- a/src/HOL/Nominal/Examples/Compile.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Compile.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -3,7 +3,7 @@
 (* A challenge suggested by Adam Chlipala *)
 
 theory Compile
-imports "Nominal"
+imports "../Nominal"
 begin
 
 atom_decl name 
@@ -84,19 +84,6 @@
 
 text {* capture-avoiding substitution *}
 
-lemma perm_if:
-  fixes pi::"'a prm"
-  shows "pi\<bullet>(if b then c1 else c2) = (if (pi\<bullet>b) then (pi\<bullet>c1) else (pi\<bullet>c2))"
-apply(simp add: perm_fun_def)
-done
-
-lemma eq_eqvt:
-  fixes pi::"name prm"
-  and   x::"'a::pt_name"
-  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
-  apply(simp add: perm_bool perm_bij)
-  done
-
 consts
   subst :: "'a \<Rightarrow> name \<Rightarrow> 'a \<Rightarrow> 'a"  ("_[_::=_]" [100,100,100] 100)
 
@@ -113,54 +100,10 @@
   "\<lbrakk>z\<noteq>x; x\<sharp>y; x\<sharp>e; x\<sharp>e2; z\<sharp>y; z\<sharp>e; z\<sharp>e1; x\<sharp>t'; z\<sharp>t'\<rbrakk> \<Longrightarrow>
      (Case e of inl x \<rightarrow> e1 | inr z \<rightarrow> e2)[y::=t'] =
        (Case (e[y::=t']) of inl x \<rightarrow> (e1[y::=t']) | inr z \<rightarrow> (e2[y::=t']))"
-  apply (finite_guess add: fs_name1 perm_if eq_eqvt)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess only: fs_name1)
-  apply perm_simp
-  apply (simp add: supp_unit)
-  apply (rule TrueI)+
-  apply (simp add: abs_fresh)
-  apply (simp_all add: abs_fresh)
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess only: fs_name1)
-  apply perm_simp
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess only: fs_name1)
-  apply perm_simp
-  apply (fresh_guess only: fs_name1)
-  apply perm_simp
+  apply(finite_guess)+
+  apply(rule TrueI)+
+  apply(simp add: abs_fresh)+
+  apply(fresh_guess)+
   done
 
 nominal_primrec (Isubst)
@@ -174,41 +117,20 @@
   "(IRef e)[y::=t'] = IRef (e[y::=t'])"
   "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])"
   "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])"
-  apply (finite_guess add: fs_name1 perm_if eq_eqvt)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess only: fs_name1)
-  apply perm_simp
-  apply (simp add: supp_unit)
-  apply (rule TrueI)+
-  apply (simp add: abs_fresh)
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess only: fs_name1)
-  apply perm_simp
+  apply(finite_guess)+
+  apply(rule TrueI)+
+  apply(simp add: abs_fresh)+
+  apply(fresh_guess)+
   done
 
-lemma Isubst_eqvt:
+lemma Isubst_eqvt[eqvt]:
   fixes pi::"name prm"
   and   t1::"trmI"
   and   t2::"trmI"
   and   x::"name"
   shows "pi\<bullet>(t1[x::=t2]) = ((pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)])"
   apply (nominal_induct t1 avoiding: x t2 rule: trmI.induct)
-  apply (simp_all add: Isubst.simps perm_if eq_eqvt fresh_bij)
+  apply (simp_all add: Isubst.simps eqvt fresh_bij)
   done
 
 lemma Isubst_supp:
@@ -294,50 +216,10 @@
         let v1 = (trans e1) in
         let v2 = (trans e2) in 
         Iif (IRef (ISucc v)) (v2[x2::=IRef (ISucc (ISucc v))]) (v1[x1::=IRef (ISucc (ISucc v))]))"
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1 Let_def perm_nat_def)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1)
-  apply (finite_guess add: fs_name1 Let_def perm_nat_def)
-  apply (finite_guess add: fs_name1 Let_def perm_nat_def)
-  apply (finite_guess add: fs_name1 Let_def Isubst_eqvt)
-  apply (simp add: supp_unit)
-  apply (rule TrueI)+
-  apply (simp add: abs_fresh)
-  apply (simp_all add: abs_fresh Isubst_fresh)
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def Isubst_eqvt)
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1 perm_if eq_eqvt)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def perm_nat_def)
-  apply (fresh_guess add: fs_name1 Let_def Isubst_eqvt)
-  apply (fresh_guess add: fs_name1 Let_def Isubst_eqvt)
+  apply(finite_guess add: Let_def)+
+  apply(rule TrueI)+
+  apply(simp add: abs_fresh Isubst_fresh)+
+  apply(fresh_guess add: Let_def)+
   done
 
 consts trans_type :: "ty \<Rightarrow> tyI"
--- a/src/HOL/Nominal/Examples/Crary.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Crary.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -8,43 +8,68 @@
 (* The formalisation was done by Julien Narboux and   *)
 (* Christian Urban                                    *)
 
+(*<*)
 theory Crary
-  imports "Nominal"
+  imports "../Nominal"
 begin
 
+(* We put this def here because of some incompatibility of LatexSugar and Nominal *)
+
+syntax (Rule output)
+  "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:\mbox{}\inferrule{\mbox{>_\<^raw:}}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:\mbox{}\inferrule{>_\<^raw:}>\<^raw:{\mbox{>_\<^raw:}}>")
+
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" 
+  ("\<^raw:\mbox{>_\<^raw:}\\>/ _")
+
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+syntax (IfThen output)
+  "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\rmfamily\upshapce\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}> /\<^raw:{\rmfamily\upshape\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("\<^raw:\mbox{>_\<^raw:}>")
+
+syntax (IfThenNoBox output)
+  "==>" :: "prop \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _/ \<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
+  "_bigimpl" :: "asms \<Rightarrow> prop \<Rightarrow> prop"
+  ("\<^raw:{\rmfamily\upshape\normalsize{}>If\<^raw:\,}> _ /\<^raw:{\rmfamily\upshape\normalsize \,>then\<^raw:\,}>/ _.")
+  "_asms" :: "prop \<Rightarrow> asms \<Rightarrow> asms" ("_ /\<^raw:{\rmfamily\upshape\normalsize \,>and\<^raw:\,}>/ _")
+  "_asm" :: "prop \<Rightarrow> asms" ("_")
+
+
+(*>*)
+
+text {* 
+\section{Definition of the language} 
+\subsection{Definition of the terms and types} 
+
+First we define the type of atom names which will be used for binders.
+Each atom type is infinitely many atoms and equality is decidable. *}
+
 atom_decl name 
 
+text {* We define the datatype representing types. Although, It does not contain any binder we still use the \texttt{nominal\_datatype} command because the Nominal datatype package will prodive permutation functions and useful lemmas. *}
+
 nominal_datatype ty = TBase 
                     | TUnit 
                     | Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
 
+text {* The datatype of terms contains a binder. The notation @{text "\<guillemotleft>name\<guillemotright> trm"} means that the name is bound inside trm. *}
+
 nominal_datatype trm = Unit
                      | Var "name"
                      | Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
                      | App "trm" "trm"
                      | Const "nat"
 
-(* The next 3 lemmas should be in the nominal library *)
-lemma eq_eqvt:
-  fixes pi::"name prm"
-  and   x::"'a::pt_name"
-  shows "pi\<bullet>(x=y) = (pi\<bullet>x=pi\<bullet>y)"
-  apply(simp add: perm_bool perm_bij)
-  done
-
-lemma in_eqvt:
-  fixes pi::"name prm"
-  and   x::"'a::pt_name"
-  assumes "x\<in>X"
-  shows "pi\<bullet>x \<in> pi\<bullet>X"
-  using assms by (perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
-
-lemma set_eqvt:
-  fixes pi::"name prm"
-  and   xs::"('a::pt_name) list"
-  shows "pi\<bullet>(set xs) = set (pi\<bullet>xs)"
-  by (perm_simp add: pt_list_set_pi[OF pt_name_inst])
-(* end of library *)
+text {* As the datatype of types does not contain any binder, the application of a permutation is the identity function. In the future, this will be automatically derived by the package. *}
 
 lemma perm_ty[simp]: 
   fixes T::"ty"
@@ -60,17 +85,23 @@
 
 lemma ty_cases:
   fixes T::ty
-  shows "(\<exists> T1 T2. T=T1\<rightarrow>T2) \<or> T=TUnit \<or> T=TBase"
+  shows "(\<exists> T\<^isub>1 T\<^isub>2. T=T\<^isub>1\<rightarrow>T\<^isub>2) \<or> T=TUnit \<or> T=TBase"
 by (induct T rule:ty.induct_weak) (auto)
 
-text {* Size Functions *} 
+text {* 
+\subsection{Size functions}
+
+We define size functions for types and terms. As Isabelle allows overloading we can use the same notation for both functions.
+
+These function are automatically generated for non nominal datatypes. In the future, we need to extend the package to generate size function for nominal datatypes as well.
+ *} 
 
 instance ty :: size ..
 
 nominal_primrec
   "size (TBase) = 1"
   "size (TUnit) = 1"
-  "size (T1\<rightarrow>T2) = size T1 + size T2"
+  "size (T\<^isub>1\<rightarrow>T\<^isub>2) = size T\<^isub>1 + size T\<^isub>2"
 by (rule TrueI)+
 
 instance trm :: size ..
@@ -81,12 +112,10 @@
   "size (Lam [x].t) = size t + 1"
   "size (App t1 t2) = size t1 + size t2"
   "size (Const n) = 1"
-apply(finite_guess add: fs_name1 perm_nat_def)+
-apply(perm_full_simp add: perm_nat_def)
-apply(simp add: fs_name1)
+apply(finite_guess)+
 apply(rule TrueI)+
 apply(simp add: fresh_nat)+
-apply(fresh_guess add: fs_name1 perm_nat_def)+
+apply(fresh_guess)+
 done
 
 lemma ty_size_greater_zero[simp]:
@@ -94,56 +123,136 @@
   shows "size T > 0"
 by (nominal_induct rule:ty.induct) (simp_all)
 
-text {* valid contexts *}
+text {* 
+\subsection{Typing}
+
+\subsubsection{Typing contexts}
+
+This section contains the definition and some properties of a typing context. As the concept of context often appears in the litterature and is general, we should in the future provide these lemmas in a library.
 
+\paragraph{Definition of the Validity of contexts}\strut\\
+
+First we define what valid contexts are. Informally a context is valid is it does not contains twice the same variable.
+
+We use the following two inference rules: *}
+
+(*<*)
 inductive2
   valid :: "(name \<times> 'a::pt_name) list \<Rightarrow> bool"
 where
     v_nil[intro]:  "valid []"
   | v_cons[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"
+(*>*)
 
-lemma valid_eqvt[eqvt]:
-  fixes   pi:: "name prm"
-  assumes a: "valid \<Gamma>"
-  shows "valid (pi\<bullet>\<Gamma>)"
-  using a
-  by (induct) (auto simp add: fresh_bij)
+text {*
+\begin{center}
+\isastyle
+@{thm[mode=Rule] v_nil[no_vars]}{\sc{v\_nil}}
+\qquad
+@{thm[mode=Rule] v_cons[no_vars]}{\sc{v\_cons}}
+\end{center}
+*} 
+
+text{* We generate the equivariance lemma for the relation \texttt{valid}. *}
+
+nominal_inductive valid
+
+text{* We obtain a lemma called \texttt{valid\_eqvt}:
+\begin{center} 
+@{thm[mode=IfThen] valid_eqvt}
+\end{center}
+*}
+
+
+text{* We generate the inversion lemma for non empty lists. We add the \texttt{elim} attribute to tell the automated tactics to use it.
+ *}
 
 inductive_cases2  
   valid_cons_elim_auto[elim]:"valid ((x,T)#\<Gamma>)"
 
-text {* typing judgements for terms *}
+text{*
+The generated theorem is the following:
+\begin{center}
+@{thm "valid_cons_elim_auto"[no_vars] }
+\end{center}
+*}
+
+text{* \paragraph{Lemmas about valid contexts} 
+ Now, we can prove two useful lemmas about valid contexts. 
+*}
+
+lemma fresh_context: 
+  fixes  \<Gamma> :: "(name\<times>ty) list"
+  and    a :: "name"
+  assumes "a\<sharp>\<Gamma>"
+  shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
+using assms by (induct \<Gamma>, auto simp add: fresh_prod fresh_list_cons fresh_atm)
+
+lemma type_unicity_in_context:
+  fixes  \<Gamma> :: "(name\<times>ty)list"
+  and    pi:: "name prm"
+  and    a :: "name"
+  and    \<tau> :: "ty"
+  assumes "valid \<Gamma>" and "(x,T\<^isub>1) \<in> set \<Gamma>" and "(x,T\<^isub>2) \<in> set \<Gamma>"
+  shows "T\<^isub>1=T\<^isub>2"
+using assms by (induct \<Gamma>, auto dest!: fresh_context)
+
+text {* \paragraph{Definition of sub-contexts}
+
+The definition of sub context is standard. We do not use the subset definition to prevent the need for unfolding the definition. We include validity in the definition to shorten the statements.
+ *}
+
+
+abbreviation
+  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [55,55] 55)
+where
+  "\<Gamma>\<^isub>1 \<lless> \<Gamma>\<^isub>2 \<equiv> (\<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>\<^isub>1 \<longrightarrow> (a,\<sigma>)\<in>set \<Gamma>\<^isub>2) \<and> valid \<Gamma>\<^isub>2"
+
+text {* 
+\subsubsection{Definition of the typing relation}
+
+Now, we can define the typing judgements for terms. The rules are given in figure~\ref{typing}. *}
+
+(*<*)
 
 inductive2
   typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60) 
 where
   t_Var[intro]:   "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
-| t_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> e2 : T1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App e1 e2 : T2"
-| t_Lam[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T1)#\<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1\<rightarrow>T2"
+| t_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> e\<^isub>1 : T\<^isub>1\<rightarrow>T\<^isub>2; \<Gamma> \<turnstile> e\<^isub>2 : T\<^isub>1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App e\<^isub>1 e\<^isub>2 : T\<^isub>2"
+| t_Lam[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> t : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T\<^isub>1\<rightarrow>T\<^isub>2"
 | t_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : TBase"
 | t_Unit[intro]:  "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Unit : TUnit"
+(*>*)
+
+text {* 
+\begin{figure}
+\begin{center}
+\isastyle %smaller fonts
+@{thm[mode=Rule] t_Var[no_vars]}{\sc{t\_Var}}\qquad
+@{thm[mode=Rule] t_App[no_vars]}{\sc{t\_App}}\\
+@{thm[mode=Rule] t_Lam[no_vars]}{\sc{t\_Lam}}\\
+@{thm[mode=Rule] t_Const[no_vars]}{\sc{t\_Const}} \qquad
+@{thm[mode=Rule] t_Unit[no_vars]}{\sc{t\_Unit}}
+\end{center}
+\caption{Typing rules\label{typing}}
+\end{figure}
+*}
 
 lemma typing_valid:
   assumes "\<Gamma> \<turnstile> t : T"
   shows "valid \<Gamma>"
-  using assms
-  by (induct) (auto)
- 
-lemma typing_eqvt[eqvt]:
-  fixes pi::"name prm"
-  assumes "\<Gamma> \<turnstile> t : T"
-  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : (pi\<bullet>T)"
-  using assms
-  apply(induct)
-  apply(auto simp add: fresh_bij set_eqvt valid_eqvt) 
-  apply(rule t_Var)
-  apply(drule valid_eqvt)
-  apply(assumption)
-  apply(drule in_eqvt)
-  apply(simp add: set_eqvt)
-  done
+  using assms by (induct)(auto)
+
+nominal_inductive typing
 
-text {* Inversion lemmas for the typing judgment *}
+text {* 
+\subsubsection{Inversion lemmas for the typing relation}
+
+We generate some inversion lemmas for 
+the typing judgment and add them as elimination rules for the automatic tactics.
+During the generation of these lemmas, we need the injectivity properties of the constructor of the nominal datatypes. These are not added by default in the set of simplification rules to prevent unwanted simplifications in the rest of the development. In the future, the \texttt{inductive\_cases} will be reworked to allow to use its own set of rules instead of the whole 'simpset'.
+*}
 
 declare trm.inject [simp add]
 declare ty.inject  [simp add]
@@ -157,16 +266,21 @@
 declare trm.inject [simp del]
 declare ty.inject [simp del]
 
+text {* 
+\subsubsection{Strong induction principle}
+
+Now, we define a strong induction principle. This induction principle allow us to avoid some terms during the induction. The bound variable The avoided terms ($c$)  *}
+
 lemma typing_induct_strong
   [consumes 1, case_names t_Var t_App t_Lam t_Const t_Unit]:
     fixes  P::"'a::fs_name \<Rightarrow> (name\<times>ty) list \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow>bool"
     and    x :: "'a::fs_name"
     assumes a: "\<Gamma> \<turnstile> e : T"
     and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T"
-    and a2: "\<And>\<Gamma> e1 T1 T2 e2 c. \<lbrakk>\<And>c. P c \<Gamma> e1 (T1\<rightarrow>T2); \<And>c. P c \<Gamma> e2 T1\<rbrakk> 
-             \<Longrightarrow> P c \<Gamma> (App e1 e2) T2"
-    and a3: "\<And>x \<Gamma> T1 t T2 c.\<lbrakk>x\<sharp>(\<Gamma>,c); \<And>c. P c ((x,T1)#\<Gamma>) t T2\<rbrakk>
-             \<Longrightarrow> P c \<Gamma> (Lam [x].t) (T1\<rightarrow>T2)"
+    and a2: "\<And>\<Gamma> e\<^isub>1 T\<^isub>1 T\<^isub>2 e\<^isub>2 c. \<lbrakk>\<And>c. P c \<Gamma> e\<^isub>1 (T\<^isub>1\<rightarrow>T\<^isub>2); \<And>c. P c \<Gamma> e\<^isub>2 T\<^isub>1\<rbrakk> 
+             \<Longrightarrow> P c \<Gamma> (App e\<^isub>1 e\<^isub>2) T\<^isub>2"
+    and a3: "\<And>x \<Gamma> T\<^isub>1 t T\<^isub>2 c.\<lbrakk>x\<sharp>(\<Gamma>,c); \<And>c. P c ((x,T\<^isub>1)#\<Gamma>) t T\<^isub>2\<rbrakk>
+             \<Longrightarrow> P c \<Gamma> (Lam [x].t) (T\<^isub>1\<rightarrow>T\<^isub>2)"
     and a4: "\<And>\<Gamma> n c. valid \<Gamma> \<Longrightarrow> P c \<Gamma> (Const n) TBase"
     and a5: "\<And>\<Gamma> c. valid \<Gamma> \<Longrightarrow> P c \<Gamma> Unit TUnit"
     shows "P c \<Gamma> e T"
@@ -179,30 +293,30 @@
     moreover
     have "(x,T)\<in>set \<Gamma>" by fact
     then have "pi\<bullet>(x,T)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
-    then have "(pi\<bullet>x,T)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: set_eqvt)
+    then have "(pi\<bullet>x,T)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: eqvt)
     ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Var x)) (pi\<bullet>T)" using a1 by simp
   next
-    case (t_App \<Gamma> e1 T1 T2 e2 pi c)
-    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(App e1 e2)) (pi\<bullet>T2)" using a2 
+    case (t_App \<Gamma> e\<^isub>1 T\<^isub>1 T\<^isub>2 e\<^isub>2 pi c)
+    thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(App e\<^isub>1 e\<^isub>2)) (pi\<bullet>T\<^isub>2)" using a2 
       by (simp only: eqvt) (blast)
   next
-    case (t_Lam x \<Gamma> T1 t T2 pi c)
+    case (t_Lam x \<Gamma> T\<^isub>1 t T\<^isub>2 pi c)
     obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>\<Gamma>,pi\<bullet>t,c)" by (erule exists_fresh[OF fs_name1])
     let ?sw = "[(pi\<bullet>x,y)]"
     let ?pi' = "?sw@pi"
     have f0: "x\<sharp>\<Gamma>" by fact
     have f1: "(pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>)" using f0 by (simp add: fresh_bij)
     have f2: "y\<sharp>?pi'\<bullet>\<Gamma>" by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
-      have ih1: "\<And>c. P c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>T2)" by fact
-    then have "P c (?pi'\<bullet>\<Gamma>) (Lam [y].(?pi'\<bullet>t)) (T1\<rightarrow>T2)" using fs f2 a3
+      have ih1: "\<And>c. P c (?pi'\<bullet>((x,T\<^isub>1)#\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>T\<^isub>2)" by fact
+    then have "P c (?pi'\<bullet>\<Gamma>) (Lam [y].(?pi'\<bullet>t)) (T\<^isub>1\<rightarrow>T\<^isub>2)" using fs f2 a3
       by (simp add: calc_atm)
-    then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) (T1\<rightarrow>T2)" 
+    then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) (T\<^isub>1\<rightarrow>T\<^isub>2)" 
       by (simp del: append_Cons add: calc_atm pt_name2)
     moreover have "(?sw\<bullet>pi\<bullet>\<Gamma>) = (pi\<bullet>\<Gamma>)"
       by (rule perm_fresh_fresh) (simp_all add: fs f1)
     moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) = Lam [(pi\<bullet>x)].(pi\<bullet>t)"
       by (rule perm_fresh_fresh) (simp_all add: fs f1 fresh_atm abs_fresh)
-    ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [x].t)) (pi\<bullet>T1\<rightarrow>T2)" by simp
+    ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [x].t)) (pi\<bullet>T\<^isub>1\<rightarrow>T\<^isub>2)" by simp
   next
     case (t_Const \<Gamma> n pi c)
     thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Const n)) (pi\<bullet>TBase)" using a4 by (simp, blast intro: valid_eqvt)
@@ -214,7 +328,13 @@
   then show "P c \<Gamma> e T" by simp
 qed
 
-text {* capture-avoiding substitution *}
+text {* 
+\subsection{Capture-avoiding substitutions}
+
+In this section we define parallel substitution. The usual substitution will be derived as a special case of parallel substitution. But first we define a function to lookup for the term corresponding to a type in an association list. Note that if the term does not appear in the list then we return a variable of that name.\\
+Should we return Some or None and put that in a library ?
+ *}
+
 
 fun
   lookup :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm"   
@@ -222,18 +342,17 @@
   "lookup [] X        = Var X"
   "lookup ((Y,T)#\<theta>) X = (if X=Y then T else lookup \<theta> X)"
 
-lemma lookup_eqvt:
+lemma lookup_eqvt[eqvt]:
   fixes pi::"name prm"
   and   \<theta>::"(name\<times>trm) list"
   and   X::"name"
   shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)"
-by (induct \<theta>)
-   (auto simp add: perm_bij)
+by (induct \<theta>) (auto simp add: perm_bij)
 
 lemma lookup_fresh:
   fixes z::"name"
   assumes "z\<sharp>\<theta>" "z\<sharp>x"
-  shows "z \<sharp>lookup \<theta> x"
+  shows "z\<sharp> lookup \<theta> x"
 using assms 
 by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)
 
@@ -244,23 +363,30 @@
 by (induct rule: lookup.induct)
    (auto simp add: fresh_list_cons fresh_prod fresh_atm)
 
+text {* \subsubsection{Parallel substitution} *}
+
 consts
   psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm"  ("_<_>" [60,100] 100)
  
 nominal_primrec
   "\<theta><(Var x)> = (lookup \<theta> x)"
-  "\<theta><(App t1 t2)> = App (\<theta><t1>) (\<theta><t2>)"
+  "\<theta><(App t\<^isub>1 t\<^isub>2)> = App (\<theta><t\<^isub>1>) (\<theta><t\<^isub>2>)"
   "x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].t)> = Lam [x].(\<theta><t>)"
   "\<theta><(Const n)> = Const n"
   "\<theta><(Unit)> = Unit"
-apply(finite_guess add: fs_name1 lookup_eqvt)+
-apply(perm_full_simp)
-apply(simp add: fs_name1)
+apply(finite_guess)+
 apply(rule TrueI)+
 apply(simp add: abs_fresh)+
-apply(fresh_guess add: fs_name1 lookup_eqvt)+
+apply(fresh_guess)+
 done
 
+
+text {* \subsubsection{Substitution} 
+
+The substitution function is defined just as a special case of parallel substitution. 
+
+*}
+
 abbreviation 
  subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
 where
@@ -268,7 +394,7 @@
 
 lemma subst[simp]:
   shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
-  and   "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
+  and   "(App t\<^isub>1 t\<^isub>2)[y::=t'] = App (t\<^isub>1[y::=t']) (t\<^isub>2[y::=t'])"
   and   "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
   and   "Const n[y::=t'] = Const n"
   and   "Unit [y::=t'] = Unit"
@@ -281,13 +407,15 @@
   by (nominal_induct t avoiding: x t' rule: trm.induct)
      (perm_simp add: fresh_bij)+
 
+text {* \subsubsection{Lemmas about freshness and substitutions} *}
+
 lemma subst_rename: 
   fixes c::"name"
   and   t1::"trm"
-  assumes a: "c\<sharp>t1"
-  shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]"
+  assumes a: "c\<sharp>t\<^isub>1"
+  shows "t\<^isub>1[a::=t\<^isub>2] = ([(c,a)]\<bullet>t\<^isub>1)[c::=t\<^isub>2]"
   using a
-  apply(nominal_induct t1 avoiding: a c t2 rule: trm.induct)
+  apply(nominal_induct t\<^isub>1 avoiding: a c t\<^isub>2 rule: trm.induct)
   apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def)+
   done
 
@@ -304,28 +432,28 @@
   fixes z::"name"
   and   t1::"trm"
   and   t2::"trm"
-  assumes "z\<sharp>t2"
-  shows "z\<sharp>t1[z::=t2]"
+  assumes "z\<sharp>t\<^isub>2"
+  shows "z\<sharp>t\<^isub>1[z::=t\<^isub>2]"
 using assms 
-by (nominal_induct t1 avoiding: t2 z rule: trm.induct)
+by (nominal_induct t\<^isub>1 avoiding: t\<^isub>2 z rule: trm.induct)
    (auto simp add: abs_fresh fresh_nat fresh_atm)
 
 lemma fresh_subst':
   fixes z::"name"
   and   t1::"trm"
   and   t2::"trm"
-  assumes "z\<sharp>[y].t1" "z\<sharp>t2"
-  shows "z\<sharp>t1[y::=t2]"
+  assumes "z\<sharp>[y].t\<^isub>1" "z\<sharp>t\<^isub>2"
+  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
 using assms 
-by (nominal_induct t1 avoiding: y t2 z rule: trm.induct)
+by (nominal_induct t\<^isub>1 avoiding: y t\<^isub>2 z rule: trm.induct)
    (auto simp add: abs_fresh fresh_nat fresh_atm)
 
 lemma fresh_subst:
   fixes z::"name"
   and   t1::"trm"
   and   t2::"trm"
-  assumes "z\<sharp>t1" "z\<sharp>t2"
-  shows "z\<sharp>t1[y::=t2]"
+  assumes "z\<sharp>t\<^isub>1" "z\<sharp>t\<^isub>2"
+  shows "z\<sharp>t\<^isub>1[y::=t\<^isub>2]"
 using assms fresh_subst'
 by (auto simp add: abs_fresh) 
 
@@ -346,415 +474,6 @@
   ultimately show "(x,u)#\<theta><Lam [y].t> = \<theta><Lam [y].t>" by auto
 qed (auto simp add: fresh_atm abs_fresh)
 
-text {* Equivalence (defined) *}
-
-inductive2
-  equiv_def :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ == _ : _" [60,60] 60) 
-where
-  Q_Refl[intro]:  "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> t == t : T"
-| Q_Symm[intro]:  "\<Gamma> \<turnstile> t == s : T \<Longrightarrow> \<Gamma> \<turnstile> s == t : T"
-| Q_Trans[intro]: "\<lbrakk>\<Gamma> \<turnstile> s == t : T; \<Gamma> \<turnstile> t == u : T\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> s == u : T"
-| Q_Abs[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T1)#\<Gamma> \<turnstile> s2 == t2 : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. s2 ==  Lam [x]. t2 : T1 \<rightarrow> T2"
-| Q_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> s1 == t1 : T1 \<rightarrow> T2 ; \<Gamma> \<turnstile> s2 == t2 : T1\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> App s1 s2 == App t1 t2 : T2"
-| Q_Beta[intro]:  "\<lbrakk>x\<sharp>(\<Gamma>,s2,t2); (x,T1)#\<Gamma> \<turnstile> s12 == t12 : T2 ; \<Gamma> \<turnstile> s2 == t2 : T1\<rbrakk> 
-                   \<Longrightarrow>  \<Gamma> \<turnstile> App (Lam [x]. s12) s2 ==  t12[x::=t2] : T2"
-| Q_Ext[intro]:   "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T1)#\<Gamma> \<turnstile> App s (Var x) == App t (Var x) : T2\<rbrakk> 
-                   \<Longrightarrow> \<Gamma> \<turnstile> s == t : T1 \<rightarrow> T2"
-
-lemma equiv_def_valid:
-  assumes "\<Gamma> \<turnstile> t == s : T"
-  shows "valid \<Gamma>"
-using assms by (induct,auto elim:typing_valid)
-
-lemma equiv_def_eqvt[eqvt]:
-  fixes pi::"name prm"
-  assumes a: "\<Gamma> \<turnstile> s == t : T"
-  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) == (pi\<bullet>t) : (pi\<bullet>T)"
-using a
-apply(induct)
-apply(auto intro: typing_eqvt valid_eqvt simp add: fresh_bij subst_eqvt simp del: perm_ty)
-apply(rule_tac x="pi\<bullet>x" in Q_Ext)
-apply(simp add: fresh_bij)+
-done
-
-lemma equiv_def_strong_induct
-  [consumes 1, case_names Q_Refl Q_Symm Q_Trans Q_Abs Q_App Q_Beta Q_Ext]:
-  fixes c::"'a::fs_name"
-  assumes a0: "\<Gamma> \<turnstile> s == t : T" 
-  and     a1: "\<And>\<Gamma> t T c.  \<Gamma> \<turnstile> t : T  \<Longrightarrow> P c \<Gamma> t t T"
-  and     a2: "\<And>\<Gamma> t s T c. \<lbrakk>\<And>d. P d \<Gamma> t s T\<rbrakk> \<Longrightarrow>  P c \<Gamma> s t T"
-  and     a3: "\<And>\<Gamma> s t T u c. \<lbrakk>\<And>d. P d \<Gamma> s t T; \<And>d. P d \<Gamma> t u T\<rbrakk> 
-               \<Longrightarrow> P c \<Gamma> s u T"
-  and     a4: "\<And>x \<Gamma> T1 s2 t2 T2 c. \<lbrakk>x\<sharp>\<Gamma>; x\<sharp>c; \<And>d. P d ((x,T1)#\<Gamma>) s2 t2 T2\<rbrakk>
-               \<Longrightarrow> P c \<Gamma> (Lam [x].s2) (Lam [x].t2) (T1\<rightarrow>T2)"
-  and     a5: "\<And>\<Gamma> s1 t1 T1 T2 s2 t2 c. \<lbrakk>\<And>d. P d \<Gamma> s1 t1 (T1\<rightarrow>T2); \<And>d. P d \<Gamma> s2 t2 T1\<rbrakk> 
-               \<Longrightarrow> P c \<Gamma> (App s1 s2) (App t1 t2) T2"
-  and     a6: "\<And>x \<Gamma> T1 s12 t12 T2 s2 t2 c.
-               \<lbrakk>x\<sharp>(\<Gamma>,s2,t2); x\<sharp>c; \<And>d. P d ((x,T1)#\<Gamma>) s12 t12 T2; \<And>d. P d \<Gamma> s2 t2 T1\<rbrakk> 
-               \<Longrightarrow> P c \<Gamma> (App (Lam [x].s12) s2) (t12[x::=t2]) T2"
-  and     a7: "\<And>x \<Gamma> s t T1 T2 c.
-               \<lbrakk>x\<sharp>(\<Gamma>,s,t); \<And>d. P d ((x,T1)#\<Gamma>) (App s (Var x)) (App t (Var x)) T2\<rbrakk>
-               \<Longrightarrow> P c \<Gamma> s t (T1\<rightarrow>T2)"
- shows "P c \<Gamma>  s t T"
-proof -
-  from a0 have "\<And>(pi::name prm) c. P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (pi\<bullet>T)" 
-    proof(induct)
-      case (Q_Refl \<Gamma> t T pi c)
-      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>t) (pi\<bullet>t) (pi\<bullet>T)" using a1 by (simp only: eqvt)
-    next
-      case (Q_Symm \<Gamma> t s T pi c)
-      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (pi\<bullet>T)" using a2 by (simp only: equiv_def_eqvt)
-    next
-      case (Q_Trans \<Gamma> s t T u pi c)
-      then show " P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>u) (pi\<bullet>T)" using a3 by (blast intro: equiv_def_eqvt)
-    next
-      case (Q_App \<Gamma> s1 t1 T1 T2 s2 t2 pi c)
-      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>App s1 s2) (pi\<bullet>App t1 t2) (pi\<bullet>T2)" using a5 
-	by (simp only: eqvt) (blast)
-    next
-      case (Q_Ext x \<Gamma> s t T1 T2 pi c)
-      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (pi\<bullet>T1\<rightarrow>T2)"  
-	by (auto intro!: a7 simp add: fresh_bij)
-    next
-      case (Q_Abs x \<Gamma> T1 s2 t2 T2 pi c)
-      obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>s2,pi\<bullet>t2,pi\<bullet>\<Gamma>,c)" by (rule exists_fresh[OF fs_name1])
-      let ?sw="[(pi\<bullet>x,y)]"
-      let ?pi'="?sw@pi"
-      have f1: "x\<sharp>\<Gamma>" by fact
-      have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>)" using f1 by (simp add: fresh_bij)
-      have f3: "y\<sharp>?pi'\<bullet>\<Gamma>" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
-      have ih1: "\<And>c. P c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>s2) (?pi'\<bullet>t2) (?pi'\<bullet>T2)" by fact
-      then have "\<And>c. P c ((y,T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>s2) (?pi'\<bullet>t2) T2" by (simp add: calc_atm)
-      then have "P c  (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>Lam [x].s2) (?pi'\<bullet>Lam [x].t2) (T1\<rightarrow>T2)" using a4 f3 fs
-	by (simp add: calc_atm)
-      then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>Lam [(pi\<bullet>x)].(pi\<bullet>s2)) (?sw\<bullet>Lam [(pi\<bullet>x)].(pi\<bullet>t2)) (T1\<rightarrow>T2)" 
-	by (simp del: append_Cons add: calc_atm pt_name2)
-      moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs f2)
-      moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>s2))) = Lam [(pi\<bullet>x)].(pi\<bullet>s2)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
-      moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t2))) = Lam [(pi\<bullet>x)].(pi\<bullet>t2)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
-      ultimately have "P c (pi\<bullet>\<Gamma>) (Lam [(pi\<bullet>x)].(pi\<bullet>s2)) (Lam [(pi\<bullet>x)].(pi\<bullet>t2)) (T1\<rightarrow>T2)" by simp
-      then show  "P c (pi\<bullet>\<Gamma>) (pi\<bullet>Lam [x].s2) (pi\<bullet>Lam [x].t2) (pi\<bullet>T1\<rightarrow>T2)" by simp 
-    next 
-      case (Q_Beta x \<Gamma> s2 t2 T1 s12 t12 T2 pi c) 
-      obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>s12,pi\<bullet>t12,pi\<bullet>s2,pi\<bullet>t2,pi\<bullet>\<Gamma>,c)" 
-	by (rule exists_fresh[OF fs_name1])
-      let ?sw="[(pi\<bullet>x,y)]"
-      let ?pi'="?sw@pi"
-      have f1: "x\<sharp>(\<Gamma>,s2,t2)" by fact 
-      have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>(\<Gamma>,s2,t2))" using f1 by (simp add: fresh_bij)
-      have f3: "y\<sharp>(?pi'\<bullet>(\<Gamma>,s2,t2))" using f1 
-	by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp fresh_prod)
-      have ih1: "\<And>c. P c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>s12) (?pi'\<bullet>t12) (?pi'\<bullet>T2)" by fact
-      then have "\<And>c. P c ((y,T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>s12) (?pi'\<bullet>t12) (?pi'\<bullet>T2)" by (simp add: calc_atm)
-      moreover
-      have ih2: "\<And>c. P c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>s2) (?pi'\<bullet>t2) (?pi'\<bullet>T1)" by fact
-      ultimately have "P c  (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>(App (Lam [x].s12) s2)) (?pi'\<bullet>(t12[x::=t2])) (?pi'\<bullet>T2)" 
-	using a6 f3 fs by (force simp add: subst_eqvt calc_atm del: perm_ty)
-      then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2))) 
-	(?sw\<bullet>((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)])) T2" 
-	by (simp del: append_Cons add: calc_atm pt_name2 subst_eqvt)
-      moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
-      moreover have "(?sw\<bullet>(App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2))) = App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified] abs_fresh)
-      moreover have "(?sw\<bullet>((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)])) = ((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)])" 
-	by (rule perm_fresh_fresh) 
-	   (simp_all add: fs[simplified] f2[simplified] fresh_subst fresh_subst'')
-      ultimately have "P c (pi\<bullet>\<Gamma>) (App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2)) ((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)]) T2"
-	by simp
-      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>App (Lam [x].s12) s2) (pi\<bullet>t12[x::=t2]) (pi\<bullet>T2)" by (simp add: subst_eqvt)
-    qed
-  then have "P c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>s) (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>T)" by blast
-  then show "P c \<Gamma> s t T" by simp
-qed
-
-text {* Weak head reduction *}
-
-inductive2
-  whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80) 
-where
-  QAR_Beta[intro]: "App (Lam [x]. t12) t2 \<leadsto> t12[x::=t2]"
-| QAR_App[intro]: "t1 \<leadsto> t1' \<Longrightarrow> App t1 t2 \<leadsto> App t1' t2"
-
-declare trm.inject  [simp add]
-declare ty.inject  [simp add]
-
-inductive_cases2 whr_Gen[elim]: "t \<leadsto> t'"
-inductive_cases2 whr_Lam[elim]: "Lam [x].t \<leadsto> t'"
-inductive_cases2 whr_App_Lam[elim]: "App (Lam [x].t12) t2 \<leadsto> t"
-inductive_cases2 whr_Var[elim]: "Var x \<leadsto> t"
-inductive_cases2 whr_Const[elim]: "Const n \<leadsto> t"
-inductive_cases2 whr_App[elim]: "App p q \<leadsto> t"
-inductive_cases2 whr_Const_Right[elim]: "t \<leadsto> Const n"
-inductive_cases2 whr_Var_Right[elim]: "t \<leadsto> Var x"
-inductive_cases2 whr_App_Right[elim]: "t \<leadsto> App p q"
-
-declare trm.inject  [simp del]
-declare ty.inject  [simp del]
-
-text {* Weak head normalization *}
-
-abbreviation 
- nf :: "trm \<Rightarrow> bool" ("_ \<leadsto>|" [100] 100)
-where
-  "t\<leadsto>|  \<equiv> \<not>(\<exists> u. t \<leadsto> u)" 
-
-inductive2
-  whn_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
-where
-  QAN_Reduce[intro]: "\<lbrakk>s \<leadsto> t; t \<Down> u\<rbrakk> \<Longrightarrow> s \<Down> u"
-| QAN_Normal[intro]: "t\<leadsto>|  \<Longrightarrow> t \<Down> t"
-
-lemma whr_eqvt:
-  fixes pi::"name prm"
-  assumes a: "t \<leadsto> t'"
-  shows "(pi\<bullet>t) \<leadsto> (pi\<bullet>t')"
-using a
-by (induct) (auto simp add: subst_eqvt fresh_bij)
-
-lemma whn_eqvt[eqvt]:
-  fixes pi::"name prm"
-  assumes a: "t \<Down> t'"
-  shows "(pi\<bullet>t) \<Down> (pi\<bullet>t')"
-using a
-apply(induct)
-apply(rule QAN_Reduce)
-apply(rule whr_eqvt)
-apply(assumption)+
-apply(rule QAN_Normal)
-apply(auto)
-apply(drule_tac pi="rev pi" in whr_eqvt)
-apply(perm_simp)
-done
-
-text {* Algorithmic term equivalence and algorithmic path equivalence *}
-
-inductive2
-  alg_equiv :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ <=> _ : _" [60,60] 60) 
-and
-  alg_path_equiv :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60] 60) 
-where
-  QAT_Base[intro]:  "\<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : TBase"
-| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>\<Gamma>; x\<sharp>s; x\<sharp>t; (x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2\<rbrakk> 
-                     \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T1 \<rightarrow> T2"
-| QAT_One[intro]:   "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : TUnit"
-| QAP_Var[intro]:   "\<lbrakk>valid \<Gamma>;(x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T"
-| QAP_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T1 \<rightarrow> T2; \<Gamma> \<turnstile> s <=> t : T1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App p s \<leftrightarrow> App q t : T2"
-| QAP_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n \<leftrightarrow> Const n : TBase"
-
-lemma alg_equiv_alg_path_equiv_eqvt[eqvt]:
-  fixes pi::"name prm"
-  shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) <=> (pi\<bullet>t) : (pi\<bullet>T)" 
-  and   "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>p) \<leftrightarrow> (pi\<bullet>q) : (pi\<bullet>T)"
-apply(induct rule: alg_equiv_alg_path_equiv.inducts)
-apply(auto intro: whn_eqvt simp add: fresh_bij valid_eqvt)
-apply(rule_tac x="pi\<bullet>x" in  QAT_Arrow)
-apply(auto simp add: fresh_bij)
-apply(rule QAP_Var)
-apply(simp add: valid_eqvt)
-apply(simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
-apply(perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
-done
-
-lemma alg_equiv_alg_path_equiv_strong_induct
-  [case_names QAT_Base QAT_Arrow QAT_One QAP_Var QAP_App QAP_Const]:
-  fixes c::"'a::fs_name"
-  assumes a1: "\<And>s p t q \<Gamma> c. \<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase; \<And>d. P2 d \<Gamma> p q TBase\<rbrakk> 
-               \<Longrightarrow> P1 c \<Gamma> s t TBase"
-  and     a2: "\<And>x \<Gamma> s t T1 T2 c.
-               \<lbrakk>x\<sharp>\<Gamma>; x\<sharp>s; x\<sharp>t; x\<sharp>c; (x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2;
-               \<And>d. P1 d ((x,T1)#\<Gamma>) (App s (Var x)) (App t (Var x)) T2\<rbrakk>
-               \<Longrightarrow> P1 c \<Gamma> s t (T1\<rightarrow>T2)"
-  and     a3: "\<And>\<Gamma> s t c. valid \<Gamma> \<Longrightarrow> P1 c \<Gamma> s t TUnit"
-  and     a4: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P2 c \<Gamma> (Var x) (Var x) T"
-  and     a5: "\<And>\<Gamma> p q T1 T2 s t c.
-               \<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T1\<rightarrow>T2; \<And>d. P2 d \<Gamma> p q (T1\<rightarrow>T2); \<Gamma> \<turnstile> s <=> t : T1; \<And>d. P1 d \<Gamma> s t T1\<rbrakk>
-               \<Longrightarrow> P2 c \<Gamma> (App p s) (App q t) T2"
-  and     a6: "\<And>\<Gamma> n c. valid \<Gamma> \<Longrightarrow> P2 c \<Gamma> (Const n) (Const n) TBase"
-  shows "(\<Gamma> \<turnstile> s <=> t : T \<longrightarrow>P1 c \<Gamma> s t T) \<and> (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> P2 c \<Gamma>' s' t' T')"
-proof -
-  have "(\<Gamma> \<turnstile> s <=> t : T \<longrightarrow> (\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) T)) \<and> 
-        (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> (\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>') (pi\<bullet>s') (pi\<bullet>t') T'))"
-  proof(induct rule: alg_equiv_alg_path_equiv.induct)
-    case (QAT_Base s q t p \<Gamma>)
-    then show "\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) TBase"
-      apply(auto)
-      apply(rule_tac p="pi\<bullet>q" and q="pi\<bullet>p" in  a1)
-      apply(simp_all add: whn_eqvt alg_equiv_alg_path_equiv_eqvt[simplified])
-      done
-  next
-    case (QAT_Arrow x \<Gamma> s t T1 T2)
-    show ?case
-    proof (rule allI, rule allI)
-      fix c::"'a::fs_name" and pi::"name prm"
-      obtain y::"name" where fs: "y\<sharp>(pi\<bullet>s,pi\<bullet>t,pi\<bullet>\<Gamma>,c)" by (rule exists_fresh[OF fs_name1])
-      let ?sw="[(pi\<bullet>x,y)]"
-      let ?pi'="?sw@pi"
-      have f0: "x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
-      then have f1: "x\<sharp>(\<Gamma>,s,t)" by simp
-      have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>(\<Gamma>,s,t))" using f1 by (simp add: fresh_bij)
-      have f3: "y\<sharp>?pi'\<bullet>(\<Gamma>,s,t)" using f1 
-	by (simp only: pt_name2 fresh_left, auto simp add: perm_pi_simp calc_atm)
-      then have f3': "y\<sharp>?pi'\<bullet>\<Gamma>" "y\<sharp>?pi'\<bullet>s" "y\<sharp>?pi'\<bullet>t" by (auto simp add: fresh_prod)
-      have pr1: "(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2" by fact
-      then have "(?pi'\<bullet>((x,T1)#\<Gamma>)) \<turnstile> (?pi'\<bullet>(App s (Var x))) <=> (?pi'\<bullet>(App t (Var x))) : (?pi'\<bullet>T2)" 
-	by (rule alg_equiv_alg_path_equiv_eqvt)
-      then have "(y,T1)#(?pi'\<bullet>\<Gamma>) \<turnstile> (App (?pi'\<bullet>s) (Var y)) <=> (App (?pi'\<bullet>t) (Var y)) : T2" 
-	by (simp add: calc_atm)
-      moreover    
-      have ih1: "\<forall>c (pi::name prm).  P1 c (pi\<bullet>((x,T1)#\<Gamma>)) (pi\<bullet>App s (Var x)) (pi\<bullet>App t (Var x)) T2"
-	by fact
-      then have "\<And>c.  P1 c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>App s (Var x)) (?pi'\<bullet>App t (Var x)) T2"
-	by auto
-      then have "\<And>c.  P1 c ((y,T1)#(?pi'\<bullet>\<Gamma>)) (App (?pi'\<bullet>s) (Var y)) (App (?pi'\<bullet>t) (Var y)) T2"
-	by (simp add: calc_atm)     
-      ultimately have "P1 c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>s) (?pi'\<bullet>t) (T1\<rightarrow>T2)" using a2 f3' fs by simp
-      then have "P1 c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>pi\<bullet>s) (?sw\<bullet>pi\<bullet>t) (T1\<rightarrow>T2)" 
-	by (simp del: append_Cons add: calc_atm pt_name2)
-      moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
-      moreover have "(?sw\<bullet>(pi\<bullet>s)) = (pi\<bullet>s)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
-      moreover have "(?sw\<bullet>(pi\<bullet>t)) = (pi\<bullet>t)" 
-	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
-      ultimately show "P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (T1\<rightarrow>T2)" by (simp)
-    qed
-  next
-    case (QAT_One \<Gamma> s t)
-    then show "\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) TUnit"
-      by (auto intro!: a3 simp add: valid_eqvt)
-  next
-    case (QAP_Var \<Gamma> x T)
-    then show "\<forall>c (pi::name prm). P2 c (pi \<bullet> \<Gamma>) (pi \<bullet> Var x) (pi \<bullet> Var x) T"
-      apply(auto intro!: a4 simp add: valid_eqvt)
-      apply(simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
-      apply(perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
-      done
-  next
-    case (QAP_App \<Gamma> p q T1 T2 s t)
-    then show "\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>) (pi\<bullet>App p s) (pi\<bullet>App q t) T2"
-      by (auto intro!: a5 simp add: alg_equiv_alg_path_equiv_eqvt[simplified])
-  next
-    case (QAP_Const \<Gamma> n)
-    then show "\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>) (pi\<bullet>Const n) (pi\<bullet>Const n) TBase"
-      by (auto intro!: a6 simp add: valid_eqvt)
-  qed
-  then have "(\<Gamma> \<turnstile> s <=> t : T \<longrightarrow> P1 c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>s) (([]::name prm)\<bullet>t) T) \<and> 
-             (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> P2 c (([]::name prm)\<bullet>\<Gamma>') (([]::name prm)\<bullet>s') (([]::name prm)\<bullet>t') T')"
-    by blast
-  then show ?thesis by simp
-qed
-
-(* post-processing of the strong induction principle proved above; *) 
-(* the code extracts the strong_inducts-version from strong_induct *)
-setup {*
-  PureThy.add_thmss
-    [(("alg_equiv_alg_path_equiv_strong_inducts",
-      ProjectRule.projects (ProofContext.init (the_context())) [1,2]
-        (thm "alg_equiv_alg_path_equiv_strong_induct")), [])] #> snd
-*}
-
-declare trm.inject  [simp add]
-declare ty.inject  [simp add]
-
-inductive_cases2 whn_inv_auto[elim]: "t \<Down> t'"
-
-inductive_cases2 alg_equiv_Base_inv_auto[elim]: "\<Gamma> \<turnstile> s<=>t : TBase"
-inductive_cases2 alg_equiv_Arrow_inv_auto[elim]: "\<Gamma> \<turnstile> s<=>t : T1 \<rightarrow> T2"
-
-inductive_cases2 alg_path_equiv_Base_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : TBase"
-inductive_cases2 alg_path_equiv_Unit_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : TUnit"
-inductive_cases2 alg_path_equiv_Arrow_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : T1 \<rightarrow> T2"
-
-inductive_cases2 alg_path_equiv_Var_left_inv_auto[elim]: "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T"
-inductive_cases2 alg_path_equiv_Var_left_inv_auto'[elim]: "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T'"
-inductive_cases2 alg_path_equiv_Var_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T"
-inductive_cases2 alg_path_equiv_Var_right_inv_auto'[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T'"
-inductive_cases2 alg_path_equiv_Const_left_inv_auto[elim]: "\<Gamma> \<turnstile> Const n \<leftrightarrow> t : T"
-inductive_cases2 alg_path_equiv_Const_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Const n : T"
-inductive_cases2 alg_path_equiv_App_left_inv_auto[elim]: "\<Gamma> \<turnstile> App p s \<leftrightarrow> t : T"
-inductive_cases2 alg_path_equiv_App_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> App q t : T"
-inductive_cases2 alg_path_equiv_Lam_left_inv_auto[elim]: "\<Gamma> \<turnstile> Lam[x].s \<leftrightarrow> t : T"
-inductive_cases2 alg_path_equiv_Lam_right_inv_auto[elim]: "\<Gamma> \<turnstile> t \<leftrightarrow> Lam[x].s : T"
-
-declare trm.inject [simp del]
-declare ty.inject [simp del]
-
-text {* Subcontext. *}
-
-abbreviation
-  "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [55,55] 55)
-where
-  "\<Gamma>1 \<lless> \<Gamma>2 \<equiv> (\<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow>  (a,\<sigma>)\<in>set \<Gamma>2) \<and> valid \<Gamma>2"
-
-lemma alg_equiv_valid:
-  shows  "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> valid \<Gamma>" and  "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> valid \<Gamma>"
-by (induct rule : alg_equiv_alg_path_equiv.inducts, auto)
-
-lemma algorithmic_monotonicity:
-  fixes \<Gamma>':: "(name\<times>ty) list"
-  shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s <=> t : T"
-  and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<leftrightarrow> t : T"
-proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: \<Gamma>' rule: alg_equiv_alg_path_equiv_strong_inducts)
-  case (QAT_Arrow x \<Gamma> s t T1 T2 \<Gamma>')
-  have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
-  have h2:"\<Gamma>\<lless>\<Gamma>'" by fact
-  have ih:"\<And>\<Gamma>'. \<lbrakk>(x,T1)#\<Gamma> \<lless> \<Gamma>'\<rbrakk>  \<Longrightarrow> \<Gamma>' \<turnstile> App s (Var x) <=> App t (Var x) : T2" by fact
-  have "x\<sharp>\<Gamma>'" by fact
-  then have sub:"(x,T1)#\<Gamma> \<lless> (x,T1)#\<Gamma>'" using h2 by auto
-  then have "(x,T1)#\<Gamma>' \<turnstile> App s (Var x) <=> App t (Var x) : T2" using ih by auto
-  then show "\<Gamma>' \<turnstile> s <=> t : T1\<rightarrow>T2" using sub fs by auto
-qed (auto)
-
-text {* Logical equivalence. *}
-
-function log_equiv :: "((name\<times>ty) list \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)"   
-                      ("_ \<turnstile> _ is _ : _" [60,60,60,60] 60) 
-where    
-   "\<Gamma> \<turnstile> s is t : TUnit = valid \<Gamma>"
- | "\<Gamma> \<turnstile> s is t : TBase = \<Gamma> \<turnstile> s <=> t : TBase"
- | "\<Gamma> \<turnstile> s is t : (T1 \<rightarrow> T2) =  
-           (valid \<Gamma> \<and> (\<forall>\<Gamma>' s' t'. \<Gamma>\<lless>\<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T1 \<longrightarrow>  (\<Gamma>' \<turnstile> (App s s') is (App t t') : T2)))"
-apply (auto simp add: ty.inject)
-apply (subgoal_tac "(\<exists>T1 T2. b=T1 \<rightarrow> T2) \<or> b=TUnit \<or> b=TBase" )
-apply (force)
-apply (rule ty_cases)
-done
-
-termination
-apply(relation "measure (\<lambda>(_,_,_,T). size T)")
-apply(auto)
-done
-
-lemma log_equiv_valid: 
-  assumes "\<Gamma> \<turnstile> s is t : T"
-  shows "valid \<Gamma>"
-using assms 
-by (induct rule: log_equiv.induct,auto elim: alg_equiv_valid)
-
-lemma logical_monotonicity :
- fixes T::ty
- assumes "\<Gamma> \<turnstile> s is t : T" "\<Gamma>\<lless>\<Gamma>'" 
- shows "\<Gamma>' \<turnstile> s is t : T"
-using assms
-proof (induct arbitrary: \<Gamma>' rule: log_equiv.induct)
-  case (2 \<Gamma> s t \<Gamma>')
-  then show "\<Gamma>' \<turnstile> s is t : TBase" using algorithmic_monotonicity by auto
-next
-  case (3 \<Gamma> s t T1 T2 \<Gamma>')
-  have h1:"\<Gamma> \<turnstile> s is t : T1\<rightarrow>T2" by fact
-  have h2:"\<Gamma>\<lless>\<Gamma>'" by fact
-  {
-    fix \<Gamma>'' s' t'
-    assume "\<Gamma>'\<lless>\<Gamma>''" "\<Gamma>'' \<turnstile> s' is t' : T1"
-    then have "\<Gamma>'' \<turnstile> (App s s') is (App t t') : T2" using h1 h2 by auto
-  }
-  then show "\<Gamma>' \<turnstile> s is t : T1\<rightarrow>T2" using h2 by auto
-qed (auto)
-  
 lemma forget: 
   fixes x::"name"
   and   L::"trm"
@@ -766,17 +485,17 @@
 
 lemma subst_fun_eq:
   fixes u::trm
-  assumes h:"[x].t1 = [y].t2"
-  shows "t1[x::=u] = t2[y::=u]"
+  assumes h:"[x].t\<^isub>1 = [y].t\<^isub>2"
+  shows "t\<^isub>1[x::=u] = t\<^isub>2[y::=u]"
 proof -
   { 
-    assume "x=y" and "t1=t2"
+    assume "x=y" and "t\<^isub>1=t\<^isub>2"
     then have ?thesis using h by simp
   }
   moreover 
   {
-    assume h1:"x \<noteq> y" and h2:"t1=[(x,y)] \<bullet> t2" and h3:"x \<sharp> t2"
-    then have "([(x,y)] \<bullet> t2)[x::=u] = t2[y::=u]" by (simp add: subst_rename)
+    assume h1:"x \<noteq> y" and h2:"t\<^isub>1=[(x,y)] \<bullet> t\<^isub>2" and h3:"x \<sharp> t\<^isub>2"
+    then have "([(x,y)] \<bullet> t\<^isub>2)[x::=u] = t\<^isub>2[y::=u]" by (simp add: subst_rename)
     then have ?thesis using h2 by simp 
   }
   ultimately show ?thesis using alpha h by blast
@@ -787,7 +506,7 @@
   by (nominal_induct t rule: trm.induct) (auto simp add: fresh_list_nil)
 
 lemma psubst_subst_psubst:
-  assumes h:"c \<sharp> \<theta>"
+  assumes h:"c\<sharp>\<theta>"
   shows "\<theta><t>[c::=s] = (c,s)#\<theta><t>"
   using h
 by (nominal_induct t avoiding: \<theta> c s rule: trm.induct)
@@ -830,13 +549,473 @@
   ultimately have "\<theta><(Lam [n].t)[x::=u]> = (Lam [n].\<theta><t>)[x::=\<theta><u>]" using fs by auto
   then show "\<theta><(Lam [n].t)[x::=u]> = \<theta><Lam [n].t>[x::=\<theta><u>]" using fs by auto
 qed (auto)
+
+text {* 
+\section{Equivalence}
+\subsection{Definition of the equivalence relation}
+ *}
+
+(*<*)
+inductive2
+  equiv_def :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<equiv> _ : _" [60,60] 60) 
+where
+  Q_Refl[intro]:  "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<equiv> t : T"
+| Q_Symm[intro]:  "\<Gamma> \<turnstile> t \<equiv> s : T \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"
+| Q_Trans[intro]: "\<lbrakk>\<Gamma> \<turnstile> s \<equiv> t : T; \<Gamma> \<turnstile> t \<equiv> u : T\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> s \<equiv> u : T"
+| Q_Abs[intro]:   "\<lbrakk>x\<sharp>\<Gamma>; (x,T\<^isub>1)#\<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. s\<^isub>2 \<equiv>  Lam [x]. t\<^isub>2 : T\<^isub>1 \<rightarrow> T\<^isub>2"
+| Q_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> s\<^isub>1 \<equiv> t\<^isub>1 : T\<^isub>1 \<rightarrow> T\<^isub>2 ; \<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1\<rbrakk> \<Longrightarrow>  \<Gamma> \<turnstile> App s\<^isub>1 s\<^isub>2 \<equiv> App t\<^isub>1 t\<^isub>2 : T\<^isub>2"
+| Q_Beta[intro]:  "\<lbrakk>x\<sharp>(\<Gamma>,s\<^isub>2,t\<^isub>2); (x,T\<^isub>1)#\<Gamma> \<turnstile> s12 \<equiv> t12 : T\<^isub>2 ; \<Gamma> \<turnstile> s\<^isub>2 \<equiv> t\<^isub>2 : T\<^isub>1\<rbrakk> 
+                   \<Longrightarrow>  \<Gamma> \<turnstile> App (Lam [x]. s12) s\<^isub>2 \<equiv>  t12[x::=t\<^isub>2] : T\<^isub>2"
+| Q_Ext[intro]:   "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<equiv> App t (Var x) : T\<^isub>2\<rbrakk> 
+                   \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+(*>*)
+
+text {* 
+\begin{center}
+\isastyle
+@{thm[mode=Rule] Q_Refl[no_vars]}{\sc{Q\_Refl}}\qquad
+@{thm[mode=Rule] Q_Symm[no_vars]}{\sc{Q\_Symm}}\\
+@{thm[mode=Rule] Q_Trans[no_vars]}{\sc{Q\_Trans}}\\
+@{thm[mode=Rule] Q_App[no_vars]}{\sc{Q\_App}}\\
+@{thm[mode=Rule] Q_Abs[no_vars]}{\sc{Q\_Abs}}\\
+@{thm[mode=Rule] Q_Beta[no_vars]}{\sc{Q\_Beta}}\\
+@{thm[mode=Rule] Q_Ext[no_vars]}{\sc{Q\_Ext}}\\
+\end{center}
+*}
+
+text {* It is now a tradition, we derive the lemma about validity, and we generate the equivariance lemma. *}
+
+lemma equiv_def_valid:
+  assumes "\<Gamma> \<turnstile> t \<equiv> s : T"
+  shows "valid \<Gamma>"
+using assms by (induct,auto elim:typing_valid)
+
+nominal_inductive equiv_def
+
+text{*
+\subsection{Strong induction principle for the equivalence relation}
+*}
+
+lemma equiv_def_strong_induct
+  [consumes 1, case_names Q_Refl Q_Symm Q_Trans Q_Abs Q_App Q_Beta Q_Ext]:
+  fixes c::"'a::fs_name"
+  assumes a0: "\<Gamma> \<turnstile> s \<equiv> t : T" 
+  and     a1: "\<And>\<Gamma> t T c.  \<Gamma> \<turnstile> t : T  \<Longrightarrow> P c \<Gamma> t t T"
+  and     a2: "\<And>\<Gamma> t s T c. \<lbrakk>\<And>d. P d \<Gamma> t s T\<rbrakk> \<Longrightarrow>  P c \<Gamma> s t T"
+  and     a3: "\<And>\<Gamma> s t T u c. \<lbrakk>\<And>d. P d \<Gamma> s t T; \<And>d. P d \<Gamma> t u T\<rbrakk> 
+               \<Longrightarrow> P c \<Gamma> s u T"
+  and     a4: "\<And>x \<Gamma> T\<^isub>1 s\<^isub>2 t\<^isub>2 T\<^isub>2 c. \<lbrakk>x\<sharp>\<Gamma>; x\<sharp>c; \<And>d. P d ((x,T\<^isub>1)#\<Gamma>) s\<^isub>2 t\<^isub>2 T\<^isub>2\<rbrakk>
+               \<Longrightarrow> P c \<Gamma> (Lam [x].s\<^isub>2) (Lam [x].t\<^isub>2) (T\<^isub>1\<rightarrow>T\<^isub>2)"
+  and     a5: "\<And>\<Gamma> s\<^isub>1 t\<^isub>1 T\<^isub>1 T\<^isub>2 s\<^isub>2 t\<^isub>2 c. \<lbrakk>\<And>d. P d \<Gamma> s\<^isub>1 t\<^isub>1 (T\<^isub>1\<rightarrow>T\<^isub>2); \<And>d. P d \<Gamma> s\<^isub>2 t\<^isub>2 T\<^isub>1\<rbrakk> 
+               \<Longrightarrow> P c \<Gamma> (App s\<^isub>1 s\<^isub>2) (App t\<^isub>1 t\<^isub>2) T\<^isub>2"
+  and     a6: "\<And>x \<Gamma> T\<^isub>1 s12 t12 T\<^isub>2 s\<^isub>2 t\<^isub>2 c.
+               \<lbrakk>x\<sharp>(\<Gamma>,s\<^isub>2,t\<^isub>2); x\<sharp>c; \<And>d. P d ((x,T\<^isub>1)#\<Gamma>) s12 t12 T\<^isub>2; \<And>d. P d \<Gamma> s\<^isub>2 t\<^isub>2 T\<^isub>1\<rbrakk> 
+               \<Longrightarrow> P c \<Gamma> (App (Lam [x].s12) s\<^isub>2) (t12[x::=t\<^isub>2]) T\<^isub>2"
+  and     a7: "\<And>x \<Gamma> s t T\<^isub>1 T\<^isub>2 c.
+               \<lbrakk>x\<sharp>(\<Gamma>,s,t); \<And>d. P d ((x,T\<^isub>1)#\<Gamma>) (App s (Var x)) (App t (Var x)) T\<^isub>2\<rbrakk>
+               \<Longrightarrow> P c \<Gamma> s t (T\<^isub>1\<rightarrow>T\<^isub>2)"
+ shows "P c \<Gamma>  s t T"
+proof -
+  from a0 have "\<And>(pi::name prm) c. P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (pi\<bullet>T)" 
+    proof(induct)
+      case (Q_Refl \<Gamma> t T pi c)
+      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>t) (pi\<bullet>t) (pi\<bullet>T)" using a1 by (simp only: eqvt)
+    next
+      case (Q_Symm \<Gamma> t s T pi c)
+      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (pi\<bullet>T)" using a2 by (simp only: equiv_def_eqvt)
+    next
+      case (Q_Trans \<Gamma> s t T u pi c)
+      then show " P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>u) (pi\<bullet>T)" using a3 by (blast intro: equiv_def_eqvt)
+    next
+      case (Q_App \<Gamma> s\<^isub>1 t\<^isub>1 T\<^isub>1 T\<^isub>2 s\<^isub>2 t\<^isub>2 pi c)
+      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>App s\<^isub>1 s\<^isub>2) (pi\<bullet>App t\<^isub>1 t\<^isub>2) (pi\<bullet>T\<^isub>2)" using a5 
+	by (simp only: eqvt) (blast)
+    next
+      case (Q_Ext x \<Gamma> s t T\<^isub>1 T\<^isub>2 pi c)
+      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (pi\<bullet>T\<^isub>1\<rightarrow>T\<^isub>2)"  
+	by (auto intro!: a7 simp add: fresh_bij)
+    next
+      case (Q_Abs x \<Gamma> T\<^isub>1 s\<^isub>2 t\<^isub>2 T\<^isub>2 pi c)
+      obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>s\<^isub>2,pi\<bullet>t\<^isub>2,pi\<bullet>\<Gamma>,c)" by (rule exists_fresh[OF fs_name1])
+      let ?sw="[(pi\<bullet>x,y)]"
+      let ?pi'="?sw@pi"
+      have f1: "x\<sharp>\<Gamma>" by fact
+      have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>)" using f1 by (simp add: fresh_bij)
+      have f3: "y\<sharp>?pi'\<bullet>\<Gamma>" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
+      have ih1: "\<And>c. P c (?pi'\<bullet>((x,T\<^isub>1)#\<Gamma>)) (?pi'\<bullet>s\<^isub>2) (?pi'\<bullet>t\<^isub>2) (?pi'\<bullet>T\<^isub>2)" by fact
+      then have "\<And>c. P c ((y,T\<^isub>1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>s\<^isub>2) (?pi'\<bullet>t\<^isub>2) T\<^isub>2" by (simp add: calc_atm)
+      then have "P c  (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>Lam [x].s\<^isub>2) (?pi'\<bullet>Lam [x].t\<^isub>2) (T\<^isub>1\<rightarrow>T\<^isub>2)" using a4 f3 fs
+	by (simp add: calc_atm)
+      then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>Lam [(pi\<bullet>x)].(pi\<bullet>s\<^isub>2)) (?sw\<bullet>Lam [(pi\<bullet>x)].(pi\<bullet>t\<^isub>2)) (T\<^isub>1\<rightarrow>T\<^isub>2)" 
+	by (simp del: append_Cons add: calc_atm pt_name2)
+      moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs f2)
+      moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>s\<^isub>2))) = Lam [(pi\<bullet>x)].(pi\<bullet>s\<^isub>2)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
+      moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t\<^isub>2))) = Lam [(pi\<bullet>x)].(pi\<bullet>t\<^isub>2)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
+      ultimately have "P c (pi\<bullet>\<Gamma>) (Lam [(pi\<bullet>x)].(pi\<bullet>s\<^isub>2)) (Lam [(pi\<bullet>x)].(pi\<bullet>t\<^isub>2)) (T\<^isub>1\<rightarrow>T\<^isub>2)" by simp
+      then show  "P c (pi\<bullet>\<Gamma>) (pi\<bullet>Lam [x].s\<^isub>2) (pi\<bullet>Lam [x].t\<^isub>2) (pi\<bullet>T\<^isub>1\<rightarrow>T\<^isub>2)" by simp 
+    next 
+      case (Q_Beta x \<Gamma> s\<^isub>2 t\<^isub>2 T\<^isub>1 s12 t12 T\<^isub>2 pi c) 
+      obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>s12,pi\<bullet>t12,pi\<bullet>s\<^isub>2,pi\<bullet>t\<^isub>2,pi\<bullet>\<Gamma>,c)" 
+	by (rule exists_fresh[OF fs_name1])
+      let ?sw="[(pi\<bullet>x,y)]"
+      let ?pi'="?sw@pi"
+      have f1: "x\<sharp>(\<Gamma>,s\<^isub>2,t\<^isub>2)" by fact 
+      have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>(\<Gamma>,s\<^isub>2,t\<^isub>2))" using f1 by (simp add: fresh_bij)
+      have f3: "y\<sharp>(?pi'\<bullet>(\<Gamma>,s\<^isub>2,t\<^isub>2))" using f1 
+	by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp fresh_prod)
+      have ih1: "\<And>c. P c (?pi'\<bullet>((x,T\<^isub>1)#\<Gamma>)) (?pi'\<bullet>s12) (?pi'\<bullet>t12) (?pi'\<bullet>T\<^isub>2)" by fact
+      then have "\<And>c. P c ((y,T\<^isub>1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>s12) (?pi'\<bullet>t12) (?pi'\<bullet>T\<^isub>2)" by (simp add: calc_atm)
+      moreover
+      have ih2: "\<And>c. P c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>s\<^isub>2) (?pi'\<bullet>t\<^isub>2) (?pi'\<bullet>T\<^isub>1)" by fact
+      ultimately have "P c  (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>(App (Lam [x].s12) s\<^isub>2)) (?pi'\<bullet>(t12[x::=t\<^isub>2])) (?pi'\<bullet>T\<^isub>2)" 
+	using a6 f3 fs by (force simp add: subst_eqvt calc_atm del: perm_ty)
+      then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s\<^isub>2))) 
+	(?sw\<bullet>((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t\<^isub>2)])) T\<^isub>2" 
+	by (simp del: append_Cons add: calc_atm pt_name2 subst_eqvt)
+      moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
+      moreover have "(?sw\<bullet>(App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s\<^isub>2))) = App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s\<^isub>2)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified] abs_fresh)
+      moreover have "(?sw\<bullet>((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t\<^isub>2)])) = ((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t\<^isub>2)])" 
+	by (rule perm_fresh_fresh) 
+	   (simp_all add: fs[simplified] f2[simplified] fresh_subst fresh_subst'')
+      ultimately have "P c (pi\<bullet>\<Gamma>) (App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s\<^isub>2)) ((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t\<^isub>2)]) T\<^isub>2"
+	by simp
+      then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>App (Lam [x].s12) s\<^isub>2) (pi\<bullet>t12[x::=t\<^isub>2]) (pi\<bullet>T\<^isub>2)" by (simp add: subst_eqvt)
+    qed
+  then have "P c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>s) (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>T)" by blast
+  then show "P c \<Gamma> s t T" by simp
+qed
+
+
+text {* \section{Type-driven equivalence algorithm} 
+
+We follow the original presentation. The algorithm is described using inference rules only.
+
+*}
+
+text {* \subsection{Weak head reduction} *}
+
+(*<*)
+inductive2
+  whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80) 
+where
+  QAR_Beta[intro]: "App (Lam [x]. t12) t\<^isub>2 \<leadsto> t12[x::=t\<^isub>2]"
+| QAR_App[intro]: "t1 \<leadsto> t1' \<Longrightarrow> App t1 t\<^isub>2 \<leadsto> App t1' t\<^isub>2"
+(*>*)
+
+text {* 
+\begin{figure}
+\begin{center}
+\isastyle
+@{thm[mode=Rule] QAR_Beta[no_vars]}{\sc{QAR\_Beta}} \qquad
+@{thm[mode=Rule] QAR_App[no_vars]}{\sc{QAR\_App}}
+\end{center}
+\end{figure}
+*}
+
+text {* \subsubsection{Inversion lemma for weak head reduction} *}
+
+declare trm.inject  [simp add]
+declare ty.inject  [simp add]
+
+inductive_cases2 whr_Gen[elim]: "t \<leadsto> t'"
+inductive_cases2 whr_Lam[elim]: "Lam [x].t \<leadsto> t'"
+inductive_cases2 whr_App_Lam[elim]: "App (Lam [x].t12) t2 \<leadsto> t"
+inductive_cases2 whr_Var[elim]: "Var x \<leadsto> t"
+inductive_cases2 whr_Const[elim]: "Const n \<leadsto> t"
+inductive_cases2 whr_App[elim]: "App p q \<leadsto> t"
+inductive_cases2 whr_Const_Right[elim]: "t \<leadsto> Const n"
+inductive_cases2 whr_Var_Right[elim]: "t \<leadsto> Var x"
+inductive_cases2 whr_App_Right[elim]: "t \<leadsto> App p q"
+
+declare trm.inject  [simp del]
+declare ty.inject  [simp del]
+
+nominal_inductive whr_def
+
+text {* 
+\subsection{Weak head normalization} 
+\subsubsection{Definition of the normal form}
+*}
+
+abbreviation 
+ nf :: "trm \<Rightarrow> bool" ("_ \<leadsto>|" [100] 100)
+where
+  "t\<leadsto>|  \<equiv> \<not>(\<exists> u. t \<leadsto> u)" 
+
+text{* \subsubsection{Definition of weak head normal form} *}
+
+(*<*)
+inductive2
+  whn_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80) 
+where
+  QAN_Reduce[intro]: "\<lbrakk>s \<leadsto> t; t \<Down> u\<rbrakk> \<Longrightarrow> s \<Down> u"
+| QAN_Normal[intro]: "t\<leadsto>|  \<Longrightarrow> t \<Down> t"
+
+(*>*)
+
+text {* 
+\begin{center}
+@{thm[mode=Rule] QAN_Reduce[no_vars]}{\sc{QAN\_Reduce}}\qquad
+@{thm[mode=Rule] QAN_Normal[no_vars]}{\sc{QAN\_Normal}}
+\end{center}
+*}
+
+lemma whn_eqvt[eqvt]:
+  fixes pi::"name prm"
+  assumes a: "t \<Down> t'"
+  shows "(pi\<bullet>t) \<Down> (pi\<bullet>t')"
+using a
+apply(induct)
+apply(rule QAN_Reduce)
+apply(rule whr_def_eqvt)
+apply(assumption)+
+apply(rule QAN_Normal)
+apply(auto)
+apply(drule_tac pi="rev pi" in whr_def_eqvt)
+apply(perm_simp)
+done
+
+text {* \subsection{Algorithmic term equivalence and algorithmic path equivalence} *}
+
+(*<*)
+inductive2
+  alg_equiv :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<Leftrightarrow> _ : _" [60,60] 60) 
+and
+  alg_path_equiv :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60] 60) 
+where
+  QAT_Base[intro]:  "\<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
+| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>\<Gamma>; x\<sharp>s; x\<sharp>t; (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2\<rbrakk> 
+                     \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+| QAT_One[intro]:   "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : TUnit"
+| QAP_Var[intro]:   "\<lbrakk>valid \<Gamma>;(x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T"
+| QAP_App[intro]:   "\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1 \<rightarrow> T\<^isub>2; \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App p s \<leftrightarrow> App q t : T\<^isub>2"
+| QAP_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n \<leftrightarrow> Const n : TBase"
+(*>*)
+
+text {* 
+\begin{center}
+@{thm[mode=Rule] QAT_Base[no_vars]}{\sc{QAT\_Base}}\\
+@{thm[mode=Rule] QAT_Arrow[no_vars]}{\sc{QAT\_Arrow}}\\
+@{thm[mode=Rule] QAP_Var[no_vars]}{\sc{QAP\_Var}}\\
+@{thm[mode=Rule] QAP_App[no_vars]}{\sc{QAP\_App}}\\
+@{thm[mode=Rule] QAP_Const[no_vars]}{\sc{QAP\_Const}}\\
+\end{center}
+*}
+
+text {* Again we generate the equivariance lemma. *}
+
+nominal_inductive  alg_equiv
+
+text {* \subsubsection{Strong induction principle for algorithmic term and path equivalences} *}
+
+lemma alg_equiv_alg_path_equiv_strong_induct
+  [case_names QAT_Base QAT_Arrow QAT_One QAP_Var QAP_App QAP_Const]:
+  fixes c::"'a::fs_name"
+  assumes a1: "\<And>s p t q \<Gamma> c. \<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase; \<And>d. P2 d \<Gamma> p q TBase\<rbrakk> 
+               \<Longrightarrow> P1 c \<Gamma> s t TBase"
+  and     a2: "\<And>x \<Gamma> s t T\<^isub>1 T\<^isub>2 c.
+               \<lbrakk>x\<sharp>\<Gamma>; x\<sharp>s; x\<sharp>t; x\<sharp>c; (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2;
+               \<And>d. P1 d ((x,T\<^isub>1)#\<Gamma>) (App s (Var x)) (App t (Var x)) T\<^isub>2\<rbrakk>
+               \<Longrightarrow> P1 c \<Gamma> s t (T\<^isub>1\<rightarrow>T\<^isub>2)"
+  and     a3: "\<And>\<Gamma> s t c. valid \<Gamma> \<Longrightarrow> P1 c \<Gamma> s t TUnit"
+  and     a4: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P2 c \<Gamma> (Var x) (Var x) T"
+  and     a5: "\<And>\<Gamma> p q T\<^isub>1 T\<^isub>2 s t c.
+               \<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2; \<And>d. P2 d \<Gamma> p q (T\<^isub>1\<rightarrow>T\<^isub>2); \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1; \<And>d. P1 d \<Gamma> s t T\<^isub>1\<rbrakk>
+               \<Longrightarrow> P2 c \<Gamma> (App p s) (App q t) T\<^isub>2"
+  and     a6: "\<And>\<Gamma> n c. valid \<Gamma> \<Longrightarrow> P2 c \<Gamma> (Const n) (Const n) TBase"
+  shows "(\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<longrightarrow>P1 c \<Gamma> s t T) \<and> (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> P2 c \<Gamma>' s' t' T')"
+proof -
+  have "(\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<longrightarrow> (\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) T)) \<and> 
+        (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> (\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>') (pi\<bullet>s') (pi\<bullet>t') T'))"
+  proof(induct rule: alg_equiv_alg_path_equiv.induct)
+    case (QAT_Base s q t p \<Gamma>)
+    then show "\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) TBase"
+      apply(auto)
+      apply(rule_tac p="pi\<bullet>q" and q="pi\<bullet>p" in  a1)
+      apply(simp_all add: eqvt[simplified])
+      done
+  next
+    case (QAT_Arrow x \<Gamma> s t T\<^isub>1 T\<^isub>2)
+    show ?case
+    proof (rule allI, rule allI)
+      fix c::"'a::fs_name" and pi::"name prm"
+      obtain y::"name" where fs: "y\<sharp>(pi\<bullet>s,pi\<bullet>t,pi\<bullet>\<Gamma>,c)" by (rule exists_fresh[OF fs_name1])
+      let ?sw="[(pi\<bullet>x,y)]"
+      let ?pi'="?sw@pi"
+      have f0: "x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
+      then have f1: "x\<sharp>(\<Gamma>,s,t)" by simp
+      have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>(\<Gamma>,s,t))" using f1 by (simp add: fresh_bij)
+      have f3: "y\<sharp>?pi'\<bullet>(\<Gamma>,s,t)" using f1 
+	by (simp only: pt_name2 fresh_left, auto simp add: perm_pi_simp calc_atm)
+      then have f3': "y\<sharp>?pi'\<bullet>\<Gamma>" "y\<sharp>?pi'\<bullet>s" "y\<sharp>?pi'\<bullet>t" by (auto simp add: fresh_prod)
+      have pr1: "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" by fact
+      then have "?pi'\<bullet> ((x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2)" using perm_bool by simp
+      then have "(?pi'\<bullet>((x,T\<^isub>1)#\<Gamma>)) \<turnstile> (?pi'\<bullet>(App s (Var x))) \<Leftrightarrow> (?pi'\<bullet>(App t (Var x))) : (?pi'\<bullet>T\<^isub>2)" 
+	by (auto simp add: eqvt)
+      then have "(y,T\<^isub>1)#(?pi'\<bullet>\<Gamma>) \<turnstile> (App (?pi'\<bullet>s) (Var y)) \<Leftrightarrow> (App (?pi'\<bullet>t) (Var y)) : T\<^isub>2" 
+	by (simp add: calc_atm)
+      moreover    
+      have ih1: "\<forall>c (pi::name prm).  P1 c (pi\<bullet>((x,T\<^isub>1)#\<Gamma>)) (pi\<bullet>App s (Var x)) (pi\<bullet>App t (Var x)) T\<^isub>2"
+	by fact
+      then have "\<And>c.  P1 c (?pi'\<bullet>((x,T\<^isub>1)#\<Gamma>)) (?pi'\<bullet>App s (Var x)) (?pi'\<bullet>App t (Var x)) T\<^isub>2"
+	by auto
+      then have "\<And>c.  P1 c ((y,T\<^isub>1)#(?pi'\<bullet>\<Gamma>)) (App (?pi'\<bullet>s) (Var y)) (App (?pi'\<bullet>t) (Var y)) T\<^isub>2"
+	by (simp add: calc_atm)     
+      ultimately have "P1 c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>s) (?pi'\<bullet>t) (T\<^isub>1\<rightarrow>T\<^isub>2)" using a2 f3' fs by simp
+      then have "P1 c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>pi\<bullet>s) (?sw\<bullet>pi\<bullet>t) (T\<^isub>1\<rightarrow>T\<^isub>2)" 
+	by (simp del: append_Cons add: calc_atm pt_name2)
+      moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
+      moreover have "(?sw\<bullet>(pi\<bullet>s)) = (pi\<bullet>s)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
+      moreover have "(?sw\<bullet>(pi\<bullet>t)) = (pi\<bullet>t)" 
+	by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
+      ultimately show "P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (T\<^isub>1\<rightarrow>T\<^isub>2)" by (simp)
+    qed
+  next
+    case (QAT_One \<Gamma> s t)
+    then show "\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) TUnit"
+      by (auto intro!: a3 simp add: valid_eqvt)
+  next
+    case (QAP_Var \<Gamma> x T)
+    then show "\<forall>c (pi::name prm). P2 c (pi \<bullet> \<Gamma>) (pi \<bullet> Var x) (pi \<bullet> Var x) T"
+      apply(auto intro!: a4 simp add: valid_eqvt)
+      apply(simp add: set_eqvt[symmetric])
+      apply(perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
+      done
+  next
+    case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t)
+    then show "\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>) (pi\<bullet>App p s) (pi\<bullet>App q t) T\<^isub>2"
+      by (auto intro!: a5 simp add: eqvt[simplified])
+  next
+    case (QAP_Const \<Gamma> n)
+    then show "\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>) (pi\<bullet>Const n) (pi\<bullet>Const n) TBase"
+      by (auto intro!: a6 simp add: eqvt)
+  qed
+  then have "(\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<longrightarrow> P1 c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>s) (([]::name prm)\<bullet>t) T) \<and> 
+             (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> P2 c (([]::name prm)\<bullet>\<Gamma>') (([]::name prm)\<bullet>s') (([]::name prm)\<bullet>t') T')"
+    by blast
+  then show ?thesis by simp
+qed
+
+(* post-processing of the strong induction principle proved above; *) 
+(* the code extracts the strong_inducts-version from strong_induct *)
+(*<*)
+setup {*
+  PureThy.add_thmss
+    [(("alg_equiv_alg_path_equiv_strong_inducts",
+      ProjectRule.projects (ProofContext.init (the_context())) [1,2]
+        (thm "alg_equiv_alg_path_equiv_strong_induct")), [])] #> snd
+*}
+(*>*)
+
+lemma alg_equiv_valid:
+  shows  "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> valid \<Gamma>" and  "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> valid \<Gamma>"
+by (induct rule : alg_equiv_alg_path_equiv.inducts, auto)
+
+text{* \subsubsection{Inversion lemmas for algorithmic term and path equivalences} *}
+
+declare trm.inject  [simp add]
+declare ty.inject  [simp add]
+
+inductive_cases2 whn_inv_auto[elim]: "t \<Down> t'"
+
+inductive_cases2 alg_equiv_Base_inv_auto[elim]: "\<Gamma> \<turnstile> s\<Leftrightarrow>t : TBase"
+inductive_cases2 alg_equiv_Arrow_inv_auto[elim]: "\<Gamma> \<turnstile> s\<Leftrightarrow>t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+
+inductive_cases2 alg_path_equiv_Base_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : TBase"
+inductive_cases2 alg_path_equiv_Unit_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : TUnit"
+inductive_cases2 alg_path_equiv_Arrow_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : T\<^isub>1 \<rightarrow> T\<^isub>2"
+
+inductive_cases2 alg_path_equiv_Var_left_inv_auto[elim]: "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T"
+inductive_cases2 alg_path_equiv_Var_left_inv_auto'[elim]: "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T'"
+inductive_cases2 alg_path_equiv_Var_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T"
+inductive_cases2 alg_path_equiv_Var_right_inv_auto'[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T'"
+inductive_cases2 alg_path_equiv_Const_left_inv_auto[elim]: "\<Gamma> \<turnstile> Const n \<leftrightarrow> t : T"
+inductive_cases2 alg_path_equiv_Const_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Const n : T"
+inductive_cases2 alg_path_equiv_App_left_inv_auto[elim]: "\<Gamma> \<turnstile> App p s \<leftrightarrow> t : T"
+inductive_cases2 alg_path_equiv_App_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> App q t : T"
+inductive_cases2 alg_path_equiv_Lam_left_inv_auto[elim]: "\<Gamma> \<turnstile> Lam[x].s \<leftrightarrow> t : T"
+inductive_cases2 alg_path_equiv_Lam_right_inv_auto[elim]: "\<Gamma> \<turnstile> t \<leftrightarrow> Lam[x].s : T"
+
+declare trm.inject [simp del]
+declare ty.inject [simp del]
+
+text {* \section{Completeness of the algorithm} *}
+
+lemma algorithmic_monotonicity:
+  fixes \<Gamma>':: "(name\<times>ty) list"
+  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<Leftrightarrow> t : T"
+  and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<leftrightarrow> t : T"
+proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: \<Gamma>' rule: alg_equiv_alg_path_equiv_strong_inducts)
+ case (QAT_Arrow x \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>')
+  have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
+  have h2:"\<Gamma>\<lless>\<Gamma>'" by fact
+  have ih:"\<And>\<Gamma>'. \<lbrakk>(x,T\<^isub>1)#\<Gamma> \<lless> \<Gamma>'\<rbrakk>  \<Longrightarrow> \<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" by fact
+  have "x\<sharp>\<Gamma>'" by fact
+  then have sub:"(x,T\<^isub>1)#\<Gamma> \<lless> (x,T\<^isub>1)#\<Gamma>'" using h2 by auto
+  then have "(x,T\<^isub>1)#\<Gamma>' \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" using ih by auto
+  then show "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rightarrow>T\<^isub>2" using sub fs by auto
+qed (auto)
+
+lemma valid_monotonicity[elim]:
+ assumes "\<Gamma> \<lless> \<Gamma>'" and "x\<sharp>\<Gamma>'"
+ shows "(x,T\<^isub>1)#\<Gamma> \<lless> (x,T\<^isub>1)#\<Gamma>'"
+using assms by auto
+
+lemma algorithmic_monotonicity_auto:
+  fixes \<Gamma>':: "(name\<times>ty) list"
+  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<Leftrightarrow> t : T"
+  and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<leftrightarrow> t : T"
+by (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: \<Gamma>' rule: alg_equiv_alg_path_equiv_strong_inducts, auto simp add: valid_monotonicity)
  
-lemma fresh_subst_fresh:
-    assumes "a\<sharp>e"
-    shows "a\<sharp>t[a::=e]"
+text {* \subsection{Definition of the logical relation} *}
+
+function log_equiv :: "((name\<times>ty) list \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)"   
+                      ("_ \<turnstile> _ is _ : _" [60,60,60,60] 60) 
+where    
+   "\<Gamma> \<turnstile> s is t : TUnit = valid \<Gamma>"
+ | "\<Gamma> \<turnstile> s is t : TBase = \<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase"
+ | "\<Gamma> \<turnstile> s is t : (T\<^isub>1 \<rightarrow> T\<^isub>2) =  
+           (valid \<Gamma> \<and> (\<forall>\<Gamma>' s' t'. \<Gamma>\<lless>\<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^isub>1 \<longrightarrow>  (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^isub>2)))"
+apply (auto simp add: ty.inject)
+apply (subgoal_tac "(\<exists>T\<^isub>1 T\<^isub>2. b=T\<^isub>1 \<rightarrow> T\<^isub>2) \<or> b=TUnit \<or> b=TBase" )
+apply (force)
+apply (rule ty_cases)
+done
+
+termination
+apply(relation "measure (\<lambda>(_,_,_,T). size T)")
+apply(auto)
+done
+
+lemma log_equiv_valid: 
+  assumes "\<Gamma> \<turnstile> s is t : T"
+  shows "valid \<Gamma>"
 using assms 
-by (nominal_induct t avoiding: a e rule: trm.induct)
-   (auto simp add: fresh_atm abs_fresh fresh_nat) 
+by (induct rule: log_equiv.induct,auto elim: alg_equiv_valid)
+
+text {* \subsubsection{Monotonicity of the logical equivalence relation} *}
+
+lemma logical_monotonicity :
+ fixes T::ty
+ assumes "\<Gamma> \<turnstile> s is t : T" "\<Gamma>\<lless>\<Gamma>'" 
+ shows "\<Gamma>' \<turnstile> s is t : T"
+using assms
+proof (induct arbitrary: \<Gamma>' rule: log_equiv.induct)
+  case (2 \<Gamma> s t \<Gamma>')
+  then show "\<Gamma>' \<turnstile> s is t : TBase" using algorithmic_monotonicity by auto
+next
+  case (3 \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>')
+  then show "\<Gamma>' \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by force 
+qed (auto)
+
+text {* If two terms are path equivalent then they are in normal form. *}
 
 lemma path_equiv_implies_nf:
   assumes "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
@@ -844,44 +1023,38 @@
 using assms
 by (induct rule: alg_equiv_alg_path_equiv.inducts(2)) (simp, auto)
 
-lemma path_equiv_implies_nf:
-  shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> True"
-    and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> s \<leadsto>|"
-        "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> t \<leadsto>|"
-by (induct rule: alg_equiv_alg_path_equiv.inducts) (auto)
-
 lemma main_lemma:
-  shows "\<Gamma> \<turnstile> s is t : T \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T" and "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> \<Gamma> \<turnstile> p is q : T"
+  shows "\<Gamma> \<turnstile> s is t : T \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T" and "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> \<Gamma> \<turnstile> p is q : T"
 proof (nominal_induct T arbitrary: \<Gamma> s t p q rule:ty.induct)
-  case (Arrow T1 T2)
+  case (Arrow T\<^isub>1 T\<^isub>2)
   { 
     case (1 \<Gamma> s t)
-    have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s is t : T2 \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T2" by fact
-    have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T1 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T1" by fact
-    have h:"\<Gamma> \<turnstile> s is t : T1\<rightarrow>T2" by fact
+    have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s is t : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>2" by fact
+    have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^isub>1" by fact
+    have h:"\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
     obtain x::name where fs:"x\<sharp>(\<Gamma>,s,t)" by (erule exists_fresh[OF fs_name1])
     have "valid \<Gamma>" using h log_equiv_valid by auto
-    then have v:"valid ((x,T1)#\<Gamma>)" using fs by auto
-    then have "(x,T1)#\<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T1" by auto
-    then have "(x,T1)#\<Gamma> \<turnstile> Var x is Var x : T1" using ih2 by auto
-    then have "(x,T1)#\<Gamma> \<turnstile> App s (Var x) is App t (Var x) : T2" using h v by auto
-    then have "(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2" using ih1 by auto
-    then show "\<Gamma> \<turnstile> s <=> t : T1\<rightarrow>T2" using fs by (auto simp add: fresh_prod)
+    then have v:"valid ((x,T\<^isub>1)#\<Gamma>)" using fs by auto
+    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T\<^isub>1" by auto
+    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> Var x is Var x : T\<^isub>1" using ih2 by auto
+    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) is App t (Var x) : T\<^isub>2" using h v by auto
+    then have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App t (Var x) : T\<^isub>2" using ih1 by auto
+    then show "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by (auto simp add: fresh_prod)
   next
     case (2 \<Gamma> p q)
-    have h:"\<Gamma> \<turnstile> p \<leftrightarrow> q : T1\<rightarrow>T2" by fact
-    have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T2 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T2" by fact
-    have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s is t : T1 \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T1" by fact
+    have h:"\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+    have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T\<^isub>2" by fact
+    have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s is t : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" by fact
     {
       fix \<Gamma>' s t
-      assume "\<Gamma>\<lless>\<Gamma>'" and hl:"\<Gamma>' \<turnstile> s is t : T1"
-      then have "\<Gamma>' \<turnstile> p \<leftrightarrow> q : T1 \<rightarrow> T2" using h algorithmic_monotonicity by auto
-      moreover have "\<Gamma>' \<turnstile> s <=> t : T1" using ih2 hl by auto
-      ultimately have "\<Gamma>' \<turnstile> App p s \<leftrightarrow> App q t : T2" by auto
-      then have "\<Gamma>' \<turnstile> App p s is App q t : T2" using ih1 by auto
+      assume "\<Gamma>\<lless>\<Gamma>'" and hl:"\<Gamma>' \<turnstile> s is t : T\<^isub>1"
+      then have "\<Gamma>' \<turnstile> p \<leftrightarrow> q : T\<^isub>1 \<rightarrow> T\<^isub>2" using h algorithmic_monotonicity by auto
+      moreover have "\<Gamma>' \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" using ih2 hl by auto
+      ultimately have "\<Gamma>' \<turnstile> App p s \<leftrightarrow> App q t : T\<^isub>2" by auto
+      then have "\<Gamma>' \<turnstile> App p s is App q t : T\<^isub>2" using ih1 by auto
     }
     moreover have "valid \<Gamma>" using h alg_equiv_valid by auto
-    ultimately show "\<Gamma> \<turnstile> p is q : T1\<rightarrow>T2"  by auto
+    ultimately show "\<Gamma> \<turnstile> p is q : T\<^isub>1\<rightarrow>T\<^isub>2"  by auto
   }
 next
   case TBase
@@ -889,30 +1062,20 @@
     have h:"\<Gamma> \<turnstile> s \<leftrightarrow> t : TBase" by fact
     then have "s \<leadsto>|" and "t \<leadsto>|" using path_equiv_implies_nf by auto
     then have "s \<Down> s" and "t \<Down> t" by auto
-    then have "\<Gamma> \<turnstile> s <=> t : TBase" using h by auto
+    then have "\<Gamma> \<turnstile> s \<Leftrightarrow> t : TBase" using h by auto
     then show "\<Gamma> \<turnstile> s is t : TBase" by auto
   }
 qed (auto elim:alg_equiv_valid)
 
 corollary corollary_main:
   assumes a: "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
-  shows "\<Gamma> \<turnstile> s <=> t : T"
+  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T"
 using a main_lemma by blast
 
-lemma algorithmic_symmetry_aux:
-  shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> \<Gamma> \<turnstile> t <=> s : T" 
-  and   "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<leftrightarrow> s : T"
-by (induct rule: alg_equiv_alg_path_equiv.inducts) (auto)
-
 lemma algorithmic_symmetry:
-  assumes a: "\<Gamma> \<turnstile> s <=> t : T"
-  shows "\<Gamma> \<turnstile> t <=> s : T"
-using a by (simp add: algorithmic_symmetry_aux)
-
-lemma algorithmic_path_symmetry:
-  assumes a: "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
-  shows "\<Gamma> \<turnstile> t \<leftrightarrow> s : T"
-using a by (simp add: algorithmic_symmetry_aux)
+  shows "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<Leftrightarrow> s : T" 
+  and   "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<leftrightarrow> s : T"
+by (induct rule: alg_equiv_alg_path_equiv.inducts) (auto)
 
 lemma red_unicity : 
   assumes a: "x \<leadsto> a" 
@@ -943,72 +1106,49 @@
   then show "a=b" using ih hl by auto
 qed (auto)
 
-lemma Q_eqvt :
-  fixes pi:: "name prm"
-  shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) <=> (pi\<bullet>t) : T"
-  and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) \<leftrightarrow> (pi\<bullet>t) : T"
-using assms
-proof (induct rule:alg_equiv_alg_path_equiv.inducts)
-  case (QAP_Var \<Gamma> x T)
-  then have "pi\<bullet>(x,T) \<in> (pi\<bullet>(set \<Gamma>))" using in_eqvt by blast
-  then have "(pi\<bullet>x,T) \<in> (set (pi\<bullet>\<Gamma>))" using set_eqvt perm_ty by auto
-  moreover have "valid \<Gamma>" by fact
-  ultimately show ?case using valid_eqvt by auto
-next
-  case (QAT_Arrow x \<Gamma> s t T1 T2)
-  then have h:"((pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>))" "((pi\<bullet>x)\<sharp>(pi\<bullet>s))" "((pi\<bullet>x)\<sharp>(pi\<bullet>t))"  using fresh_bij by auto
-  have "(pi\<bullet>((x,T1)#\<Gamma>)) \<turnstile> pi \<bullet> (App s (Var x)) <=> pi \<bullet> (App t (Var x)) : T2" by fact
-  then have "(pi\<bullet>((x,T1)#\<Gamma>)) \<turnstile> (App (pi\<bullet>s) (Var (pi\<bullet>x))) <=> (App (pi\<bullet>t) (Var (pi\<bullet>x))) : T2" by auto
-  moreover have "pi\<bullet>((x,T1)#\<Gamma>) = (pi\<bullet>x,pi\<bullet>T1)#(pi\<bullet>\<Gamma>)" by auto
-  ultimately have "((pi\<bullet>x,T1)#(pi\<bullet>\<Gamma>))  \<turnstile> (App (pi\<bullet>s) (Var (pi\<bullet>x))) <=> (App (pi\<bullet>t) (Var (pi\<bullet>x))) : T2" 
-    using perm_ty by auto
-  then show ?case using h by auto
-qed (auto elim:whn_eqvt valid_eqvt)
- 
+nominal_inductive alg_equiv
+
 (* FIXME this lemma should not be here I am reinventing the wheel I guess *)
-
+(*<*)
 lemma perm_basic[simp]:
  fixes x y::"name"
 shows "[(x,y)]\<bullet>y = x"
 using assms using calc_atm by perm_simp
+(*>*)
+
+text{* \subsection{Strong inversion lemma for algorithmic equivalence} *}
 
 lemma Q_Arrow_strong_inversion:
-  assumes fs:"x\<sharp>\<Gamma>" "x\<sharp>t" "x\<sharp>u" and h:"\<Gamma> \<turnstile> t <=> u : T1\<rightarrow>T2"
-  shows "(x,T1)#\<Gamma> \<turnstile> App t (Var x) <=> App u (Var x) : T2"
+  assumes fs:"x\<sharp>\<Gamma>" "x\<sharp>t" "x\<sharp>u" and h:"\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2"
+  shows "(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2"
 proof -
-  obtain y where  fs2:"y\<sharp>\<Gamma>" "y\<sharp>u" "y\<sharp>t" and "(y,T1)#\<Gamma> \<turnstile> App t (Var y) <=> App u (Var y) : T2" 
+  obtain y where  fs2:"y\<sharp>\<Gamma>" "y\<sharp>u" "y\<sharp>t" and "(y,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var y) \<Leftrightarrow> App u (Var y) : T\<^isub>2" 
     using h by auto
-  then have "([(x,y)]\<bullet>((y,T1)#\<Gamma>)) \<turnstile> [(x,y)]\<bullet> App t (Var y) <=> [(x,y)]\<bullet> App u (Var y) : T2" 
-    using Q_eqvt by blast
+  then have "([(x,y)]\<bullet>((y,T\<^isub>1)#\<Gamma>)) \<turnstile> [(x,y)]\<bullet> App t (Var y) \<Leftrightarrow> [(x,y)]\<bullet> App u (Var y) : T\<^isub>2" 
+    using  alg_equiv_eqvt[simplified] by blast
   then show ?thesis using fs fs2 by (perm_simp)
 qed
 
-lemma fresh_context: 
-  fixes  \<Gamma> :: "(name\<times>ty) list"
-  and    a :: "name"
-  assumes "a\<sharp>\<Gamma>"
-  shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
-using assms by(induct \<Gamma>, auto simp add: fresh_prod fresh_list_cons fresh_atm)
+text{* 
+For the \texttt{algorithmic\_transitivity} lemma we need a unicity property.
+But one has to be cautious, because this unicity property is true only for algorithmic path. 
+Indeed the following lemma is \textbf{false}:\\
+\begin{center}
+@{prop "\<lbrakk> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<Leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"}
+\end{center}
+Here is the counter example :\\
+\begin{center}
+@{prop "\<Gamma> \<turnstile> Const n \<Leftrightarrow> Const n : Tbase"} and @{prop "\<Gamma> \<turnstile> Const n \<Leftrightarrow> Const n : TUnit"}
+\end{center}
+*}
 
-lemma type_unicity_in_context:
-  fixes  \<Gamma> :: "(name\<times>ty)list"
-  and    pi:: "name prm"
-  and    a :: "name"
-  and    \<tau> :: "ty"
-  assumes "valid \<Gamma>" and "(x,T1) \<in> set \<Gamma>" and "(x,T2) \<in> set \<Gamma>"
-  shows "T1=T2"
-using assms by (induct \<Gamma>, auto dest!: fresh_context)
-
-(* 
-
-Warning: This lemma is false.
-
+(*
 lemma algorithmic_type_unicity:
-  shows "\<lbrakk> \<Gamma> \<turnstile> s <=> t : T ; \<Gamma> \<turnstile> s <=> u : T' \<rbrakk> \<Longrightarrow> T = T'"
+  shows "\<lbrakk> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<Leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"
   and "\<lbrakk> \<Gamma> \<turnstile> s \<leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"
 
 Here is the counter example : 
-\<Gamma> \<turnstile> Const n <=> Const n : Tbase and \<Gamma> \<turnstile> Const n <=> Const n : TUnit
+\<Gamma> \<turnstile> Const n \<Leftrightarrow> Const n : Tbase and \<Gamma> \<turnstile> Const n \<Leftrightarrow> Const n : TUnit
 
 *)
 
@@ -1022,56 +1162,56 @@
   moreover have "valid \<Gamma>" "(x,T) \<in> set \<Gamma>" by fact
   ultimately show "T=T'" using type_unicity_in_context by auto
 next
-  case (QAP_App \<Gamma> p q T1 T2 s t u T2')
-  have ih:"\<And>u T. \<Gamma> \<turnstile> p \<leftrightarrow> u : T \<Longrightarrow> T1\<rightarrow>T2 = T" by fact
-  have "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T2'" by fact
-  then obtain r t T1' where "u = App r t"  "\<Gamma> \<turnstile> p \<leftrightarrow> r : T1' \<rightarrow> T2'" by auto
-  then have "T1\<rightarrow>T2 = T1' \<rightarrow> T2'" by auto
-  then show "T2=T2'" using ty.inject by auto
+  case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t u T\<^isub>2')
+  have ih:"\<And>u T. \<Gamma> \<turnstile> p \<leftrightarrow> u : T \<Longrightarrow> T\<^isub>1\<rightarrow>T\<^isub>2 = T" by fact
+  have "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^isub>2'" by fact
+  then obtain r t T\<^isub>1' where "u = App r t"  "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1' \<rightarrow> T\<^isub>2'" by auto
+  then have "T\<^isub>1\<rightarrow>T\<^isub>2 = T\<^isub>1' \<rightarrow> T\<^isub>2'" by auto
+  then show "T\<^isub>2=T\<^isub>2'" using ty.inject by auto
 qed (auto)
 
 lemma algorithmic_transitivity:
-  shows "\<lbrakk> \<Gamma> \<turnstile> s <=> t : T ; \<Gamma> \<turnstile> t <=> u : T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s <=> u : T"
+  shows "\<lbrakk> \<Gamma> \<turnstile> s \<Leftrightarrow> t : T ; \<Gamma> \<turnstile> t \<Leftrightarrow> u : T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> u : T"
   and  "\<lbrakk> \<Gamma> \<turnstile> s \<leftrightarrow> t : T ; \<Gamma> \<turnstile> t \<leftrightarrow> u : T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<leftrightarrow> u : T"
 proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: u rule: alg_equiv_alg_path_equiv_strong_inducts)
   case (QAT_Base s p t q \<Gamma> u)
   have h:"s \<Down> p" "t \<Down> q" by fact
   have ih:"\<And>u. \<Gamma> \<turnstile> q \<leftrightarrow> u : TBase \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> u : TBase" by fact
-  have "\<Gamma> \<turnstile> t <=> u : TBase" by fact
+  have "\<Gamma> \<turnstile> t \<Leftrightarrow> u : TBase" by fact
   then obtain r q' where hl:"t \<Down> q'" "u \<Down> r" "\<Gamma> \<turnstile> q' \<leftrightarrow> r : TBase" by auto
   moreover have eq:"q=q'" using nf_unicity hl h by auto
   ultimately have "\<Gamma> \<turnstile> p \<leftrightarrow> r : TBase" using ih by auto
-  then show "\<Gamma> \<turnstile> s <=> u : TBase" using hl h by auto
+  then show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : TBase" using hl h by auto
 next
-  case (QAT_Arrow  x \<Gamma> s t T1 T2 u)
+  case (QAT_Arrow  x \<Gamma> s t T\<^isub>1 T\<^isub>2 u)
   have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
-  moreover have h:"\<Gamma> \<turnstile> t <=> u : T1\<rightarrow>T2" by fact
-  moreover then obtain y where "y\<sharp>\<Gamma>" "y\<sharp>t" "y\<sharp>u" and "(y,T1)#\<Gamma> \<turnstile> App t (Var y) <=> App u (Var y) : T2" 
+  moreover have h:"\<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+  moreover then obtain y where "y\<sharp>\<Gamma>" "y\<sharp>t" "y\<sharp>u" and "(y,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var y) \<Leftrightarrow> App u (Var y) : T\<^isub>2" 
     by auto
   moreover have fs2:"x\<sharp>u" by fact
-  ultimately have "(x,T1)#\<Gamma> \<turnstile> App t (Var x) <=> App u (Var x) : T2" using Q_Arrow_strong_inversion by blast
-  moreover have ih:"\<And> v. (x,T1)#\<Gamma> \<turnstile> App t (Var x) <=> v : T2 \<Longrightarrow> (x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> v : T2" 
+  ultimately have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" using Q_Arrow_strong_inversion by blast
+  moreover have ih:"\<And> v. (x,T\<^isub>1)#\<Gamma> \<turnstile> App t (Var x) \<Leftrightarrow> v : T\<^isub>2 \<Longrightarrow> (x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> v : T\<^isub>2" 
     by fact
-  ultimately have "(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App u (Var x) : T2" by auto
-  then show "\<Gamma> \<turnstile> s <=> u : T1\<rightarrow>T2" using fs fs2 by auto
+  ultimately have "(x,T\<^isub>1)#\<Gamma> \<turnstile> App s (Var x) \<Leftrightarrow> App u (Var x) : T\<^isub>2" by auto
+  then show "\<Gamma> \<turnstile> s \<Leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" using fs fs2 by auto
 next
-  case (QAP_App \<Gamma> p q T1 T2 s t u)
-  have h1:"\<Gamma> \<turnstile> p \<leftrightarrow> q : T1\<rightarrow>T2" by fact
-  have ih1:"\<And>u. \<Gamma> \<turnstile> q \<leftrightarrow> u : T1\<rightarrow>T2 \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> u : T1\<rightarrow>T2" by fact
-  have "\<Gamma> \<turnstile> s <=> t : T1" by fact
-  have ih2:"\<And>u. \<Gamma> \<turnstile> t <=> u : T1 \<Longrightarrow> \<Gamma> \<turnstile> s <=> u : T1" by fact
-  have "\<Gamma> \<turnstile> App q t \<leftrightarrow> u : T2" by fact
-  then obtain r T1' v where ha:"\<Gamma> \<turnstile> q \<leftrightarrow> r : T1'\<rightarrow>T2" and hb:"\<Gamma> \<turnstile> t <=> v : T1'" and eq:"u = App r v" 
+  case (QAP_App \<Gamma> p q T\<^isub>1 T\<^isub>2 s t u)
+  have h1:"\<Gamma> \<turnstile> p \<leftrightarrow> q : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+  have ih1:"\<And>u. \<Gamma> \<turnstile> q \<leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2 \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> u : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+  have "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T\<^isub>1" by fact
+  have ih2:"\<And>u. \<Gamma> \<turnstile> t \<Leftrightarrow> u : T\<^isub>1 \<Longrightarrow> \<Gamma> \<turnstile> s \<Leftrightarrow> u : T\<^isub>1" by fact
+  have "\<Gamma> \<turnstile> App q t \<leftrightarrow> u : T\<^isub>2" by fact
+  then obtain r T\<^isub>1' v where ha:"\<Gamma> \<turnstile> q \<leftrightarrow> r : T\<^isub>1'\<rightarrow>T\<^isub>2" and hb:"\<Gamma> \<turnstile> t \<Leftrightarrow> v : T\<^isub>1'" and eq:"u = App r v" 
     by auto
-  moreover have "\<Gamma> \<turnstile> q \<leftrightarrow> p : T1\<rightarrow>T2" using h1 algorithmic_path_symmetry by auto
-  ultimately have "T1'\<rightarrow>T2 = T1\<rightarrow>T2" using algorithmic_path_type_unicity by auto
-  then have "T1' = T1" using ty.inject by auto
-  then have "\<Gamma> \<turnstile> s <=> v : T1" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T1\<rightarrow>T2" using ih1 ih2 ha hb by auto
-  then show "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T2" using eq by auto
+  moreover have "\<Gamma> \<turnstile> q \<leftrightarrow> p : T\<^isub>1\<rightarrow>T\<^isub>2" using h1 algorithmic_symmetry by auto
+  ultimately have "T\<^isub>1'\<rightarrow>T\<^isub>2 = T\<^isub>1\<rightarrow>T\<^isub>2" using algorithmic_path_type_unicity by auto
+  then have "T\<^isub>1' = T\<^isub>1" using ty.inject by auto
+  then have "\<Gamma> \<turnstile> s \<Leftrightarrow> v : T\<^isub>1" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T\<^isub>1\<rightarrow>T\<^isub>2" using ih1 ih2 ha hb by auto
+  then show "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T\<^isub>2" using eq by auto
 qed (auto)
 
 lemma algorithmic_weak_head_closure:
-  shows "\<lbrakk>\<Gamma> \<turnstile> s <=> t : T ; s' \<leadsto> s; t' \<leadsto> t\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' <=> t' : T"
+  shows "\<lbrakk>\<Gamma> \<turnstile> s \<Leftrightarrow> t : T ; s' \<leadsto> s; t' \<leadsto> t\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' \<Leftrightarrow> t' : T"
 by (nominal_induct \<Gamma> s t T avoiding: s' t'  
     rule: alg_equiv_alg_path_equiv_strong_inducts(1) [of _ _ _ _ "%a b c d e. True"])
    (auto)
@@ -1090,32 +1230,29 @@
   case TBase
   then show "\<Gamma> \<turnstile> s is u : TBase" by (auto elim:  algorithmic_transitivity)
 next 
-  case (Arrow T1 T2 \<Gamma> s t u)
-  have h1:"\<Gamma> \<turnstile> s is t : T1 \<rightarrow> T2" by fact
-  have h2:"\<Gamma> \<turnstile> t is u : T1 \<rightarrow> T2" by fact
-  have ih1:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T1; \<Gamma> \<turnstile> t is u : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T1" by fact
-  have ih2:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T2; \<Gamma> \<turnstile> t is u : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T2" by fact
+  case (Arrow T\<^isub>1 T\<^isub>2 \<Gamma> s t u)
+  have h1:"\<Gamma> \<turnstile> s is t : T\<^isub>1 \<rightarrow> T\<^isub>2" by fact
+  have h2:"\<Gamma> \<turnstile> t is u : T\<^isub>1 \<rightarrow> T\<^isub>2" by fact
+  have ih1:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>1; \<Gamma> \<turnstile> t is u : T\<^isub>1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^isub>1" by fact
+  have ih2:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; \<Gamma> \<turnstile> t is u : T\<^isub>2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T\<^isub>2" by fact
   {
     fix \<Gamma>' s' u'
-    assume hsub:"\<Gamma>\<lless>\<Gamma>'" and hl:"\<Gamma>' \<turnstile> s' is u' : T1"
-    then have "\<Gamma>' \<turnstile> u' is s' : T1" using logical_symmetry by blast
-    then have "\<Gamma>' \<turnstile> u' is u' : T1" using ih1 hl by blast
-    then have "\<Gamma>' \<turnstile> App t u' is App u u' : T2" using h2 hsub by auto
-    moreover have "\<Gamma>' \<turnstile>  App s s' is App t u' : T2" using h1 hsub hl by auto
-    ultimately have "\<Gamma>' \<turnstile>  App s s' is App u u' : T2" using ih2 by blast
+    assume hsub:"\<Gamma>\<lless>\<Gamma>'" and hl:"\<Gamma>' \<turnstile> s' is u' : T\<^isub>1"
+    then have "\<Gamma>' \<turnstile> u' is s' : T\<^isub>1" using logical_symmetry by blast
+    then have "\<Gamma>' \<turnstile> u' is u' : T\<^isub>1" using ih1 hl by blast
+    then have "\<Gamma>' \<turnstile> App t u' is App u u' : T\<^isub>2" using h2 hsub by auto
+    moreover have "\<Gamma>' \<turnstile>  App s s' is App t u' : T\<^isub>2" using h1 hsub hl by auto
+    ultimately have "\<Gamma>' \<turnstile>  App s s' is App u u' : T\<^isub>2" using ih2 by blast
   }
   moreover have "valid \<Gamma>" using h1 alg_equiv_valid by auto
-  ultimately show "\<Gamma> \<turnstile> s is u : T1 \<rightarrow> T2" by auto
+  ultimately show "\<Gamma> \<turnstile> s is u : T\<^isub>1 \<rightarrow> T\<^isub>2" by auto
 qed (auto)
 
 lemma logical_weak_head_closure:
-  assumes a: "\<Gamma> \<turnstile> s is t : T" "s' \<leadsto> s" "t' \<leadsto> t"
+  assumes "\<Gamma> \<turnstile> s is t : T" "s' \<leadsto> s" "t' \<leadsto> t"
   shows "\<Gamma> \<turnstile> s' is t' : T"
-using a
-apply(nominal_induct arbitrary: \<Gamma> s t s' t' rule:ty.induct)
-apply(auto simp add: algorithmic_weak_head_closure)
-apply(blast)
-done
+using assms algorithmic_weak_head_closure 
+by (nominal_induct arbitrary: \<Gamma> s t s' t' rule:ty.induct) (auto, blast)
 
 lemma logical_weak_head_closure':
   assumes "\<Gamma> \<turnstile> s is t : T" "s' \<leadsto> s" 
@@ -1127,22 +1264,22 @@
 next
   case (TUnit \<Gamma> s t s')
   then show ?case by auto
-next
-  case (Arrow T1 T2 \<Gamma> s t s')
+next 
+  case (Arrow T\<^isub>1 T\<^isub>2 \<Gamma> s t s')
   have h1:"s' \<leadsto> s" by fact
-  have ih:"\<And>\<Gamma> s t s'. \<lbrakk>\<Gamma> \<turnstile> s is t : T2; s' \<leadsto> s\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' is t : T2" by fact
-  have h2:"\<Gamma> \<turnstile> s is t : T1\<rightarrow>T2" by fact
-  then have hb:"\<forall>\<Gamma>' s' t'. \<Gamma>\<lless>\<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T1 \<longrightarrow>  (\<Gamma>' \<turnstile> (App s s') is (App t t') : T2)" by auto
+  have ih:"\<And>\<Gamma> s t s'. \<lbrakk>\<Gamma> \<turnstile> s is t : T\<^isub>2; s' \<leadsto> s\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' is t : T\<^isub>2" by fact
+  have h2:"\<Gamma> \<turnstile> s is t : T\<^isub>1\<rightarrow>T\<^isub>2" by fact
+  then have hb:"\<forall>\<Gamma>' s' t'. \<Gamma>\<lless>\<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T\<^isub>1 \<longrightarrow>  (\<Gamma>' \<turnstile> (App s s') is (App t t') : T\<^isub>2)" by auto
   {
-    fix \<Gamma>' s2 t2
-    assume "\<Gamma>\<lless>\<Gamma>'" and "\<Gamma>' \<turnstile> s2 is t2 : T1"
-    then have "\<Gamma>' \<turnstile> (App s s2) is (App t t2) : T2" using hb by auto
-    moreover have "(App s' s2)  \<leadsto> (App s s2)" using h1 by auto  
-    ultimately have "\<Gamma>' \<turnstile> App s' s2 is App t t2 : T2" using ih by auto
+    fix \<Gamma>' s\<^isub>2 t\<^isub>2
+    assume "\<Gamma>\<lless>\<Gamma>'" and "\<Gamma>' \<turnstile> s\<^isub>2 is t\<^isub>2 : T\<^isub>1"
+    then have "\<Gamma>' \<turnstile> (App s s\<^isub>2) is (App t t\<^isub>2) : T\<^isub>2" using hb by auto
+    moreover have "(App s' s\<^isub>2)  \<leadsto> (App s s\<^isub>2)" using h1 by auto  
+    ultimately have "\<Gamma>' \<turnstile> App s' s\<^isub>2 is App t t\<^isub>2 : T\<^isub>2" using ih by auto
   }
   moreover have "valid \<Gamma>" using h2 log_equiv_valid by auto
-  ultimately show "\<Gamma> \<turnstile> s' is t : T1\<rightarrow>T2" by auto
-qed
+  ultimately show "\<Gamma> \<turnstile> s' is t : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
+qed 
 
 abbreviation 
  log_equiv_subst :: "(name\<times>ty) list \<Rightarrow> (name\<times>trm) list \<Rightarrow>  (name\<times>trm) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool"  
@@ -1200,28 +1337,28 @@
   shows "\<Gamma>' \<turnstile> \<gamma><t> is \<theta><t> : T"   
 using assms
 proof (nominal_induct avoiding:\<Gamma>' \<gamma> \<theta>  rule: typing_induct_strong)
-case (t_Lam x \<Gamma> T1 t2 T2 \<Gamma>' \<gamma> \<theta>)
+case (t_Lam x \<Gamma> T\<^isub>1 t\<^isub>2 T\<^isub>2 \<Gamma>' \<gamma> \<theta>)
 have fs:"x\<sharp>\<gamma>" "x\<sharp>\<theta>" "x\<sharp>\<Gamma>" by fact
 have h:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
-have ih:"\<And> \<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><t2> is \<theta><t2> : T2" by fact
+have ih:"\<And> \<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T\<^isub>1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><t\<^isub>2> is \<theta><t\<^isub>2> : T\<^isub>2" by fact
 {
   fix \<Gamma>'' s' t'
-  assume "\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T1"
+  assume "\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T\<^isub>1"
   then have "\<Gamma>'' \<turnstile> \<gamma> is \<theta> over \<Gamma>" using logical_subst_monotonicity h by blast
-  then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext hl fs by blast
-  then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><t2> is (x,t')#\<theta><t2> : T2" using ih by auto
-  then have "\<Gamma>''\<turnstile>\<gamma><t2>[x::=s'] is \<theta><t2>[x::=t'] : T2" using psubst_subst_psubst fs by simp
-  moreover have "App (Lam [x].\<gamma><t2>) s' \<leadsto> \<gamma><t2>[x::=s']" by auto
-  moreover have "App (Lam [x].\<theta><t2>) t' \<leadsto> \<theta><t2>[x::=t']" by auto
-  ultimately have "\<Gamma>''\<turnstile> App (Lam [x].\<gamma><t2>) s' is App (Lam [x].\<theta><t2>) t' : T2" 
+  then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
+  then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><t\<^isub>2> is (x,t')#\<theta><t\<^isub>2> : T\<^isub>2" using ih by auto
+  then have "\<Gamma>''\<turnstile>\<gamma><t\<^isub>2>[x::=s'] is \<theta><t\<^isub>2>[x::=t'] : T\<^isub>2" using psubst_subst_psubst fs by simp
+  moreover have "App (Lam [x].\<gamma><t\<^isub>2>) s' \<leadsto> \<gamma><t\<^isub>2>[x::=s']" by auto
+  moreover have "App (Lam [x].\<theta><t\<^isub>2>) t' \<leadsto> \<theta><t\<^isub>2>[x::=t']" by auto
+  ultimately have "\<Gamma>''\<turnstile> App (Lam [x].\<gamma><t\<^isub>2>) s' is App (Lam [x].\<theta><t\<^isub>2>) t' : T\<^isub>2" 
     using logical_weak_head_closure by auto
 }
 moreover have "valid \<Gamma>'" using h by auto
-ultimately show "\<Gamma>' \<turnstile> \<gamma><Lam [x].t2> is \<theta><Lam [x].t2> : T1\<rightarrow>T2" using fs by auto 
+ultimately show "\<Gamma>' \<turnstile> \<gamma><Lam [x].t\<^isub>2> is \<theta><Lam [x].t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" using fs by auto 
 qed (auto)
 
 theorem fundamental_theorem_2:
-  assumes h1: "\<Gamma> \<turnstile> s == t : T" 
+  assumes h1: "\<Gamma> \<turnstile> s \<equiv> t : T" 
   and     h2: "\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>"
   shows "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T"
 using h1 h2
@@ -1245,69 +1382,70 @@
   moreover have "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T" using ih1 h by auto
   ultimately show "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><u> : T" using logical_transitivity by blast
 next
-  case (Q_Abs x \<Gamma> T1 s2 t2 T2 \<Gamma>' \<gamma> \<theta>)
+  case (Q_Abs x \<Gamma> T\<^isub>1 s\<^isub>2 t\<^isub>2 T\<^isub>2 \<Gamma>' \<gamma> \<theta>)
   have fs:"x\<sharp>\<Gamma>" by fact
   have fs2: "x\<sharp>\<gamma>" "x\<sharp>\<theta>" by fact 
   have h2:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
-  have ih:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s2> is \<theta><t2> : T2" by fact
+  have ih:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T\<^isub>1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s\<^isub>2> is \<theta><t\<^isub>2> : T\<^isub>2" by fact
   {
     fix \<Gamma>'' s' t'
-    assume "\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T1"
+    assume "\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T\<^isub>1"
     then have "\<Gamma>'' \<turnstile> \<gamma> is \<theta> over \<Gamma>" using h2 logical_subst_monotonicity by blast
-    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext hl fs by blast
-    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><s2> is (x,t')#\<theta><t2> : T2" using ih by blast
-    then have "\<Gamma>''\<turnstile> \<gamma><s2>[x::=s'] is \<theta><t2>[x::=t'] : T2" using fs2 psubst_subst_psubst by auto
-    moreover have "App (Lam [x]. \<gamma><s2>) s' \<leadsto>  \<gamma><s2>[x::=s']" 
-              and "App (Lam [x].\<theta><t2>) t' \<leadsto> \<theta><t2>[x::=t']" by auto
-    ultimately have "\<Gamma>'' \<turnstile> App (Lam [x]. \<gamma><s2>) s' is App (Lam [x].\<theta><t2>) t' : T2" 
+    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
+    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><s\<^isub>2> is (x,t')#\<theta><t\<^isub>2> : T\<^isub>2" using ih by blast
+    then have "\<Gamma>''\<turnstile> \<gamma><s\<^isub>2>[x::=s'] is \<theta><t\<^isub>2>[x::=t'] : T\<^isub>2" using fs2 psubst_subst_psubst by auto
+    moreover have "App (Lam [x]. \<gamma><s\<^isub>2>) s' \<leadsto>  \<gamma><s\<^isub>2>[x::=s']" 
+              and "App (Lam [x].\<theta><t\<^isub>2>) t' \<leadsto> \<theta><t\<^isub>2>[x::=t']" by auto
+    ultimately have "\<Gamma>'' \<turnstile> App (Lam [x]. \<gamma><s\<^isub>2>) s' is App (Lam [x].\<theta><t\<^isub>2>) t' : T\<^isub>2" 
       using logical_weak_head_closure by auto
   }
   moreover have "valid \<Gamma>'" using h2 by auto
-  ultimately have "\<Gamma>' \<turnstile> Lam [x].\<gamma><s2> is Lam [x].\<theta><t2> : T1\<rightarrow>T2" by auto
-  then show "\<Gamma>' \<turnstile> \<gamma><Lam [x].s2> is \<theta><Lam [x].t2> : T1\<rightarrow>T2" using fs2 by auto
+  ultimately have "\<Gamma>' \<turnstile> Lam [x].\<gamma><s\<^isub>2> is Lam [x].\<theta><t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
+  then show "\<Gamma>' \<turnstile> \<gamma><Lam [x].s\<^isub>2> is \<theta><Lam [x].t\<^isub>2> : T\<^isub>1\<rightarrow>T\<^isub>2" using fs2 by auto
 next
-  case (Q_App \<Gamma> s1 t1 T1 T2 s2 t2 \<Gamma>' \<gamma> \<theta>)
-  then show "\<Gamma>' \<turnstile> \<gamma><App s1 s2> is \<theta><App t1 t2> : T2" by auto 
+  case (Q_App \<Gamma> s\<^isub>1 t\<^isub>1 T\<^isub>1 T\<^isub>2 s\<^isub>2 t\<^isub>2 \<Gamma>' \<gamma> \<theta>)
+  then show "\<Gamma>' \<turnstile> \<gamma><App s\<^isub>1 s\<^isub>2> is \<theta><App t\<^isub>1 t\<^isub>2> : T\<^isub>2" by auto 
 next
-  case (Q_Beta x \<Gamma> T1 s12 t12 T2 s2 t2 \<Gamma>' \<gamma> \<theta>)
+  case (Q_Beta x \<Gamma> T\<^isub>1 s12 t12 T\<^isub>2 s\<^isub>2 t\<^isub>2 \<Gamma>' \<gamma> \<theta>)
   have h:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
   have fs:"x\<sharp>\<Gamma>" by fact
   have fs2:"x\<sharp>\<gamma>" "x\<sharp>\<theta>" by fact 
-  have ih1:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s2> is \<theta><t2> : T1" by fact
-  have ih2:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s12> is \<theta><t12> : T2" by fact
-  have "\<Gamma>' \<turnstile> \<gamma><s2> is \<theta><t2> : T1" using ih1 h by auto
-  then have "\<Gamma>' \<turnstile> (x,\<gamma><s2>)#\<gamma> is (x,\<theta><t2>)#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext h fs by blast
-  then have "\<Gamma>' \<turnstile> (x,\<gamma><s2>)#\<gamma><s12> is (x,\<theta><t2>)#\<theta><t12> : T2" using ih2 by auto
-  then have "\<Gamma>' \<turnstile> \<gamma><s12>[x::=\<gamma><s2>] is \<theta><t12>[x::=\<theta><t2>] : T2" using fs2 psubst_subst_psubst by auto
-  then have "\<Gamma>' \<turnstile> \<gamma><s12>[x::=\<gamma><s2>] is \<theta><t12[x::=t2]> : T2" using fs2 psubst_subst_propagate by auto
-  moreover have "App (Lam [x].\<gamma><s12>) (\<gamma><s2>) \<leadsto> \<gamma><s12>[x::=\<gamma><s2>]" by auto 
-  ultimately have "\<Gamma>' \<turnstile> App (Lam [x].\<gamma><s12>) (\<gamma><s2>) is \<theta><t12[x::=t2]> : T2" 
+  have ih1:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s\<^isub>2> is \<theta><t\<^isub>2> : T\<^isub>1" by fact
+  have ih2:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T\<^isub>1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s12> is \<theta><t12> : T\<^isub>2" by fact
+  have "\<Gamma>' \<turnstile> \<gamma><s\<^isub>2> is \<theta><t\<^isub>2> : T\<^isub>1" using ih1 h by auto
+  then have "\<Gamma>' \<turnstile> (x,\<gamma><s\<^isub>2>)#\<gamma> is (x,\<theta><t\<^isub>2>)#\<theta> over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext h fs by blast
+  then have "\<Gamma>' \<turnstile> (x,\<gamma><s\<^isub>2>)#\<gamma><s12> is (x,\<theta><t\<^isub>2>)#\<theta><t12> : T\<^isub>2" using ih2 by auto
+  then have "\<Gamma>' \<turnstile> \<gamma><s12>[x::=\<gamma><s\<^isub>2>] is \<theta><t12>[x::=\<theta><t\<^isub>2>] : T\<^isub>2" using fs2 psubst_subst_psubst by auto
+  then have "\<Gamma>' \<turnstile> \<gamma><s12>[x::=\<gamma><s\<^isub>2>] is \<theta><t12[x::=t\<^isub>2]> : T\<^isub>2" using fs2 psubst_subst_propagate by auto
+  moreover have "App (Lam [x].\<gamma><s12>) (\<gamma><s\<^isub>2>) \<leadsto> \<gamma><s12>[x::=\<gamma><s\<^isub>2>]" by auto 
+  ultimately have "\<Gamma>' \<turnstile> App (Lam [x].\<gamma><s12>) (\<gamma><s\<^isub>2>) is \<theta><t12[x::=t\<^isub>2]> : T\<^isub>2" 
     using logical_weak_head_closure' by auto
-  then show "\<Gamma>' \<turnstile> \<gamma><App (Lam [x].s12) s2> is \<theta><t12[x::=t2]> : T2" using fs2 by simp
+  then show "\<Gamma>' \<turnstile> \<gamma><App (Lam [x].s12) s\<^isub>2> is \<theta><t12[x::=t\<^isub>2]> : T\<^isub>2" using fs2 by simp
 next
-  case (Q_Ext x \<Gamma> s t T1 T2 \<Gamma>' \<gamma> \<theta>)
+  case (Q_Ext x \<Gamma> s t T\<^isub>1 T\<^isub>2 \<Gamma>' \<gamma> \<theta>)
   have h2:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
   have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
-  have ih:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><App s (Var x)> is \<theta><App t (Var x)> : T2" 
+  have ih:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T\<^isub>1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><App s (Var x)> is \<theta><App t (Var x)> : T\<^isub>2" 
     by fact
    {
     fix \<Gamma>'' s' t'
-    assume hsub:"\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T1"
+    assume hsub:"\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T\<^isub>1"
     then have "\<Gamma>'' \<turnstile> \<gamma> is \<theta> over \<Gamma>" using h2 logical_subst_monotonicity by blast
-    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext hl fs by blast
-    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><App s (Var x)>  is (x,t')#\<theta><App t (Var x)> : T2" using ih by blast
-    then have "\<Gamma>'' \<turnstile> App (((x,s')#\<gamma>)<s>) (((x,s')#\<gamma>)<(Var x)>) is App ((x,t')#\<theta><t>) ((x,t')#\<theta><(Var x)>) : T2"
+    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T\<^isub>1)#\<Gamma>" using equiv_subst_ext hl fs by blast
+    then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><App s (Var x)>  is (x,t')#\<theta><App t (Var x)> : T\<^isub>2" using ih by blast
+    then have "\<Gamma>'' \<turnstile> App (((x,s')#\<gamma>)<s>) (((x,s')#\<gamma>)<(Var x)>) is App ((x,t')#\<theta><t>) ((x,t')#\<theta><(Var x)>) : T\<^isub>2"
       by auto
-    then have "\<Gamma>'' \<turnstile> App ((x,s')#\<gamma><s>) s'  is App ((x,t')#\<theta><t>) t' : T2" by auto
-    then have "\<Gamma>'' \<turnstile> App (\<gamma><s>) s' is App (\<theta><t>) t' : T2" using fs fresh_psubst_simpl by auto
+    then have "\<Gamma>'' \<turnstile> App ((x,s')#\<gamma><s>) s'  is App ((x,t')#\<theta><t>) t' : T\<^isub>2" by auto
+    then have "\<Gamma>'' \<turnstile> App (\<gamma><s>) s' is App (\<theta><t>) t' : T\<^isub>2" using fs fresh_psubst_simpl by auto
   }
   moreover have "valid \<Gamma>'" using h2 by auto
-  ultimately show "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T1\<rightarrow>T2" by auto
+  ultimately show "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T\<^isub>1\<rightarrow>T\<^isub>2" by auto
 qed
 
+
 theorem completeness:
-  assumes "\<Gamma> \<turnstile> s == t : T"
-  shows "\<Gamma> \<turnstile> s <=> t : T"
+  assumes "\<Gamma> \<turnstile> s \<equiv> t : T"
+  shows   "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T"
 using assms
 proof -
   {
@@ -1319,16 +1457,17 @@
   ultimately have "\<Gamma> \<turnstile> [] is [] over \<Gamma>" by auto 
   then have "\<Gamma> \<turnstile> []<s> is []<t> : T" using fundamental_theorem_2 assms by blast
   then have "\<Gamma> \<turnstile> s is t : T" by simp
-  then show  "\<Gamma> \<turnstile> s <=> t : T" using main_lemma by simp
+  then show  "\<Gamma> \<turnstile> s \<Leftrightarrow> t : T" using main_lemma by simp
 qed
 
-(* Soundness is left as an exercise - like in the book - for the avid formalist 
+text{*
+\section{About soundness}
+*}
 
-theorem soundness:
-  shows "\<lbrakk>\<Gamma> \<turnstile> s <=> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s == t : T"
-  and   "\<lbrakk>\<Gamma> \<turnstile> s \<leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s == t : T"
-
-*)
+text {* We leave soundness as an exercise - like in the book :-) \\ 
+ @{prop[mode=IfThen] "\<lbrakk>\<Gamma> \<turnstile> s \<Leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"} \\
+ @{prop "\<lbrakk>\<Gamma> \<turnstile> s \<leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<equiv> t : T"} 
+*}
 
 end
 
--- a/src/HOL/Nominal/Examples/Fsub.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Fsub.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -90,7 +90,7 @@
 lemma domain_eqvt:
   fixes pi::"tyvrs prm"
   shows "pi\<bullet>(domain \<Gamma>) = domain (pi\<bullet>\<Gamma>)" 
-  by (induct \<Gamma>, simp_all add: perm_empty perm_insert perm_fst)
+  by (induct \<Gamma>) (simp_all add: eqvt)
 
 lemma finite_domain:
   shows "finite (domain \<Gamma>)"
@@ -239,17 +239,10 @@
   "size_ty (Top) = 1"
   "size_ty (T1 \<rightarrow> T2) = (size_ty T1) + (size_ty T2) + 1"
   "X\<sharp>T1 \<Longrightarrow> size_ty (\<forall>[X<:T1].T2) = (size_ty T1) + (size_ty T2) + 1"
-  apply (finite_guess add: perm_nat_def)+
+  apply (finite_guess)+
   apply (rule TrueI)+
-  apply (simp add: fresh_def supp_nat)
-  apply (assumption | fresh_guess add: perm_nat_def)+
-  done
-
-lemma eq_eqvt:
-  fixes pi::"tyvrs prm"
-  and   x::"'a::pt_tyvrs"
-  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
-  apply(simp add: perm_bool perm_bij)
+  apply (simp add: fresh_nat)
+  apply (fresh_guess)+
   done
 
 consts subst_ty :: "tyvrs \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> ty"
@@ -265,20 +258,10 @@
   "(Top)[Y:=T]\<^isub>t\<^isub>y = Top"
   "(T\<^isub>1 \<rightarrow> T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (T\<^isub>1[Y:=T]\<^isub>t\<^isub>y) \<rightarrow> (T\<^isub>2[Y:=T]\<^isub>t\<^isub>y)"
   "\<lbrakk>X\<sharp>(Y,T); X\<sharp>T\<^isub>1\<rbrakk> \<Longrightarrow> (\<forall>[X<:T\<^isub>1].T\<^isub>2)[Y:=T]\<^isub>t\<^isub>y = (\<forall>[X<:(T\<^isub>1[Y:=T]\<^isub>t\<^isub>y)].(T\<^isub>2[Y:=T]\<^isub>t\<^isub>y))"
-  apply (finite_guess add: perm_if eq_eqvt fs_tyvrs1)
-  apply finite_guess
-  apply finite_guess
-  apply finite_guess
-  apply perm_full_simp
-  apply (simp add: supp_unit)
+  apply (finite_guess)+
   apply (rule TrueI)+
   apply (simp add: abs_fresh)
-  apply (fresh_guess add: fs_tyvrs1 perm_if eq_eqvt)
-  apply fresh_guess
-  apply fresh_guess
-  apply fresh_guess
-  apply perm_full_simp
-  apply assumption
+  apply (fresh_guess)+
   done
 
 consts subst_tyc :: "ty_context \<Rightarrow> tyvrs \<Rightarrow> ty \<Rightarrow> ty_context" ("_[_:=_]\<^isub>t\<^isub>y\<^isub>c" [100,100,100] 100)
@@ -357,11 +340,11 @@
   fixes pi::"tyvrs prm"
   shows "\<Gamma> \<turnstile> S <: T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>S) <: (pi\<bullet>T)"
 apply(erule subtype_of.induct)
-apply(force intro: valid_eqvt closed_in_eqvt)
+apply(force intro: valid_eqvt closed_in_eqvt eqvt)
 apply(force intro!: S_Var
             intro: closed_in_eqvt valid_eqvt 
             dest: pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst]
-            simp add: pt_list_set_pi[OF pt_tyvrs_inst, symmetric])
+            simp add: set_eqvt)
 apply(force intro: valid_eqvt
             dest: pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst]
             simp add: domain_eqvt)
@@ -400,7 +383,7 @@
     case (S_Var S T X \<Gamma>)
     have "(X,S) \<in> set \<Gamma>" by fact
     hence "pi\<bullet>(X,S)\<in>pi\<bullet>(set \<Gamma>)" by (rule pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
-    hence "(pi\<bullet>X,pi\<bullet>S)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_tyvrs_inst])
+    hence "(pi\<bullet>X,pi\<bullet>S)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: set_eqvt)
     moreover
     have "\<Gamma> \<turnstile> S <: T" by fact
     hence "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>S) <: (pi\<bullet>T)" by (rule subtype_eqvt)
@@ -416,7 +399,7 @@
     moreover
     have "X \<in> domain \<Gamma>" by fact
     hence "(pi\<bullet>X)\<in>pi\<bullet>domain \<Gamma>" by (rule pt_set_bij2[OF pt_tyvrs_inst, OF at_tyvrs_inst])
-    hence "(pi\<bullet>X)\<in>domain (pi\<bullet>\<Gamma>)" by (simp add: domain_eqvt pt_list_set_pi[OF pt_tyvrs_inst])
+    hence "(pi\<bullet>X)\<in>domain (pi\<bullet>\<Gamma>)" by (simp add: domain_eqvt set_eqvt)
     ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Tvar X)) (pi\<bullet>(Tvar X))" by (simp add: a3)
   next
     case S_Arrow thus ?case by (auto intro: a4 subtype_eqvt)
--- a/src/HOL/Nominal/Examples/Height.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Height.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -1,9 +1,9 @@
 (* $Id$ *)
 
-(*  Simple, but artificial, problem suggested by D. Wang *)
+(*  Simple problem suggested by D. Wang *)
 
 theory Height
-imports "Nominal"
+imports "../Nominal"
 begin
 
 atom_decl name
@@ -12,7 +12,7 @@
                      | App "lam" "lam"
                      | Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100)
 
-text {* definition of the height-function *} 
+text {* definition of the height-function on lambda-terms *} 
 
 consts 
   height :: "lam \<Rightarrow> int"
@@ -29,13 +29,6 @@
 
 text {* definition of capture-avoiding substitution *}
 
-lemma eq_eqvt:
-  fixes pi::"name prm"
-  and   x::"'a::pt_name"
-  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
-  apply(simp add: perm_bool perm_bij)
-  done
-
 consts
   subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [100,100,100] 100)
 
@@ -43,21 +36,21 @@
   "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
   "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
   "\<lbrakk>x\<sharp>y; x\<sharp>t'\<rbrakk> \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
-apply(finite_guess add: eq_eqvt perm_if fs_name1)+
+apply(finite_guess)+
 apply(rule TrueI)+
 apply(simp add: abs_fresh)
-apply(fresh_guess add: eq_eqvt perm_if fs_name1)+
+apply(fresh_guess)+
 done
 
 text{* the next lemma is needed in the Var-case of the theorem *}
 
 lemma height_ge_one: 
   shows "1 \<le> (height e)"
-  by (nominal_induct e rule: lam.induct) 
-     (simp | arith)+
+by (nominal_induct e rule: lam.induct) 
+   (simp | arith)+
 
 text {* unlike the proplem suggested by Wang, however, the 
-        theorem is here formulated here entirely by using 
+        theorem is here formulated  entirely by using 
         functions *}
 
 theorem height_subst:
--- a/src/HOL/Nominal/Examples/Lam_Funs.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Lam_Funs.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -1,7 +1,7 @@
 (* $Id$ *)
 
 theory Lam_Funs
-imports Nominal
+imports "../Nominal"
 begin
 
 atom_decl name
@@ -16,13 +16,13 @@
   depth :: "lam \<Rightarrow> nat"
 
 nominal_primrec
-  "depth (Var x) = 1"
+  "depth (Var x) = (1::nat)"
   "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1"
-  "depth (Lam [a].t) = (depth t) + 1"
-  apply(finite_guess add: perm_nat_def)+
+  "depth (Lam [a].t) = (depth t) + (1::nat)"
+  apply(finite_guess)+
   apply(rule TrueI)+
   apply(simp add: fresh_nat)
-  apply(fresh_guess add: perm_nat_def)+
+  apply(fresh_guess)+
   done
 
 text {* free variables of a lambda-term *}
@@ -34,20 +34,13 @@
   "frees (Var a) = {a}"
   "frees (App t1 t2) = (frees t1) \<union> (frees t2)"
   "frees (Lam [a].t) = (frees t) - {a}"
-apply(finite_guess add: perm_insert perm_set_def)
-apply(finite_guess add: perm_union)
-apply(finite_guess add: pt_set_diff_eqvt[OF pt_name_inst, OF at_name_inst] perm_insert)
-apply(simp add: perm_set_def)
-apply(simp add: fs_name1)
-apply(simp)+
+apply(finite_guess)+
+apply(simp)+ 
 apply(simp add: fresh_def)
 apply(simp add: supp_of_fin_sets[OF pt_name_inst, OF at_name_inst, OF fs_at_inst[OF at_name_inst]])
 apply(simp add: supp_atm)
 apply(force)
-apply(fresh_guess add: perm_insert perm_set_def)
-apply(fresh_guess add: perm_union)
-apply(fresh_guess add: pt_set_diff_eqvt[OF pt_name_inst, OF at_name_inst] perm_insert)
-apply(simp add: perm_set_def)
+apply(fresh_guess)+
 done
 
 lemma frees_equals_support:
@@ -57,12 +50,6 @@
 
 text {* capture-avoiding substitution *}
 
-lemma eq_eqvt:
-  fixes pi::"name prm"
-  and   x::"'a::pt_name"
-  shows "pi\<bullet>(x=y) = ((pi\<bullet>x)=(pi\<bullet>y))"
-  apply(simp add: perm_bool perm_bij)
-  done
 
 consts
   subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [100,100,100] 100)
@@ -71,10 +58,10 @@
   "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
   "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
   "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
-apply(finite_guess add: eq_eqvt perm_if fs_name1)+
+apply(finite_guess)+
 apply(rule TrueI)+
 apply(simp add: abs_fresh)
-apply(fresh_guess add: eq_eqvt perm_if fs_name1)+
+apply(fresh_guess)+
 done
 
 lemma subst_eqvt[simp]:
@@ -97,43 +84,43 @@
 text{* parallel substitution *}
 
 consts
-  "domain" :: "(name\<times>lam) list \<Rightarrow> (name list)"
-primrec
-  "domain []    = []"
-  "domain (x#\<theta>) = (fst x)#(domain \<theta>)" 
+ psubst :: "(name\<times>lam) list \<Rightarrow> lam \<Rightarrow> lam" ("_<_>" [100,100] 900)
 
-consts
-  "apply_sss"  :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam" (" _ < _ >" [80,80] 80)
-primrec  
-"(x#\<theta>)<a>  = (if (a = fst x) then (snd x) else \<theta><a>)" 
+fun
+  lookup :: "(name\<times>lam) list \<Rightarrow> name \<Rightarrow> lam"   
+where
+  "lookup [] x        = Var x"
+  "lookup ((y,T)#\<theta>) x = (if x=y then T else lookup \<theta> x)"
 
-lemma apply_sss_eqvt:
+lemma lookup_eqvt[eqvt]:
   fixes pi::"name prm"
-  assumes a: "a\<in>set (domain \<theta>)"
-  shows "pi\<bullet>(\<theta><a>) = (pi\<bullet>\<theta>)<pi\<bullet>a>"
-using a
-by (induct \<theta>)
-   (auto simp add: perm_bij)
+  and   \<theta>::"(name\<times>lam) list"
+  and   x::"name"
+  shows "pi\<bullet>(lookup \<theta> x) = lookup (pi\<bullet>\<theta>) (pi\<bullet>x)"
+by (induct \<theta>) (auto simp add: perm_bij)
 
-lemma domain_eqvt:
-  fixes pi::"name prm"
-  shows "((pi\<bullet>a)\<in>set (domain \<theta>)) = (a\<in>set (domain ((rev pi)\<bullet>\<theta>)))"
-apply(induct \<theta>)
-apply(auto)
-apply(perm_simp)+
-done
+lemma lookup_fresh:
+  fixes z::"name"
+  assumes "z\<sharp>\<theta>" "z\<sharp>x"
+  shows "z\<sharp> lookup \<theta> x"
+using assms 
+by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)
 
-consts
- psubst :: "lam \<Rightarrow> (name\<times>lam) list \<Rightarrow> lam" ("_[<_>]" [100,100] 900)
+lemma lookup_fresh':
+  assumes "z\<sharp>\<theta>"
+  shows "lookup \<theta> z = Var z"
+using assms 
+by (induct rule: lookup.induct)
+   (auto simp add: fresh_list_cons fresh_prod fresh_atm)
 
 nominal_primrec
-  "(Var x)[<\<theta>>] = (if x\<in>set (domain \<theta>) then \<theta><x> else (Var x))"
-  "(App t1 t2)[<\<theta>>] = App (t1[<\<theta>>]) (t2[<\<theta>>])"
-  "x\<sharp>\<theta>\<Longrightarrow>(Lam [x].t)[<\<theta>>] = Lam [x].(t[<\<theta>>])"
-apply(finite_guess add: domain_eqvt apply_sss_eqvt fs_name1)+
+  "\<theta><(Var x)> = (lookup \<theta> x)"
+  "\<theta><(App t1 t2)> = App (\<theta><t1>) (\<theta><t2>)"
+  "x\<sharp>\<theta>\<Longrightarrow>\<theta><(Lam [x].t)> = Lam [x].(\<theta><t>)"
+apply(finite_guess)+
 apply(rule TrueI)+
 apply(simp add: abs_fresh)
-apply(fresh_guess add: domain_eqvt apply_sss_eqvt fs_name1)+
+apply(fresh_guess)+
 done
 
 end
--- a/src/HOL/Nominal/Examples/SN.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/SN.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -8,19 +8,19 @@
 
 section {* Beta Reduction *}
 
-lemma subst_rename[rule_format]: 
-  shows "c\<sharp>t1 \<longrightarrow> (t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2])"
-apply(nominal_induct t1 avoiding: a c t2 rule: lam.induct)
-apply(auto simp add: calc_atm fresh_atm abs_fresh)
-done
+lemma subst_rename: 
+  assumes a: "c\<sharp>t1"
+  shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]"
+using a
+by (nominal_induct t1 avoiding: a c t2 rule: lam.induct)
+   (auto simp add: calc_atm fresh_atm abs_fresh)
 
 lemma forget: 
   assumes a: "a\<sharp>t1"
   shows "t1[a::=t2] = t1"
   using a
-apply (nominal_induct t1 avoiding: a t2 rule: lam.induct)
-apply(auto simp add: abs_fresh fresh_atm)
-done
+by (nominal_induct t1 avoiding: a t2 rule: lam.induct)
+   (auto simp add: abs_fresh fresh_atm)
 
 lemma fresh_fact: 
   fixes a::"name"
@@ -28,9 +28,8 @@
   and     b: "a\<sharp>t2"
   shows "a\<sharp>(t1[b::=t2])"
 using a b
-apply(nominal_induct t1 avoiding: a b t2 rule: lam.induct)
-apply(auto simp add: abs_fresh fresh_atm)
-done
+by (nominal_induct t1 avoiding: a b t2 rule: lam.induct)
+   (auto simp add: abs_fresh fresh_atm)
 
 lemma subst_lemma:  
   assumes a: "x\<noteq>y"
@@ -40,10 +39,17 @@
 by (nominal_induct M avoiding: x y N L rule: lam.induct)
    (auto simp add: fresh_fact forget)
 
-lemma id_subs: "t[x::=Var x] = t"
-apply(nominal_induct t avoiding: x rule: lam.induct)
-apply(simp_all add: fresh_atm)
-done
+lemma id_subs: 
+  shows "t[x::=Var x] = t"
+  by (nominal_induct t avoiding: x rule: lam.induct)
+     (simp_all add: fresh_atm)
+
+lemma subst_eqvt[eqvt]:
+  fixes pi::"name prm" 
+  and   t::"lam"
+  shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]"
+by (nominal_induct t avoiding: x t' rule: lam.induct)
+   (perm_simp add: fresh_bij)+
 
 inductive2 Beta :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta> _" [80,80] 80)
 where
@@ -54,15 +60,9 @@
 
 abbreviation "Beta_star"  :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^isub>\<beta>\<^sup>* _" [80,80] 80) where
   "t1 \<longrightarrow>\<^isub>\<beta>\<^sup>* t2 \<equiv> Beta\<^sup>*\<^sup>* t1 t2"
+
+nominal_inductive Beta
  
-lemma eqvt_beta: 
-  fixes pi :: "name prm"
-  and   t  :: "lam"
-  and   s  :: "lam"
-  assumes a: "t\<longrightarrow>\<^isub>\<beta>s"
-  shows "(pi\<bullet>t)\<longrightarrow>\<^isub>\<beta>(pi\<bullet>s)"
-  using a by (induct, auto)
-
 lemma beta_induct[consumes 1, case_names b1 b2 b3 b4]:
   fixes  P :: "'a::fs_name\<Rightarrow>lam \<Rightarrow> lam \<Rightarrow>bool"
   and    t :: "lam"
@@ -77,9 +77,9 @@
 proof -
   from a have "\<And>(pi::name prm) x. P x (pi\<bullet>t) (pi\<bullet>s)"
   proof (induct)
-    case b1 thus ?case using a1 by (simp, blast intro: eqvt_beta)
+    case b1 thus ?case using a1 by (simp, blast intro: eqvt)
   next
-    case b2 thus ?case using a2 by (simp, blast intro: eqvt_beta)
+    case b2 thus ?case using a2 by (simp, blast intro: eqvt)
   next
     case (b3 s1 s2 a)
     have j1: "s1 \<longrightarrow>\<^isub>\<beta> s2" by fact
@@ -92,7 +92,7 @@
 	where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>x" and f3: "c\<sharp>(pi\<bullet>s1)" and f4: "c\<sharp>(pi\<bullet>s2)"
 	by (force simp add: fresh_prod fresh_atm)
       have x: "P x (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>s1)) (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>s2))"
-	using a3 f2 j1 j2 by (simp, blast intro: eqvt_beta)
+	using a3 f2 j1 j2 by (simp, blast intro: eqvt)
       have alpha1: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>s1))) = (Lam [(pi\<bullet>a)].(pi\<bullet>s1))" using f1 f3
 	by (simp add: lam.inject alpha)
       have alpha2: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>s2))) = (Lam [(pi\<bullet>a)].(pi\<bullet>s2))" using f1 f3
@@ -110,7 +110,7 @@
 	where f1: "c\<noteq>(pi\<bullet>a)" and f2: "c\<sharp>x" and f3: "c\<sharp>(pi\<bullet>s1)" and f4: "c\<sharp>(pi\<bullet>s2)"
 	by (force simp add: fresh_prod fresh_atm)
       have x: "P x (App (Lam [c].(([(c,pi\<bullet>a)]@pi)\<bullet>s1)) (pi\<bullet>s2)) ((([(c,pi\<bullet>a)]@pi)\<bullet>s1)[c::=(pi\<bullet>s2)])"
-	using a4 f2 by (blast intro!: eqvt_beta)
+	using a4 f2 by (blast intro!: eqvt)
       have alpha1: "(Lam [c].([(c,pi\<bullet>a)]\<bullet>(pi\<bullet>s1))) = (Lam [(pi\<bullet>a)].(pi\<bullet>s1))" using f1 f3
 	by (simp add: lam.inject alpha)
       have alpha2: "(([(c,pi\<bullet>a)]@pi)\<bullet>s1)[c::=(pi\<bullet>s2)] = (pi\<bullet>s1)[(pi\<bullet>a)::=(pi\<bullet>s2)]"
@@ -145,7 +145,7 @@
 apply(rule conjI)
 apply(simp add: pt_fresh_left[OF pt_name_inst, OF at_name_inst] calc_atm)
 apply(force dest!: supp_beta simp add: fresh_def)
-apply(force intro!: eqvt_beta)
+apply(force intro!: eqvt)
 done
 
 lemma beta_subst: 
@@ -158,19 +158,16 @@
 
 section {* types *}
 
-datatype ty =
-    TVar "string"
+nominal_datatype ty =
+    TVar "nat"
   | TArr "ty" "ty" (infix "\<rightarrow>" 200)
 
-primrec (unchecked)
- "pi\<bullet>(TVar s) = TVar s"
- "pi\<bullet>(\<tau> \<rightarrow> \<sigma>) = (\<tau> \<rightarrow> \<sigma>)"
-
 lemma perm_ty[simp]:
   fixes pi ::"name prm"
   and   \<tau>  ::"ty"
   shows "pi\<bullet>\<tau> = \<tau>"
-  by (cases \<tau>, simp_all)
+by (nominal_induct \<tau> rule: ty.induct) 
+   (simp_all add: perm_nat_def)
 
 lemma fresh_ty:
   fixes a ::"name"
@@ -178,16 +175,6 @@
   shows "a\<sharp>\<tau>"
   by (simp add: fresh_def supp_def)
 
-instance ty :: pt_name
-apply(intro_classes)   
-apply(simp_all)
-done
-
-instance ty :: fs_name
-apply(intro_classes)
-apply(simp add: supp_def)
-done
-
 (* valid contexts *)
 
 consts
@@ -201,18 +188,11 @@
   v1[intro]: "valid []"
 | v2[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk>\<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"
 
-lemma valid_eqvt:
-  fixes   pi:: "name prm"
-  assumes a: "valid \<Gamma>"
-  shows   "valid (pi\<bullet>\<Gamma>)"
-using a
-apply(induct)
-apply(auto simp add: fresh_bij)
-done
+nominal_inductive valid
 
 (* typing judgements *)
 
-lemma fresh_context[rule_format]: 
+lemma fresh_context: 
   fixes  \<Gamma> :: "(name\<times>ty)list"
   and    a :: "name"
   assumes a: "a\<sharp>\<Gamma>"
@@ -222,24 +202,15 @@
 apply(auto simp add: fresh_prod fresh_list_cons fresh_atm)
 done
 
-lemma valid_elim: 
-  fixes  \<Gamma> :: "(name\<times>ty)list"
-  and    pi:: "name prm"
-  and    a :: "name"
-  and    \<tau> :: "ty"
-  shows "valid ((a,\<tau>)#\<Gamma>) \<Longrightarrow> valid \<Gamma> \<and> a\<sharp>\<Gamma>"
-apply(ind_cases2 "valid ((a,\<tau>)#\<Gamma>)", simp)
-done
+inductive_cases2 valid_elim[elim]: "valid ((a,\<tau>)#\<Gamma>)"
 
-lemma valid_unicity[rule_format]: 
+lemma valid_unicity: 
   assumes a: "valid \<Gamma>"
   and     b: "(c,\<sigma>)\<in>set \<Gamma>"
   and     c: "(c,\<tau>)\<in>set \<Gamma>"
   shows "\<sigma>=\<tau>" 
 using a b c
-apply(induct \<Gamma>)
-apply(auto dest!: valid_elim fresh_context)
-done
+by (induct \<Gamma>) (auto dest!: fresh_context)
 
 inductive2 typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80)
 where
@@ -261,7 +232,7 @@
   moreover
   have "(pi\<bullet>(a,\<tau>))\<in>((pi::name prm)\<bullet>set \<Gamma>)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst])
   ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> ((pi::name prm)\<bullet>Var a) : \<tau>"
-    using typing.t1 by (force simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
+    using typing.t1 by (force simp add: set_eqvt[symmetric])
 next 
   case (t3 a \<Gamma> \<tau> t \<sigma>)
   moreover have "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)" by (simp add: fresh_bij)
@@ -290,7 +261,7 @@
     have j2: "(a,\<tau>)\<in>set \<Gamma>" by fact
     from j1 have j3: "valid (pi\<bullet>\<Gamma>)" by (rule valid_eqvt)
     from j2 have "pi\<bullet>(a,\<tau>)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
-    hence j4: "(pi\<bullet>a,\<tau>)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_name_inst])
+    hence j4: "(pi\<bullet>a,\<tau>)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: eqvt)
     show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) \<tau>" using a1 j3 j4 by simp
   next
     case (t2 \<Gamma> t1 \<tau> \<sigma> t2)
@@ -346,16 +317,14 @@
   assumes a: "(a,\<tau>)\<in>set \<Gamma>"
   shows "a\<in>set(dom_ty \<Gamma>)"
 using a
-apply(induct \<Gamma>, auto)
-done
+by (induct \<Gamma>) (auto)
 
 lemma free_vars: 
   assumes a: "\<Gamma> \<turnstile> t : \<tau>"
   shows " (supp t)\<subseteq>set(dom_ty \<Gamma>)"
 using a
-apply(nominal_induct \<Gamma> t \<tau> rule: typing_induct)
-apply(auto simp add: lam.supp abs_supp supp_atm in_ctxt)
-done
+by (nominal_induct \<Gamma> t \<tau> rule: typing_induct)
+   (auto simp add: lam.supp abs_supp supp_atm in_ctxt)
 
 lemma t1_elim: "\<Gamma> \<turnstile> Var a : \<tau> \<Longrightarrow> valid \<Gamma> \<and> (a,\<tau>) \<in> set \<Gamma>"
 apply(ind_cases2 "\<Gamma> \<turnstile> Var a : \<tau>")
@@ -381,7 +350,7 @@
 lemma typing_valid: 
   assumes a: "\<Gamma> \<turnstile> t : \<tau>" 
   shows "valid \<Gamma>"
-using a by (induct, auto dest!: valid_elim)
+using a by (induct, auto)
 
 lemma ty_subs:
   assumes a: "((c,\<sigma>)#\<Gamma>) \<turnstile> t1:\<tau>"
@@ -434,7 +403,7 @@
   then obtain \<tau>1 \<tau>2 where c11: "\<tau>=\<tau>1\<rightarrow>\<tau>2" and c12: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<turnstile> s : \<tau>2" by force
   from c12 have "valid ((a,\<tau>1)#(c,\<sigma>)#\<Gamma>)" by (rule typing_valid)
   hence ca: "valid \<Gamma>" and cb: "a\<sharp>\<Gamma>" and cc: "c\<sharp>\<Gamma>" 
-    by (auto dest: valid_elim simp add: fresh_list_cons) 
+    by (auto simp add: fresh_list_cons) 
   from c12 have c14: "((c,\<sigma>)#(a,\<tau>1)#\<Gamma>) \<turnstile> s : \<tau>2"
   proof -
     have c2: "((a,\<tau>1)#(c,\<sigma>)#\<Gamma>) \<lless> ((c,\<sigma>)#(a,\<tau>1)#\<Gamma>)" by force
@@ -486,7 +455,7 @@
   have "\<Gamma> \<turnstile> App (Lam [a].s1) s2 : \<tau>" by fact
   hence "\<exists>\<sigma>. (\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau> \<and> \<Gamma> \<turnstile> s2 : \<sigma>)" by (simp add: t2_elim)
   then obtain \<sigma> where a1: "\<Gamma> \<turnstile> (Lam [a].s1) : \<sigma>\<rightarrow>\<tau>" and a2: "\<Gamma> \<turnstile> s2 : \<sigma>" by blast
-  from a1 have "((a,\<sigma>)#\<Gamma>) \<turnstile> s1 : \<tau>" using fr by (blast dest!: t3_elim)
+  from a1 have "((a,\<sigma>)#\<Gamma>) \<turnstile> s1 : \<tau>" using fr by (auto dest!: t3_elim simp add: ty.inject) 
   with a2 show "\<Gamma> \<turnstile> s1[a::=s2] : \<tau>" by (simp add: ty_subs)
 qed
 
@@ -496,7 +465,7 @@
   shows "\<Gamma> \<turnstile> t2:\<tau>"
 using a b
 apply(nominal_induct t1 t2 avoiding: \<Gamma> \<tau> rule: beta_induct)
-apply(auto dest!: t2_elim t3_elim intro: ty_subs)
+apply(auto dest!: t2_elim t3_elim intro: ty_subs simp add: ty.inject)
 done
 
 subsection {* some facts about beta *}
@@ -533,11 +502,38 @@
 
 section {* Candidates *}
 
-consts
+lemma ty_cases:
+  fixes x::ty
+  shows "(\<exists>X. x= TVar X) \<or> (\<exists>T1 T2. x=T1\<rightarrow>T2)"
+by (induct x rule: ty.induct_weak) (auto)
+
+function
   RED :: "ty \<Rightarrow> lam set"
-primrec
- "RED (TVar X) = {t. SN(t)}"
- "RED (\<tau>\<rightarrow>\<sigma>) =   {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
+where
+  "RED (TVar X) = {t. SN(t)}"
+| "RED (\<tau>\<rightarrow>\<sigma>) =   {t. \<forall>u. (u\<in>RED \<tau> \<longrightarrow> (App t u)\<in>RED \<sigma>)}"
+apply(auto simp add: ty.inject)
+apply(subgoal_tac "(\<exists>X. x= TVar X) \<or> (\<exists>T1 T2. x=T1\<rightarrow>T2)")
+apply(blast)
+apply(rule ty_cases)
+done
+
+instance ty :: size ..
+
+nominal_primrec
+  "size (TVar X) = 1"
+  "size (T\<^isub>1\<rightarrow>T\<^isub>2) = size T\<^isub>1 + size T\<^isub>2"
+by (rule TrueI)+
+
+lemma ty_size_greater_zero[simp]:
+  fixes T::"ty"
+  shows "size T > 0"
+by (nominal_induct T rule: ty.induct) (simp_all)
+
+termination
+apply(relation "measure size") 
+apply(auto)
+done
 
 constdefs
   NEUT :: "lam \<Rightarrow> bool"
@@ -546,7 +542,7 @@
 (* a slight hack to get the first element of applications *)
 inductive2 FST :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<guillemotright> _" [80,80] 80)
 where
-fst[intro!]:  "(App t s) \<guillemotright> t"
+  fst[intro!]:  "(App t s) \<guillemotright> t"
 
 lemma fst_elim[elim!]: 
   shows "(App t s) \<guillemotright> t' \<Longrightarrow> t=t'"
@@ -620,7 +616,7 @@
 
 lemma RED_props: 
   shows "CR1 \<tau>" and "CR2 \<tau>" and "CR3 \<tau>"
-proof (induct \<tau>)
+proof (nominal_induct \<tau> rule: ty.induct)
   case (TVar a)
   { case 1 show "CR1 (TVar a)" by (simp add: CR1_def)
   next
@@ -800,6 +796,9 @@
 apply(force simp add: RED_props)
 done
 
+(* HERE *)
+
+(*
 lemma fresh_domain: 
   assumes a: "a\<sharp>\<theta>"
   shows "a\<notin>set(domain \<theta>)"
@@ -817,27 +816,30 @@
 apply(auto simp add: fresh_prod fresh_list_cons)
 done
 
-lemma psubst_subst: 
-  assumes a: "c\<sharp>\<theta>"
+lemma fresh_psubst: 
+  fixes z::"name"
+  and   t::"lam"
+  assumes "z\<sharp>t" "z\<sharp>\<theta>"
+  shows "z\<sharp>(t[<\<theta>>])"
+using assms
+by (nominal_induct t avoiding: z \<theta> t rule: lam.induct)
+   (auto simp add: abs_fresh lookup_fresh)
+
+lemma psubst_subst:
+  assumes h:"c\<sharp>\<theta>"
   shows "(t[<\<theta>>])[c::=s] = t[<((c,s)#\<theta>)>]"
-using a
-apply(nominal_induct t avoiding: \<theta> c s rule: lam.induct)
-apply(auto dest: fresh_domain)
-apply(drule fresh_at)
-apply(assumption)
-apply(rule forget)
-apply(assumption)
-apply(subgoal_tac "name\<sharp>((c,s)#\<theta>)")(*A*)
-apply(simp)
-(*A*)
-apply(simp add: fresh_list_cons fresh_prod)
-done
+  using h
+by (nominal_induct t avoiding: \<theta> c s rule: lam.induct)
+ (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst)
+
 
 lemma all_RED: 
   assumes a: "\<Gamma>\<turnstile>t:\<tau>"
   and     b: "\<forall>a \<sigma>. (a,\<sigma>)\<in>set(\<Gamma>) \<longrightarrow> (a\<in>set(domain \<theta>)\<and>\<theta><a>\<in>RED \<sigma>)" 
   shows "t[<\<theta>>]\<in>RED \<tau>"
 using a b
+sorry
+
 proof(nominal_induct t avoiding: \<Gamma> \<tau> \<theta> rule: lam.induct)
   case (Lam a t) --"lambda case"
   have ih: "\<And>\<Gamma> \<tau> \<theta>. \<Gamma> \<turnstile> t:\<tau> \<Longrightarrow> 
@@ -860,6 +862,8 @@
   and     b: "\<forall>a \<sigma>. (a,\<sigma>)\<in>set(\<Gamma>) \<longrightarrow> (a\<in>set(domain \<theta>)\<and>\<theta><a>\<in>RED \<sigma>)" 
   shows "t[<\<theta>>]\<in>RED \<tau>"
 using a b
+sorry
+
 proof(nominal_induct t avoiding: \<Gamma> \<tau> \<theta> rule: lam.induct)
   case (Lam a t) --"lambda case"
   have ih: "\<And>\<Gamma> \<tau> \<theta>. \<lbrakk>\<Gamma> \<turnstile> t:\<tau>; \<forall>c \<sigma>. (c,\<sigma>)\<in>set \<Gamma> \<longrightarrow> c\<in>set (domain \<theta>) \<and>  \<theta><c>\<in>RED \<sigma>\<rbrakk> 
@@ -874,8 +878,10 @@
   hence "(Lam [a].(t[<\<theta>>])) \<in> RED (\<tau>1 \<rightarrow> \<tau>2)" by (simp only: abs_RED)
   thus "(Lam [a].t)[<\<theta>>] \<in> RED \<tau>" using fresh \<tau>_inst by simp
 qed (force dest!: t1_elim t2_elim)+
+*)
 
 (* identity substitution generated from a context \<Gamma> *)
+(*
 consts
   "id" :: "(name\<times>ty) list \<Rightarrow> (name\<times>lam) list"
 primrec
@@ -902,11 +908,14 @@
   shows "t[<(id \<Gamma>)>] = t"
 apply(nominal_induct t avoiding: \<Gamma> rule: lam.induct)
 apply(auto)
+sorry
+
 apply(simp add: id_var)
 apply(drule id_fresh)+
 apply(simp)
 done
 
+
 lemma id_mem:
   assumes a: "(a,\<tau>)\<in>set \<Gamma>"
   shows "a \<in> set (domain (id \<Gamma>))"
@@ -920,12 +929,12 @@
 apply(simp add: id_mem)
 apply(frule id_mem)
 apply(simp add: id_var)
-apply(subgoal_tac "CR3 \<sigma>")(*A*)
+apply(subgoal_tac "CR3 \<sigma>") --"A"
 apply(drule CR3_CR4)
 apply(simp add: CR4_def)
 apply(drule_tac x="Var a" in spec)
 apply(force simp add: NEUT_def NORMAL_Var)
-(*A*)
+--"A"
 apply(rule RED_props)
 done
 
@@ -951,4 +960,5 @@
   ultimately show "SN(t)" by (simp add: CR1_def)
 qed
 
+*)
 end
\ No newline at end of file
--- a/src/HOL/Nominal/Examples/Weakening.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Examples/Weakening.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -1,7 +1,7 @@
 (* $Id$ *)
 
 theory Weakening 
-imports "Nominal" 
+imports "../Nominal" 
 begin
 
 section {* Weakening Example for the Simply-Typed Lambda-Calculus *}
@@ -20,9 +20,9 @@
 
 lemma ty_perm[simp]:
   fixes pi ::"name prm"
-  and   \<tau>  ::"ty"
-  shows "pi\<bullet>\<tau> = \<tau>"
-by (induct \<tau> rule: ty.induct_weak)
+  and   T  ::"ty"
+  shows "pi\<bullet>T = T"
+by (induct T rule: ty.induct_weak)
    (simp_all add: perm_nat_def)  
 
 text {* valid contexts *}
@@ -37,35 +37,32 @@
   assumes a: "valid \<Gamma>"
   shows   "valid (pi\<bullet>\<Gamma>)"
 using a
-by (induct)
-   (auto simp add: fresh_bij)
-
-thm eqvt
+by (induct) (auto simp add: fresh_bij)
 
 text{* typing judgements *}
 inductive2
   typing :: "(name\<times>ty) list\<Rightarrow>lam\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [80,80,80] 80) 
 where
-    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (a,\<tau>)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var a : \<tau>"
-  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : \<tau>\<rightarrow>\<sigma>; \<Gamma> \<turnstile> t2 : \<tau>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : \<sigma>"
-  | t_Lam[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,\<tau>)#\<Gamma>) \<turnstile> t : \<sigma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : \<tau>\<rightarrow>\<sigma>"
+    t_Var[intro]: "\<lbrakk>valid \<Gamma>; (a,T)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var a : T"
+  | t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> t1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> t2 : T1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App t1 t2 : T2"
+  | t_Lam[intro]: "\<lbrakk>a\<sharp>\<Gamma>;((a,T1)#\<Gamma>) \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [a].t : T1\<rightarrow>T2"
 
 lemma eqvt_typing[eqvt]: 
   fixes pi:: "name prm"
-  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
-  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : (pi\<bullet>\<tau>)"
+  assumes a: "\<Gamma> \<turnstile> t : T"
+  shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : (pi\<bullet>T)"
 using a
 proof (induct)
-  case (t_Var \<Gamma> a \<tau>)
+  case (t_Var \<Gamma> a T)
   have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt_valid)
   moreover
-  have "(pi\<bullet>(a,\<tau>))\<in>(pi\<bullet>set \<Gamma>)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst])
-  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Var a) : (pi\<bullet>\<tau>)"
-    using typing.intros by (force simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
+  have "(pi\<bullet>(a,T))\<in>(pi\<bullet>set \<Gamma>)" by (rule pt_set_bij2[OF pt_name_inst, OF at_name_inst])
+  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Var a) : (pi\<bullet>T)"
+    using typing.intros by (force simp add: set_eqvt)
 next 
-  case (t_Lam a \<Gamma> \<tau> t \<sigma>)
+  case (t_Lam a \<Gamma> T1 t T2)
   moreover have "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)" by (simp add: fresh_bij)
-  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Lam [a].t) :(pi\<bullet>\<tau>\<rightarrow>\<sigma>)" by force 
+  ultimately show "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>Lam [a].t) :(pi\<bullet>T1\<rightarrow>T2)" by force 
 qed (auto)
 
 text {* the strong induction principle needs to be derived manually *}
@@ -74,51 +71,51 @@
   fixes  P :: "'a::fs_name\<Rightarrow>(name\<times>ty) list \<Rightarrow> lam \<Rightarrow> ty \<Rightarrow>bool"
   and    \<Gamma> :: "(name\<times>ty) list"
   and    t :: "lam"
-  and    \<tau> :: "ty"
+  and    T :: "ty"
   and    x :: "'a::fs_name"
-  assumes a: "\<Gamma> \<turnstile> t : \<tau>"
-  and a1:    "\<And>\<Gamma> a \<tau> x. \<lbrakk>valid \<Gamma>; (a,\<tau>) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P x \<Gamma> (Var a) \<tau>"
-  and a2:    "\<And>\<Gamma> \<tau> \<sigma> t1 t2 x. \<lbrakk>\<And>z. P z \<Gamma> t1 (\<tau>\<rightarrow>\<sigma>); \<And>z. P z \<Gamma> t2 \<tau>\<rbrakk>
-              \<Longrightarrow> P x \<Gamma> (App t1 t2) \<sigma>"
-  and a3:    "\<And>a \<Gamma> \<tau> \<sigma> t x. \<lbrakk>a\<sharp>x; a\<sharp>\<Gamma>; \<And>z. P z ((a,\<tau>)#\<Gamma>) t \<sigma>\<rbrakk>
-              \<Longrightarrow> P x \<Gamma> (Lam [a].t) (\<tau>\<rightarrow>\<sigma>)"
-  shows "P x \<Gamma> t \<tau>"
+  assumes a: "\<Gamma> \<turnstile> t : T"
+  and a1:    "\<And>\<Gamma> a T x. \<lbrakk>valid \<Gamma>; (a,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P x \<Gamma> (Var a) T"
+  and a2:    "\<And>\<Gamma> T1 T2 t1 t2 x. \<lbrakk>\<And>z. P z \<Gamma> t1 (T1\<rightarrow>T2); \<And>z. P z \<Gamma> t2 T1\<rbrakk>
+              \<Longrightarrow> P x \<Gamma> (App t1 t2) T2"
+  and a3:    "\<And>a \<Gamma> T1 T2 t x. \<lbrakk>a\<sharp>x; a\<sharp>\<Gamma>; \<And>z. P z ((a,T1)#\<Gamma>) t T2\<rbrakk>
+              \<Longrightarrow> P x \<Gamma> (Lam [a].t) (T1\<rightarrow>T2)"
+  shows "P x \<Gamma> t T"
 proof -
-  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>t) (pi\<bullet>\<tau>)"
+  from a have "\<And>(pi::name prm) x. P x (pi\<bullet>\<Gamma>) (pi\<bullet>t) (pi\<bullet>T)"
   proof (induct)
-    case (t_Var \<Gamma> a \<tau>)
+    case (t_Var \<Gamma> a T)
     have "valid \<Gamma>" by fact
     then have "valid (pi\<bullet>\<Gamma>)" by (rule eqvt)
     moreover
-    have "(a,\<tau>)\<in>set \<Gamma>" by fact
-    then have "pi\<bullet>(a,\<tau>)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
-    then have "(pi\<bullet>a,\<tau>)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: pt_list_set_pi[OF pt_name_inst])
-    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) (pi\<bullet>\<tau>)" using a1 by simp
+    have "(a,T)\<in>set \<Gamma>" by fact
+    then have "pi\<bullet>(a,T)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])  
+    then have "(pi\<bullet>a,T)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: set_eqvt)
+    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Var a)) (pi\<bullet>T)" using a1 by simp
   next
-    case (t_App \<Gamma> t1 \<tau> \<sigma> t2)
-    thus "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(App t1 t2)) (pi\<bullet>\<sigma>)" using a2 
+    case (t_App \<Gamma> t1 T1 T2 t2)
+    thus "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(App t1 t2)) (pi\<bullet>T2)" using a2 
       by (simp only: eqvt) (blast)
   next
-    case (t_Lam a \<Gamma> \<tau> t \<sigma>)
+    case (t_Lam a \<Gamma> T1 t T2)
     obtain c::"name" where fs: "c\<sharp>(pi\<bullet>a,pi\<bullet>t,pi\<bullet>\<Gamma>,x)" by (rule exists_fresh[OF fs_name1])
     let ?sw="[(pi\<bullet>a,c)]"
     let ?pi'="?sw@pi"
     have f1: "a\<sharp>\<Gamma>" by fact
     have f2: "(pi\<bullet>a)\<sharp>(pi\<bullet>\<Gamma>)" using f1 by (simp add: fresh_bij)
     have f3: "c\<sharp>?pi'\<bullet>\<Gamma>" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
-    have ih1: "\<And>x. P x (?pi'\<bullet>((a,\<tau>)#\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>\<sigma>)" by fact
-    then have "\<And>x. P x ((c,\<tau>)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>\<sigma>)" by (simp add: calc_atm)
-    then have "P x (?pi'\<bullet>\<Gamma>) (Lam [c].(?pi'\<bullet>t)) (\<tau>\<rightarrow>\<sigma>)" using a3 f3 fs by simp
-    then have "P x (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) (\<tau>\<rightarrow>\<sigma>)" 
+    have ih1: "\<And>x. P x (?pi'\<bullet>((a,T1)#\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>T2)" by fact
+    then have "\<And>x. P x ((c,T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>t) (?pi'\<bullet>T2)" by (simp add: calc_atm)
+    then have "P x (?pi'\<bullet>\<Gamma>) (Lam [c].(?pi'\<bullet>t)) (T1\<rightarrow>T2)" using a3 f3 fs by simp
+    then have "P x (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) (T1\<rightarrow>T2)" 
       by (simp del: append_Cons add: calc_atm pt_name2)
     moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)" 
       by (rule perm_fresh_fresh) (simp_all add: fs f2)
     moreover have "(?sw\<bullet>(Lam [(pi\<bullet>a)].(pi\<bullet>t))) = Lam [(pi\<bullet>a)].(pi\<bullet>t)" 
       by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
-    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [a].t)) (pi\<bullet>\<tau>\<rightarrow>\<sigma>)" by (simp)
+    ultimately show "P x (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [a].t)) (pi\<bullet>T1\<rightarrow>T2)" by (simp)
   qed
-  hence "P x (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>\<tau>)" by blast
-  thus "P x \<Gamma> t \<tau>" by simp
+  hence "P x (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>t) (([]::name prm)\<bullet>T)" by blast
+  thus "P x \<Gamma> t T" by simp
 qed
 
 text {* definition of a subcontext *}
@@ -130,92 +127,91 @@
 text {* now it comes: The Weakening Lemma *}
 
 lemma weakening_version1: 
-  assumes a: "\<Gamma>1 \<turnstile> t : \<sigma>" 
+  assumes a: "\<Gamma>1 \<turnstile> t : T" 
   and     b: "valid \<Gamma>2" 
   and     c: "\<Gamma>1 \<lless> \<Gamma>2"
-  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
+  shows "\<Gamma>2 \<turnstile> t : T"
 using a b c
-apply(nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
-apply(auto | atomize)+
+by (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing_induct)
+   (auto | atomize)+
 (* FIXME: meta-quantifiers seem to be not as "automatic" as object-quantifiers *)
-done
 
 lemma weakening_version2: 
   fixes \<Gamma>1::"(name\<times>ty) list"
   and   t ::"lam"
   and   \<tau> ::"ty"
-  assumes a: "\<Gamma>1 \<turnstile> t:\<sigma>"
+  assumes a: "\<Gamma>1 \<turnstile> t:T"
   and     b: "valid \<Gamma>2" 
   and     c: "\<Gamma>1 \<lless> \<Gamma>2"
-  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
+  shows "\<Gamma>2 \<turnstile> t:T"
 using a b c
-proof (nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
-  case (t_Var \<Gamma>1 a \<tau>)  (* variable case *)
+proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing_induct)
+  case (t_Var \<Gamma>1 a T)  (* variable case *)
   have "\<Gamma>1 \<lless> \<Gamma>2" by fact 
   moreover  
   have "valid \<Gamma>2" by fact 
   moreover 
-  have "(a,\<tau>)\<in> set \<Gamma>1" by fact
-  ultimately show "\<Gamma>2 \<turnstile> Var a : \<tau>" by auto
+  have "(a,T)\<in> set \<Gamma>1" by fact
+  ultimately show "\<Gamma>2 \<turnstile> Var a : T" by auto
 next
-  case (t_Lam a \<Gamma>1 \<tau> \<sigma> t) (* lambda case *)
+  case (t_Lam a \<Gamma>1 T1 T2 t) (* lambda case *)
   have vc: "a\<sharp>\<Gamma>2" by fact (* variable convention *)
-  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,\<tau>)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t:\<sigma>" by fact
+  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,T1)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t:T2" by fact
   have "\<Gamma>1 \<lless> \<Gamma>2" by fact
-  then have "((a,\<tau>)#\<Gamma>1) \<lless> ((a,\<tau>)#\<Gamma>2)" by simp
+  then have "((a,T1)#\<Gamma>1) \<lless> ((a,T1)#\<Gamma>2)" by simp
   moreover
   have "valid \<Gamma>2" by fact
-  then have "valid ((a,\<tau>)#\<Gamma>2)" using vc by (simp add: v2)
-  ultimately have "((a,\<tau>)#\<Gamma>2) \<turnstile> t:\<sigma>" using ih by simp
-  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : \<tau> \<rightarrow> \<sigma>" by auto
+  then have "valid ((a,T1)#\<Gamma>2)" using vc by (simp add: v2)
+  ultimately have "((a,T1)#\<Gamma>2) \<turnstile> t:T2" using ih by simp
+  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : T1\<rightarrow>T2" by auto
 qed (auto) (* app case *)
 
 lemma weakening_version3: 
-  assumes a: "\<Gamma>1 \<turnstile> t:\<sigma>"
+  assumes a: "\<Gamma>1 \<turnstile> t : T"
   and     b: "valid \<Gamma>2" 
   and     c: "\<Gamma>1 \<lless> \<Gamma>2"
-  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
+  shows "\<Gamma>2 \<turnstile> t : T"
 using a b c
-proof (nominal_induct \<Gamma>1 t \<sigma> avoiding: \<Gamma>2 rule: typing_induct)
-  case (t_Lam a \<Gamma>1 \<tau> \<sigma> t) (* lambda case *)
+proof (nominal_induct \<Gamma>1 t T avoiding: \<Gamma>2 rule: typing_induct)
+  case (t_Lam a \<Gamma>1 T1 T2 t) (* lambda case *)
   have vc: "a\<sharp>\<Gamma>2" by fact (* variable convention *)
-  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,\<tau>)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t:\<sigma>" by fact
+  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,T1)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk> \<Longrightarrow>  \<Gamma>3 \<turnstile> t : T2" by fact
   have "\<Gamma>1 \<lless> \<Gamma>2" by fact
-  then have "((a,\<tau>)#\<Gamma>1) \<lless> ((a,\<tau>)#\<Gamma>2)" by simp
+  then have "((a,T1)#\<Gamma>1) \<lless> ((a,T1)#\<Gamma>2)" by simp
   moreover
   have "valid \<Gamma>2" by fact
-  then have "valid ((a,\<tau>)#\<Gamma>2)" using vc by (simp add: v2)
-  ultimately have "((a,\<tau>)#\<Gamma>2) \<turnstile> t:\<sigma>" using ih by simp
-  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : \<tau> \<rightarrow> \<sigma>" by auto
+  then have "valid ((a,T1)#\<Gamma>2)" using vc by (simp add: v2)
+  ultimately have "((a,T1)#\<Gamma>2) \<turnstile> t : T2" using ih by simp
+  with vc show "\<Gamma>2 \<turnstile> (Lam [a].t) : T1 \<rightarrow> T2" by auto
 qed (auto) (* app and var case *)
 
 text{* The original induction principle for the typing relation
        is not strong enough - even this simple lemma fails:     *}
 lemma weakening_too_weak: 
-  assumes a: "\<Gamma>1 \<turnstile> t:\<sigma>"
+  assumes a: "\<Gamma>1 \<turnstile> t : T"
   and     b: "valid \<Gamma>2" 
   and     c: "\<Gamma>1 \<lless> \<Gamma>2"
-  shows "\<Gamma>2 \<turnstile> t:\<sigma>"
+  shows "\<Gamma>2 \<turnstile> t : T"
 using a b c
 proof (induct arbitrary: \<Gamma>2)
-  case (t_Var \<Gamma>1 a \<tau>) (* variable case *)
+  case (t_Var \<Gamma>1 a T) (* variable case *)
   have "\<Gamma>1 \<lless> \<Gamma>2" by fact
   moreover
   have "valid \<Gamma>2" by fact
   moreover
-  have "(a,\<tau>) \<in> (set \<Gamma>1)" by fact 
-  ultimately show "\<Gamma>2 \<turnstile> Var a : \<tau>" by auto
+  have "(a,T) \<in> (set \<Gamma>1)" by fact 
+  ultimately show "\<Gamma>2 \<turnstile> Var a : T" by auto
 next
-  case (t_Lam a \<Gamma>1 \<tau> t \<sigma>) (* lambda case *)
+  case (t_Lam a \<Gamma>1 T1 t T2) (* lambda case *)
   (* all assumptions available in this case*)
   have a0: "a\<sharp>\<Gamma>1" by fact
-  have a1: "((a,\<tau>)#\<Gamma>1) \<turnstile> t : \<sigma>" by fact
+  have a1: "((a,T1)#\<Gamma>1) \<turnstile> t : T2" by fact
   have a2: "\<Gamma>1 \<lless> \<Gamma>2" by fact
   have a3: "valid \<Gamma>2" by fact
-  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,\<tau>)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk>  \<Longrightarrow>  \<Gamma>3 \<turnstile> t:\<sigma>" by fact
-  have "((a,\<tau>)#\<Gamma>1) \<lless> ((a,\<tau>)#\<Gamma>2)" using a2 by simp
+  have ih: "\<And>\<Gamma>3. \<lbrakk>valid \<Gamma>3; ((a,T1)#\<Gamma>1) \<lless> \<Gamma>3\<rbrakk>  \<Longrightarrow>  \<Gamma>3 \<turnstile> t : T2" by fact
+  have "((a,T1)#\<Gamma>1) \<lless> ((a,T1)#\<Gamma>2)" using a2 by simp
   moreover
-  have "valid ((a,\<tau>)#\<Gamma>2)" using v2 (* fails *) 
+  have "valid ((a,T1)#\<Gamma>2)" using v2 (* fails *) 
     oops
 
 end
\ No newline at end of file
--- a/src/HOL/Nominal/Nominal.thy	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/Nominal.thy	Tue Mar 06 15:28:22 2007 +0100
@@ -32,15 +32,15 @@
 defs (unchecked overloaded)
   perm_set_def:  "pi\<bullet>(X::'a set) \<equiv> {pi\<bullet>a | a. a\<in>X}"
 
-lemma perm_empty:
+lemma empty_eqvt:
   shows "pi\<bullet>{} = {}"
   by (simp add: perm_set_def)
 
-lemma perm_union:
+lemma union_eqvt:
   shows "pi \<bullet> (X \<union> Y) = (pi \<bullet> X) \<union> (pi \<bullet> Y)"
   by (auto simp add: perm_set_def)
 
-lemma perm_insert:
+lemma insert_eqvt:
   shows "pi\<bullet>(insert x X) = insert (pi\<bullet>x) (pi\<bullet>X)"
   by (auto simp add: perm_set_def)
 
@@ -51,40 +51,44 @@
 primrec (unchecked perm_prod)
   "pi\<bullet>(a,b) = (pi\<bullet>a,pi\<bullet>b)"
 
-lemma perm_fst:
+lemma fst_eqvt:
   "pi\<bullet>(fst x) = fst (pi\<bullet>x)"
  by (cases x) simp
 
-lemma perm_snd:
+lemma snd_eqvt:
   "pi\<bullet>(snd x) = snd (pi\<bullet>x)"
  by (cases x) simp
 
 (* permutation on lists *)
 primrec (unchecked perm_list)
-  perm_nil_def:  "pi\<bullet>[]     = []"
-  perm_cons_def: "pi\<bullet>(x#xs) = (pi\<bullet>x)#(pi\<bullet>xs)"
-
-lemma perm_append:
+  nil_eqvt:  "pi\<bullet>[]     = []"
+  cons_eqvt: "pi\<bullet>(x#xs) = (pi\<bullet>x)#(pi\<bullet>xs)"
+
+lemma append_eqvt:
   fixes pi :: "'x prm"
   and   l1 :: "'a list"
   and   l2 :: "'a list"
   shows "pi\<bullet>(l1@l2) = (pi\<bullet>l1)@(pi\<bullet>l2)"
   by (induct l1) auto
 
-lemma perm_rev:
+lemma rev_eqvt:
   fixes pi :: "'x prm"
   and   l  :: "'a list"
   shows "pi\<bullet>(rev l) = rev (pi\<bullet>l)"
-  by (induct l) (simp_all add: perm_append)
+  by (induct l) (simp_all add: append_eqvt)
 
 (* permutation on functions *)
 defs (unchecked overloaded)
   perm_fun_def: "pi\<bullet>(f::'a\<Rightarrow>'b) \<equiv> (\<lambda>x. pi\<bullet>f((rev pi)\<bullet>x))"
 
+lemma swap_fun: 
+  shows "[(a,b)]\<bullet>(f::'a\<Rightarrow>'b) \<equiv> (\<lambda>x. [(a,b)]\<bullet>f([(a,b)]\<bullet>x))"
+by (unfold perm_fun_def,auto)
+
 (* permutation on bools *)
 primrec (unchecked perm_bool)
-  perm_true_def:  "pi\<bullet>True  = True"
-  perm_false_def: "pi\<bullet>False = False"
+  true_eqvt:  "pi\<bullet>True  = True"
+  false_eqvt: "pi\<bullet>False = False"
 
 lemma perm_bool:
   shows "pi\<bullet>(b::bool) = b"
@@ -100,25 +104,24 @@
   shows "P"
   using a by (simp add: perm_bool)
 
-lemma perm_if:
+lemma if_eqvt:
   fixes pi::"'a prm"
   shows "pi\<bullet>(if b then c1 else c2) = (if (pi\<bullet>b) then (pi\<bullet>c1) else (pi\<bullet>c2))"
 apply(simp add: perm_fun_def)
 done
 
-
 (* permutation on options *)
 
 primrec (unchecked perm_option)
-  perm_some_def:  "pi\<bullet>Some(x) = Some(pi\<bullet>x)"
-  perm_none_def:  "pi\<bullet>None    = None"
+  some_eqvt:  "pi\<bullet>Some(x) = Some(pi\<bullet>x)"
+  none_eqvt:  "pi\<bullet>None    = None"
 
 (* a "private" copy of the option type used in the abstraction function *)
 datatype 'a noption = nSome 'a | nNone
 
 primrec (unchecked perm_noption)
-  perm_nSome_def: "pi\<bullet>nSome(x) = nSome(pi\<bullet>x)"
-  perm_nNone_def: "pi\<bullet>nNone    = nNone"
+  nSome_eqvt: "pi\<bullet>nSome(x) = nSome(pi\<bullet>x)"
+  nNone_eqvt: "pi\<bullet>nNone    = nNone"
 
 (* a "private" copy of the product type used in the nominal induct method *)
 datatype ('a,'b) nprod = nPair 'a 'b
@@ -1134,6 +1137,16 @@
   thus "pi\<bullet>x = pi\<bullet>y" by simp
 qed
 
+lemma pt_eq_eqvt:
+  fixes pi :: "'x prm"
+  and   x  :: "'a"
+  and   y  :: "'a"
+  assumes pt: "pt TYPE('a) TYPE('x)"
+  and     at: "at TYPE('x)"
+  shows "pi \<bullet> (x=y) = (pi\<bullet>x = pi\<bullet>y)"
+using assms
+by (auto simp add: pt_bij perm_bool)
+
 lemma pt_bij3:
   fixes pi :: "'x prm"
   and   x  :: "'a"
@@ -1202,6 +1215,16 @@
   shows "((pi\<bullet>x)\<in>(pi\<bullet>X)) = (x\<in>X)"
   by (simp add: perm_set_def pt_bij[OF pt, OF at])
 
+lemma pt_in_eqvt:
+  fixes pi :: "'x prm"
+  and   x  :: "'a"
+  and   X  :: "'a set"
+  assumes pt: "pt TYPE('a) TYPE('x)"
+  and     at: "at TYPE('x)"
+  shows "pi\<bullet>(x\<in>X)=((pi\<bullet>x)\<in>(pi\<bullet>X))"
+using assms
+by (auto simp add:  pt_set_bij perm_bool)
+
 lemma pt_set_bij2:
   fixes pi :: "'x prm"
   and   x  :: "'a"
@@ -1668,6 +1691,17 @@
 apply(simp add: pt_rev_pi[OF pt, OF at])
 done
 
+lemma pt_ex_eqvt:
+  fixes  pi :: "'x prm"
+  and     x :: "'a"
+  assumes pt: "pt TYPE('a) TYPE('x)"
+  and     at: "at TYPE('x)"
+  shows "pi\<bullet>(\<exists>(x::'a). P x) = (\<exists> (x::'a). (pi\<bullet>P) x)"
+apply(auto simp add: perm_bool perm_fun_def)
+apply(rule_tac x="pi\<bullet>x" in exI) 
+apply(simp add: pt_rev_pi[OF pt, OF at])
+done
+
 lemma imp_eqvt:
   fixes pi::"'x prm"
   shows "pi\<bullet>(A\<longrightarrow>B) = ((pi\<bullet>A)\<longrightarrow>(pi\<bullet>B))"
@@ -1678,6 +1712,16 @@
   shows "pi\<bullet>(A\<and>B) = ((pi\<bullet>A)\<and>(pi\<bullet>B))"
   by (simp add: perm_bool)
 
+lemma disj_eqvt:
+  fixes pi::"'x prm"
+  shows "pi\<bullet>(A\<or>B) = ((pi\<bullet>A)\<or>(pi\<bullet>B))"
+  by (simp add: perm_bool)
+
+lemma neg_eqvt:
+  fixes pi::"'x prm"
+  shows "pi\<bullet>(\<not> A) = (\<not> (pi\<bullet>A))"
+  by (simp add: perm_bool)
+
 section {* facts about supports *}
 (*==============================*)
 
@@ -2223,7 +2267,7 @@
 apply(simp add: fresh_fin_insert[OF pt, OF at, OF fs, OF f])
 done
 
-lemma pt_list_set_pi:
+lemma pt_set_eqvt:
   fixes pi :: "'x prm"
   and   xs :: "'a list"
   assumes pt: "pt TYPE('a) TYPE('x)"
@@ -2331,7 +2375,7 @@
   shows "cp TYPE ('a\<Rightarrow>'b) TYPE('x) TYPE('y)"
 using c1 c2
 apply(auto simp add: cp_def perm_fun_def expand_fun_eq)
-apply(simp add: perm_rev[symmetric])
+apply(simp add: rev_eqvt[symmetric])
 apply(simp add: pt_rev_pi[OF pt_list_inst[OF pt_prod_inst[OF pt, OF pt]], OF at])
 done
 
@@ -3003,53 +3047,94 @@
     done
 qed
 
+(************************)
+(* Various eqvt-lemmas  *)
+
+lemma Zero_nat_eqvt:
+ shows "pi\<bullet> (0::nat) = 0" 
+by (auto simp add: perm_nat_def)
+
+lemma One_nat_eqvt:
+ shows "pi\<bullet> (1::nat) = 1"
+by (simp add: perm_nat_def)
+
+lemma Suc_eqvt:
+ shows "pi\<bullet> Suc x = Suc (pi\<bullet>x)" 
+by (auto simp add: perm_nat_def)
+
+lemma numeral_nat_eqvt: 
+ shows "pi\<bullet> ((number_of n)::nat) = number_of ( n)" 
+by (simp add: perm_nat_def perm_int_def)
+
+lemma max_nat_eqvt:
+ shows "pi\<bullet>(max (x::nat) y) = max (pi\<bullet>x) (pi\<bullet>y)" 
+by (simp add:perm_nat_def) 
+
+lemma min_nat_eqvt:
+ shows "pi\<bullet> max (x::nat) y = max (pi\<bullet>x) (pi\<bullet>y)" 
+by (simp add:perm_nat_def) 
+
+lemma plus_nat_eqvt:
+ shows "pi\<bullet>((x::nat) + y) = (pi\<bullet>x) + (pi\<bullet>y)" 
+by (simp add:perm_nat_def) 
+
+lemma minus_nat_eqvt:
+ shows "pi\<bullet>((x::nat) - y) = (pi\<bullet>x) - (pi\<bullet>y)" 
+by (simp add:perm_nat_def) 
+
+lemma mult_nat_eqvt:
+ shows "pi\<bullet>((x::nat) * y) = (pi\<bullet>x) * (pi\<bullet>y)" 
+by (simp add:perm_nat_def) 
+
+lemma div_nat_eqvt:
+ shows "pi\<bullet>((x::nat) div y) = (pi\<bullet>x) div (pi\<bullet>y)" 
+by (simp add:perm_nat_def) 
+
 (*******************************************************)
 (* Setup of the theorem attributes eqvt, fresh and bij *)
 use "nominal_thmdecls.ML"
 setup "NominalThmDecls.setup"
 
-lemmas  [eqvt] = perm_list.simps perm_prod.simps 
-
-
+lemmas [eqvt] = 
+  if_eqvt
+  perm_unit.simps
+  perm_list.simps 
+  perm_prod.simps 
+  imp_eqvt disj_eqvt conj_eqvt neg_eqvt
+  Suc_eqvt Zero_nat_eqvt One_nat_eqvt
+  min_nat_eqvt max_nat_eqvt
+  plus_nat_eqvt minus_nat_eqvt mult_nat_eqvt div_nat_eqvt
+  union_eqvt empty_eqvt insert_eqvt
+  fst_eqvt snd_eqvt
+ 
+(* this lemma does not conform with the form of an eqvt-lemma but is *)
+(* still useful to have when analysing permutations on numbers       *)
+lemmas [eqvt_force] = numeral_nat_eqvt 
 
 (***************************************)
 (* setup for the individial atom-kinds *)
 (* and nominal datatypes               *)
 use "nominal_atoms.ML"
-(* permutation equality tactic *)
+setup "NominalAtoms.setup"
+
+(************************************************************)
+(* various tactics for analysing permutations, supports etc *)
 use "nominal_permeq.ML";
-use "nominal_package.ML"
-setup "NominalAtoms.setup"
-setup "NominalPackage.setup"
-
-(** primitive recursive functions on nominal datatypes **)
-use "nominal_primrec.ML"
-
-(** inductive predicates involving nominal datatypes **)
-use "nominal_inductive.ML"
-
-(*****************************************)
-(* setup for induction principles method *)
-
-use "nominal_induct.ML";
-method_setup nominal_induct =
-  {* NominalInduct.nominal_induct_method *}
-  {* nominal induction *}
 
 method_setup perm_simp =
-  {* NominalPermeq.perm_eq_meth *}
+  {* NominalPermeq.perm_simp_meth *}
   {* simp rules and simprocs for analysing permutations *}
 
 method_setup perm_simp_debug =
-  {* NominalPermeq.perm_eq_meth_debug *}
+  {* NominalPermeq.perm_simp_meth_debug *}
   {* simp rules and simprocs for analysing permutations including debugging facilities *}
 
 method_setup perm_full_simp =
-  {* NominalPermeq.perm_full_eq_meth *}
+  {* NominalPermeq.perm_full_simp_meth *}
   {* tactic for deciding equalities involving permutations *}
 
 method_setup perm_full_simp_debug =
-  {* NominalPermeq.perm_full_eq_meth_debug *}
+  {* NominalPermeq.perm_full_simp_meth_debug *}
   {* tactic for deciding equalities involving permutations including debugging facilities *}
 
 method_setup supports_simp =
@@ -3061,19 +3146,40 @@
   {* tactic for deciding whether something supports something else including debugging facilities *}
 
 method_setup finite_guess =
-  {* NominalPermeq.finite_gs_meth *}
+  {* NominalPermeq.finite_guess_meth *}
   {* tactic for deciding whether something has finite support *}
 
 method_setup finite_guess_debug =
-  {* NominalPermeq.finite_gs_meth_debug *}
+  {* NominalPermeq.finite_guess_meth_debug *}
   {* tactic for deciding whether something has finite support including debugging facilities *}
 
 method_setup fresh_guess =
-  {* NominalPermeq.fresh_gs_meth *}
+  {* NominalPermeq.fresh_guess_meth *}
   {* tactic for deciding whether an atom is fresh for something*}
 
 method_setup fresh_guess_debug =
-  {* NominalPermeq.fresh_gs_meth_debug *}
+  {* NominalPermeq.fresh_guess_meth_debug *}
   {* tactic for deciding whether an atom is fresh for something including debugging facilities *}
 
+(************************************************)
+(* main file for constructing nominal datatypes *)
+use "nominal_package.ML"
+setup "NominalPackage.setup"
+
+(******************************************************)
+(* primitive recursive functions on nominal datatypes *)
+use "nominal_primrec.ML"
+
+(****************************************************)
+(* inductive definition involving nominal datatypes *)
+use "nominal_inductive.ML"
+
+(*****************************************)
+(* setup for induction principles method *)
+use "nominal_induct.ML";
+method_setup nominal_induct =
+  {* NominalInduct.nominal_induct_method *}
+  {* nominal induction *}
+
+
 end
--- a/src/HOL/Nominal/nominal_atoms.ML	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/nominal_atoms.ML	Tue Mar 06 15:28:22 2007 +0100
@@ -8,6 +8,9 @@
 signature NOMINAL_ATOMS =
 sig
   val create_nom_typedecls : string list -> theory -> theory
+  type atom_info
+  val get_atom_infos : theory -> atom_info Symtab.table
+  val get_atom_info : theory -> string -> atom_info option
   val atoms_of : theory -> string list
   val mk_permT : typ -> typ
   val setup : theory -> theory
@@ -22,10 +25,15 @@
 
 (* data kind 'HOL/nominal' *)
 
+type atom_info =
+  {pt_class : string,
+   fs_class : string,
+   cp_classes : (string * string) list};
+
 structure NominalArgs =
 struct
   val name = "HOL/nominal";
-  type T = unit Symtab.table;
+  type T = atom_info Symtab.table;
 
   val empty = Symtab.empty;
   val copy = I;
@@ -37,6 +45,14 @@
 
 structure NominalData = TheoryDataFun(NominalArgs);
 
+fun make_atom_info ((pt_class, fs_class), cp_classes) =
+  {pt_class = pt_class,
+   fs_class = fs_class,
+   cp_classes = cp_classes};
+
+val get_atom_infos = NominalData.get;
+val get_atom_info = Symtab.lookup o NominalData.get;
+
 fun atoms_of thy = map fst (Symtab.dest (NominalData.get thy));
 
 fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T));
@@ -259,7 +275,7 @@
        in  
         AxClass.define_class_i (cl_name, [pt_name]) [] [((cl_name ^ "1", []), [axiom1])] thy            
        end) ak_names_types thy8; 
-
+	 
      (* proves that every fs_<ak>-type together with <ak>-type   *)
      (* instance of fs-type                                      *)
      (* lemma abst_<ak>_inst:                                    *)
@@ -290,7 +306,7 @@
        (* declares for every atom-kind combination an axclass            *)
        (* cp_<ak1>_<ak2> giving a composition property                   *)
        (* cp_<ak1>_<ak2>1: pi1 o pi2 o x = (pi1 o pi2) o (pi1 o x)       *)
-        val (_,thy12b) = fold_map (fn (ak_name, T) => fn thy =>
+        val (cp_ax_classes,thy12b) = fold_map (fn (ak_name, T) => fn thy =>
 	 fold_map (fn (ak_name', T') => fn thy' =>
 	     let
 	       val cl_name = "cp_"^ak_name^"_"^ak_name';
@@ -684,7 +700,12 @@
        val fresh_aux_ineq      = thm "Nominal.pt_fresh_aux_ineq";
        val fresh_aux           = thm "Nominal.pt_fresh_aux";
        val fresh_eqvt          = thm "Nominal.pt_fresh_eqvt";
+       val set_eqvt            = thm "Nominal.pt_set_eqvt";
+       val set_diff_eqvt       = thm "Nominal.pt_set_diff_eqvt";
+       val in_eqvt             = thm "Nominal.pt_in_eqvt";
+       val eq_eqvt             = thm "Nominal.pt_eq_eqvt";
        val all_eqvt            = thm "Nominal.pt_all_eqvt";
+       val ex_eqvt             = thm "Nominal.pt_ex_eqvt";
        val pt_pi_rev           = thm "Nominal.pt_pi_rev";
        val pt_rev_pi           = thm "Nominal.pt_rev_pi";
        val at_exists_fresh     = thm "Nominal.at_exists_fresh";
@@ -716,6 +737,7 @@
 
              (* FIXME: these lists do not need to be created dynamically again *)
 
+             
              (* list of all at_inst-theorems *)
              val ats = map (fn ak => PureThy.get_thm thy32 (Name ("at_"^ak^"_inst"))) ak_names
              (* list of all pt_inst-theorems *)
@@ -736,6 +758,8 @@
 	       in List.mapPartial I (map djs_fun (Library.product ak_names ak_names)) end;
              (* list of all fs_inst-theorems *)
              val fss = map (fn ak => PureThy.get_thm thy32 (Name ("fs_"^ak^"_inst"))) ak_names
+             (* list of all at_inst-theorems *)
+             val fs_axs = map (fn ak => PureThy.get_thm thy32 (Name ("fs_"^ak^"1"))) ak_names
 
              fun inst_pt thms = Library.flat (map (fn ti => instR ti pts) thms);
              fun inst_at thms = Library.flat (map (fn ti => instR ti ats) thms);
@@ -770,7 +794,8 @@
             ||>> PureThy.add_thmss [(("supp_atm", (inst_at [at_supp]) @ (inst_dj [dj_supp])),[])]
             ||>> PureThy.add_thmss [(("exists_fresh", inst_at [at_exists_fresh]),[])]
             ||>> PureThy.add_thmss [(("exists_fresh'", inst_at [at_exists_fresh']),[])]
-            ||>> PureThy.add_thmss [(("all_eqvt", inst_pt_at [all_eqvt]),[])]
+            ||>> PureThy.add_thmss [(("all_eqvt", inst_pt_at [all_eqvt]),[NominalThmDecls.eqvt_add])] 
+            ||>> PureThy.add_thmss [(("ex_eqvt", inst_pt_at [ex_eqvt]),[NominalThmDecls.eqvt_add])]
             ||>> PureThy.add_thmss 
 	      let val thms1 = inst_at [at_fresh]
 		  val thms2 = inst_dj [at_fresh_ineq]
@@ -808,19 +833,35 @@
 	      in [(("fresh_right", thms1 @ thms2),[])] end
             ||>> PureThy.add_thmss
 	      let val thms1 = inst_pt_at [fresh_bij]
-		  and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq]
+ 		  and thms2 = inst_pt_pt_at_cp [fresh_bij_ineq]
 	      in [(("fresh_bij", thms1 @ thms2),[])] end
             ||>> PureThy.add_thmss
 	      let val thms1 = inst_pt_at [fresh_eqvt]
-	      in [(("fresh_eqvt", thms1),[])] end
+	      in [(("fresh_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
+            ||>> PureThy.add_thmss
+	      let val thms1 = inst_pt_at [in_eqvt]
+	      in [(("in_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
+  	    ||>> PureThy.add_thmss
+	      let val thms1 = inst_pt_at [eq_eqvt]
+	      in [(("eq_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
+	    ||>> PureThy.add_thmss
+	      let val thms1 = inst_pt [set_eqvt]
+	      in [(("set_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
+	    ||>> PureThy.add_thmss
+	      let val thms1 = inst_pt_at [set_diff_eqvt]
+	      in [(("set_diff_eqvt", thms1),[NominalThmDecls.eqvt_add])] end
             ||>> PureThy.add_thmss
 	      let val thms1 = inst_pt_at [fresh_aux]
 		  and thms2 = inst_pt_pt_at_cp_dj [fresh_aux_ineq]
 	      in [(("fresh_aux", thms1 @ thms2),[])] end
+            ||>> PureThy.add_thmss [(("fin_supp",fs_axs),[])]
 	   end;
 
-    in NominalData.put (fold Symtab.update (map (rpair ()) full_ak_names)
-      (NominalData.get thy11)) thy33
+    in 
+      NominalData.map (fold Symtab.update (full_ak_names ~~ map make_atom_info
+        (pt_ax_classes ~~
+         fs_ax_classes ~~
+         map (fn cls => full_ak_names ~~ cls) cp_ax_classes))) thy33
     end;
 
 
--- a/src/HOL/Nominal/nominal_permeq.ML	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/nominal_permeq.ML	Tue Mar 06 15:28:22 2007 +0100
@@ -1,6 +1,6 @@
 (*  Title:      HOL/Nominal/nominal_permeq.ML
     ID:         $Id$
-    Author:     Christian Urban, TU Muenchen
+    Authors:    Christian Urban, Julien Narboux, TU Muenchen
 
 Methods for simplifying permutations and
 for analysing equations involving permutations.
@@ -33,95 +33,140 @@
   val finite_guess_tac : simpset -> int -> tactic
   val fresh_guess_tac : simpset -> int -> tactic
 
-  val perm_eq_meth : Method.src -> Proof.context -> Proof.method
-  val perm_eq_meth_debug : Method.src -> Proof.context -> Proof.method
-  val perm_full_eq_meth : Method.src -> Proof.context -> Proof.method
-  val perm_full_eq_meth_debug : Method.src -> Proof.context -> Proof.method
+  val perm_simp_meth : Method.src -> Proof.context -> Proof.method
+  val perm_simp_meth_debug : Method.src -> Proof.context -> Proof.method
+  val perm_full_simp_meth : Method.src -> Proof.context -> Proof.method
+  val perm_full_simp_meth_debug : Method.src -> Proof.context -> Proof.method
   val supports_meth : Method.src -> Proof.context -> Proof.method
   val supports_meth_debug : Method.src -> Proof.context -> Proof.method
-  val finite_gs_meth : Method.src -> Proof.context -> Proof.method
-  val finite_gs_meth_debug : Method.src -> Proof.context -> Proof.method
-  val fresh_gs_meth : Method.src -> Proof.context -> Proof.method
-  val fresh_gs_meth_debug : Method.src -> Proof.context -> Proof.method
+  val finite_guess_meth : Method.src -> Proof.context -> Proof.method
+  val finite_guess_meth_debug : Method.src -> Proof.context -> Proof.method
+  val fresh_guess_meth : Method.src -> Proof.context -> Proof.method
+  val fresh_guess_meth_debug : Method.src -> Proof.context -> Proof.method
 end
 
 structure NominalPermeq : NOMINAL_PERMEQ =
 struct
 
+(* some lemmas needed below *)
 val finite_emptyI = thm "finite.emptyI";
-val finite_Un = thm "finite_Un";
+val finite_Un     = thm "finite_Un";
+val conj_absorb   = thm "conj_absorb";
+val not_false     = thm "not_False_eq_True"
+val perm_fun_def  = thm "Nominal.perm_fun_def"
+val swap_fun      = thm "Nominal.swap_fun" (* FIXME: is this still needed? *)
+val perm_eq_app   = thm "Nominal.pt_fun_app_eq"
+val supports_def  = thm "Nominal.op supports_def";
+val fresh_def     = thm "Nominal.fresh_def";
+val fresh_prod    = thm "Nominal.fresh_prod";
+val fresh_unit    = thm "Nominal.fresh_unit";
+val supports_rule = thm "supports_finite";
+val supp_prod     = thm "supp_prod";
+val supp_unit     = thm "supp_unit";
+val pt_perm_compose_aux = thm "pt_perm_compose_aux";
+val cp1_aux             = thm "cp1_aux";
+val perm_aux_fold       = thm "perm_aux_fold"; 
+val supports_fresh_rule = thm "supports_fresh";
 
 (* pulls out dynamically a thm via the proof state *)
 fun dynamic_thms st name = PureThy.get_thms (theory_of_thm st) (Name name);
-fun dynamic_thm st name = PureThy.get_thm (theory_of_thm st) (Name name);
+fun dynamic_thm  st name = PureThy.get_thm  (theory_of_thm st) (Name name);
+
 
-(* a tactic simplyfying permutations *)
-val perm_fun_def = thm "Nominal.perm_fun_def"
-val perm_eq_app = thm "Nominal.pt_fun_app_eq"
+(* needed in the process of fully simplifying permutations *)
+val strong_congs = [thm "if_cong"]
+(* needed to avoid warnings about overwritten congs *)
+val weak_congs   = [thm "if_weak_cong"]
+
+
+(* debugging *)
+fun DEBUG_tac (msg,tac) = 
+    CHANGED (EVERY [print_tac ("before "^msg), tac, print_tac ("after "^msg)]); 
+fun NO_DEBUG_tac (_,tac) = CHANGED tac; 
+
 
-fun perm_eval_tac ss i = ("general simplification step", fn st =>
-    let
-        fun perm_eval_simproc sg ss redex =
-        let 
-	   (* the "application" case below is only applicable when the head   *)
-           (* of f is not a constant  or when it is a permuattion with two or *) 
-           (* more arguments                                                  *)
-           fun applicable t = 
-	       (case (strip_comb t) of
-		  (Const ("Nominal.perm",_),ts) => (length ts) >= 2
-		| (Const _,_) => false
-		| _ => true)
-
-	in
-        (case redex of 
+(* simproc that deals with instances of permutations in front *)
+(* of applications; just adding this rule to the simplifier   *)
+(* would loop; it also needs careful tuning with the simproc  *)
+(* for functions to avoid further possibilities for looping   *)
+fun perm_simproc_app st sg ss redex =
+  let 
+    (* the "application" case is only applicable when the head of f is not a *)
+    (* constant or when (f x) is a permuation with two or more arguments     *)
+    fun applicable_app t = 
+          (case (strip_comb t) of
+	      (Const ("Nominal.perm",_),ts) => (length ts) >= 2
+            | (Const _,_) => false
+            | _ => true)
+  in
+    case redex of 
         (* case pi o (f x) == (pi o f) (pi o x)          *)
-        (* special treatment according to the head of f  *)
         (Const("Nominal.perm",
           Type("fun",[Type("List.list",[Type("*",[Type(n,_),_])]),_])) $ pi $ (f $ x)) => 
-	   (case (applicable f) of
-                false => NONE  
-              | _ => 
-		let
-		    val name = Sign.base_name n
-		    val at_inst     = dynamic_thm st ("at_"^name^"_inst")
-		    val pt_inst     = dynamic_thm st ("pt_"^name^"_inst")  
-		in SOME ((at_inst RS (pt_inst RS perm_eq_app)) RS eq_reflection) end)
-
-        (* case pi o (%x. f x) == (%x. pi o (f ((rev pi)o x))) *)
-        | (Const("Nominal.perm",_) $ pi $ (Abs _)) => SOME (perm_fun_def)
+            (if (applicable_app f) then
+              let
+                val name = Sign.base_name n
+                val at_inst     = dynamic_thm st ("at_"^name^"_inst")
+                val pt_inst     = dynamic_thm st ("pt_"^name^"_inst")  
+              in SOME ((at_inst RS (pt_inst RS perm_eq_app)) RS eq_reflection) end
+            else NONE)
+      | _ => NONE
+  end
 
-        (* no redex otherwise *) 
-        | _ => NONE) end
-
-	val perm_eval =
-	    Simplifier.simproc (Theory.sign_of (the_context ())) "perm_eval" 
-	    ["Nominal.perm pi x"] perm_eval_simproc;
+(* a simproc that deals with instances in front of functions  *)
+fun perm_simproc_fun st sg ss redex = 
+   let 
+     fun applicable_fun t =
+       (case (strip_comb t) of
+          (Abs _ ,[]) => true
+	| (Const ("Nominal.perm",_),_) => false
+        | (Const _, _) => true
+	| _ => false)
+   in
+     case redex of 
+       (* case pi o f == (%x. pi o (f ((rev pi)o x))) *)     
+       (Const("Nominal.perm",_) $ pi $ f)  => 
+          (if (applicable_fun f) then SOME (perm_fun_def) else NONE)
+      | _ => NONE
+   end
 
-      (* these lemmas are created dynamically according to the atom types *) 
-      val perm_swap        = dynamic_thms st "perm_swap"
-      val perm_fresh       = dynamic_thms st "perm_fresh_fresh"
-      val perm_bij         = dynamic_thms st "perm_bij"
-      val perm_pi_simp     = dynamic_thms st "perm_pi_simp"
-      val ss' = ss addsimps (perm_swap@perm_fresh@perm_bij@perm_pi_simp)
+(* function for simplyfying permutations *)
+fun perm_simp_gen dyn_thms ss i = 
+    ("general simplification of permutations", fn st =>
+    let
+
+       val perm_sp_fun = Simplifier.simproc (theory_of_thm st) "perm_simproc_fun" 
+	                 ["Nominal.perm pi x"] (perm_simproc_fun st);
+
+       val perm_sp_app = Simplifier.simproc (theory_of_thm st) "perm_simproc_app" 
+	                 ["Nominal.perm pi x"] (perm_simproc_app st);
+
+       val ss' = ss addsimps (List.concat (map (dynamic_thms st) dyn_thms))
+                    delcongs weak_congs
+                    addcongs strong_congs
+                    addsimprocs [perm_sp_fun, perm_sp_app]
     in
-      asm_full_simp_tac (ss' addsimprocs [perm_eval]) i st
+      asm_full_simp_tac ss' i st
     end);
 
+(* general simplification of permutations and permutation that arose from eqvt-problems *)
+val perm_simp = perm_simp_gen ["perm_swap","perm_fresh_fresh","perm_bij","perm_pi_simp"];
+val eqvt_simp = perm_simp_gen ["perm_swap","perm_fresh_fresh","perm_pi_simp","eqvt"];
+
+(* main simplification tactics for permutations *)
+(* FIXME: perm_simp_tac should simplify more permutations *)
+fun perm_simp_tac tactical ss i = DETERM (tactical (perm_simp ss i));
+fun eqvt_simp_tac tactical ss i = DETERM (tactical (eqvt_simp ss i)); 
+
+
 (* applies the perm_compose rule such that                             *)
-(*                                                                     *)
 (*   pi o (pi' o lhs) = rhs                                            *)
-(*                                                                     *)
 (* is transformed to                                                   *) 
-(*                                                                     *)
 (*  (pi o pi') o (pi' o lhs) = rhs                                     *)
 (*                                                                     *)
 (* this rule would loop in the simplifier, so some trick is used with  *)
 (* generating perm_aux'es for the outermost permutation and then un-   *)
 (* folding the definition                                              *)
-val pt_perm_compose_aux = thm "pt_perm_compose_aux";
-val cp1_aux             = thm "cp1_aux";
-val perm_aux_fold       = thm "perm_aux_fold"; 
-
 fun perm_compose_tac ss i = 
     let
 	fun perm_compose_simproc sg ss redex =
@@ -155,7 +200,7 @@
 	Simplifier.simproc (the_context()) "perm_compose" 
 	["Nominal.perm pi1 (Nominal.perm pi2 t)"] perm_compose_simproc;
 
-      val ss' = Simplifier.theory_context (the_context ()) empty_ss	  
+      val ss' = Simplifier.theory_context (the_context ()) empty_ss (* FIXME: get rid of the_context *)	  
 
     in
 	("analysing permutation compositions on the lhs",
@@ -164,19 +209,18 @@
                 asm_full_simp_tac (HOL_basic_ss addsimps [perm_aux_fold]) i])
     end
 
+
 (* applying Stefan's smart congruence tac *)
 fun apply_cong_tac i = 
     ("application of congruence",
      (fn st => DatatypeAux.cong_tac i st handle Subscript => no_tac st));
 
+
 (* unfolds the definition of permutations     *)
 (* applied to functions such that             *)
-(*                                            *)
-(*   pi o f = rhs                             *)  
-(*                                            *)
+(*     pi o f = rhs                           *)  
 (* is transformed to                          *)
-(*                                            *)
-(*   %x. pi o (f ((rev pi) o x)) = rhs        *)
+(*     %x. pi o (f ((rev pi) o x)) = rhs      *)
 fun unfold_perm_fun_def_tac i = 
     let
 	val perm_fun_def = thm "Nominal.perm_fun_def"
@@ -187,28 +231,11 @@
 
 (* applies the ext-rule such that      *)
 (*                                     *)
-(*    f = g    goes to /\x. f x = g x  *)
+(*    f = g   goes to  /\x. f x = g x  *)
 fun ext_fun_tac i = ("extensionality expansion of functions", rtac ext i);
 
-(* FIXME FIXME FIXME *)
-(* should be able to analyse pi o fresh_fun () = fresh_fun instances *) 
-fun fresh_fun_eqvt_tac i =
-    let
-	val fresh_fun_equiv = thm "Nominal.fresh_fun_equiv_ineq"
-    in
-	("fresh_fun equivariance", rtac (fresh_fun_equiv RS trans) i)
-    end		       
-		       
-(* debugging *)
-fun DEBUG_tac (msg,tac) = 
-    CHANGED (EVERY [tac, print_tac ("after "^msg)]); 
-fun NO_DEBUG_tac (_,tac) = CHANGED tac; 
 
-(* Main Tactics *)
-fun perm_simp_tac tactical ss i = 
-    DETERM (tactical (perm_eval_tac ss i));
-
-(* perm_full_simp_tac is perm_simp_tac plus additional tactics    *)
+(* perm_full_simp_tac is perm_simp plus additional tactics        *)
 (* to decide equation that come from support problems             *)
 (* since it contains looping rules the "recursion" - depth is set *)
 (* to 10 - this seems to be sufficient in most cases              *)
@@ -217,37 +244,34 @@
 	  if n=0 then K all_tac
 	  else DETERM o 
 	       (FIRST'[fn i => tactical ("splitting conjunctions on the rhs", rtac conjI i),
-                       fn i => tactical (perm_eval_tac ss i),
+                       fn i => tactical (perm_simp ss i),
 		       fn i => tactical (perm_compose_tac ss i),
 		       fn i => tactical (apply_cong_tac i), 
                        fn i => tactical (unfold_perm_fun_def_tac i),
-                       fn i => tactical (ext_fun_tac i), 
-                       fn i => tactical (fresh_fun_eqvt_tac i)]
+                       fn i => tactical (ext_fun_tac i)]
 		      THEN_ALL_NEW (TRY o (perm_full_simp_tac_aux tactical ss (n-1))))
   in perm_full_simp_tac_aux tactical ss 10 end;
 
-(* tactic that first unfolds the support definition           *)
-(* and strips off the intros, then applies perm_full_simp_tac *)
+
+(* tactic that tries to solve "supports"-goals; first it *)
+(* unfolds the support definition and strips off the     *)
+(* intros, then applies eqvt_simp_tac                    *)
 fun supports_tac tactical ss i =
   let 
-      val supports_def = thm "Nominal.op supports_def";
-      val fresh_def    = thm "Nominal.fresh_def";
-      val fresh_prod   = thm "Nominal.fresh_prod";
-      val simps        = [supports_def,symmetric fresh_def,fresh_prod]
+     val simps        = [supports_def,symmetric fresh_def,fresh_prod]
   in
       EVERY [tactical ("unfolding of supports   ", simp_tac (HOL_basic_ss addsimps simps) i),
              tactical ("stripping of foralls    ", REPEAT_DETERM (rtac allI i)),
              tactical ("geting rid of the imps  ", rtac impI i),
              tactical ("eliminating conjuncts   ", REPEAT_DETERM (etac  conjE i)),
-             tactical ("applying perm_full_simp ", perm_full_simp_tac tactical ss i
-                                                   (*perm_simp_tac tactical ss i*))]
+             tactical ("applying eqvt_simp      ", eqvt_simp_tac tactical ss i )]
   end;
 
 
-(* tactic that guesses the finite-support of a goal       *)
-(* it collects all free variables and tries to show       *)
-(* that the support of these free variables (op supports) *)
-(* the goal                                               *)
+(* tactic that guesses the finite-support of a goal        *)
+(* it first collects all free variables and tries to show  *)
+(* that the support of these free variables (op supports)  *)
+(* the goal                                                *)
 fun collect_vars i (Bound j) vs = if j < i then vs else insert (op =) (Bound (j - i)) vs
   | collect_vars i (v as Free _) vs = insert (op =) v vs
   | collect_vars i (v as Var _) vs = insert (op =) v vs
@@ -255,11 +279,6 @@
   | collect_vars i (Abs (_, _, t)) vs = collect_vars (i+1) t vs
   | collect_vars i (t $ u) vs = collect_vars i u (collect_vars i t vs);
 
-val supports_rule = thm "supports_finite";
-
-val supp_prod = thm "supp_prod";
-val supp_unit = thm "supp_unit";
-
 fun finite_guess_tac tactical ss i st =
     let val goal = List.nth(cprems_of st, i-1)
     in
@@ -280,7 +299,8 @@
               Logic.strip_assums_concl (hd (prems_of supports_rule'));
             val supports_rule'' = Drule.cterm_instantiate
               [(cert (head_of S), cert s')] supports_rule'
-            val ss' = ss addsimps [supp_prod, supp_unit, finite_Un, finite_emptyI]
+            val fin_supp = dynamic_thms st ("fin_supp")
+            val ss' = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
           in
             (tactical ("guessing of the right supports-set",
                       EVERY [compose_tac (false, supports_rule'', 2) i,
@@ -291,14 +311,16 @@
     end
     handle Subscript => Seq.empty
 
-val supports_fresh_rule = thm "supports_fresh";
-val fresh_def           = thm "Nominal.fresh_def";
-val fresh_prod          = thm "Nominal.fresh_prod";
-val fresh_unit          = thm "Nominal.fresh_unit";
-
+(* tactic that guesses whether an atom is fresh for an expression  *)
+(* it first collects all free variables and tries to show that the *) 
+(* support of these free variables (op supports) the goal          *)
 fun fresh_guess_tac tactical ss i st =
     let 
 	val goal = List.nth(cprems_of st, i-1)
+        val fin_supp = dynamic_thms st ("fin_supp")
+        val fresh_atm = dynamic_thms st ("fresh_atm")
+	val ss1 = ss addsimps [symmetric fresh_def,fresh_prod,fresh_unit,conj_absorb,not_false]@fresh_atm
+        val ss2 = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI,conj_absorb]@fin_supp
     in
       case Logic.strip_assums_concl (term_of goal) of
           _ $ (Const ("Nominal.fresh", Type ("fun", [T, _])) $ _ $ t) => 
@@ -317,37 +339,45 @@
               Logic.strip_assums_concl (hd (prems_of supports_fresh_rule'));
             val supports_fresh_rule'' = Drule.cterm_instantiate
               [(cert (head_of S), cert s')] supports_fresh_rule'
-	    val ss1 = ss addsimps [symmetric fresh_def,fresh_prod,fresh_unit]
-            val ss2 = ss addsimps [supp_prod,supp_unit,finite_Un,finite_emptyI]
-            (* FIXME sometimes these rewrite rules are already in the ss, *)
-            (* which produces a warning                                   *)
           in
-            (tactical ("guessing of the right set that supports the goal",
-                      EVERY [compose_tac (false, supports_fresh_rule'', 3) i,
+            (tactical ("guessing of the right set that supports the goal", 
+                      (EVERY [compose_tac (false, supports_fresh_rule'', 3) i,
                              asm_full_simp_tac ss1 (i+2),
                              asm_full_simp_tac ss2 (i+1), 
-                             supports_tac tactical ss i])) st
+                             supports_tac tactical ss i]))) st
           end
-        | _ => Seq.empty
+          (* when a term-constructor contains more than one binder, it is useful    *) 
+          (* in nominal_primrecs to try whether the goal can be solved by an hammer *)
+        | _ => (tactical ("if it is not of the form _\<sharp>_, then try the simplifier",   
+                          (asm_full_simp_tac (HOL_ss addsimps [fresh_prod]@fresh_atm) i))) st
     end
-    handle Subscript => Seq.empty
+    handle Subscript => Seq.empty;
+
 
-fun simp_meth_setup tac =
+(* setup so that the simpset is used which is active at the moment when the tactic is called *)
+fun local_simp_meth_setup tac =
   Method.only_sectioned_args (Simplifier.simp_modifiers' @ Splitter.split_modifiers)
-  (Method.SIMPLE_METHOD' o tac o local_simpset_of);
+  (Method.SIMPLE_METHOD' o tac o local_simpset_of) ;
 
-val perm_eq_meth            = simp_meth_setup (perm_simp_tac NO_DEBUG_tac);
-val perm_eq_meth_debug      = simp_meth_setup (perm_simp_tac DEBUG_tac);
-val perm_full_eq_meth       = simp_meth_setup (perm_full_simp_tac NO_DEBUG_tac);
-val perm_full_eq_meth_debug = simp_meth_setup (perm_full_simp_tac DEBUG_tac);
-val supports_meth           = simp_meth_setup (supports_tac NO_DEBUG_tac);
-val supports_meth_debug     = simp_meth_setup (supports_tac DEBUG_tac);
-val finite_gs_meth          = simp_meth_setup (finite_guess_tac NO_DEBUG_tac);
-val finite_gs_meth_debug    = simp_meth_setup (finite_guess_tac DEBUG_tac);
-val fresh_gs_meth           = simp_meth_setup (fresh_guess_tac NO_DEBUG_tac);
-val fresh_gs_meth_debug     = simp_meth_setup (fresh_guess_tac DEBUG_tac);
+(* uses HOL_basic_ss only *)
+fun basic_simp_meth_setup tac =
+  Method.sectioned_args 
+   (fn (ctxt,l) => ((),((Simplifier.map_ss (fn _ => HOL_basic_ss) ctxt),l)))
+   (Simplifier.simp_modifiers' @ Splitter.split_modifiers)
+   (fn _ => Method.SIMPLE_METHOD' o tac o local_simpset_of);
+
 
-(* FIXME: get rid of "debug" versions? *)
+val perm_simp_meth            = local_simp_meth_setup (perm_simp_tac NO_DEBUG_tac);
+val perm_simp_meth_debug      = local_simp_meth_setup (perm_simp_tac DEBUG_tac);
+val perm_full_simp_meth       = local_simp_meth_setup (perm_full_simp_tac NO_DEBUG_tac);
+val perm_full_simp_meth_debug = local_simp_meth_setup (perm_full_simp_tac DEBUG_tac);
+val supports_meth             = local_simp_meth_setup (supports_tac NO_DEBUG_tac);
+val supports_meth_debug       = local_simp_meth_setup (supports_tac DEBUG_tac);
+val finite_guess_meth         = basic_simp_meth_setup (finite_guess_tac NO_DEBUG_tac);
+val finite_guess_meth_debug   = basic_simp_meth_setup (finite_guess_tac DEBUG_tac);
+val fresh_guess_meth          = basic_simp_meth_setup (fresh_guess_tac NO_DEBUG_tac);
+val fresh_guess_meth_debug    = basic_simp_meth_setup (fresh_guess_tac DEBUG_tac);
+
 val perm_simp_tac = perm_simp_tac NO_DEBUG_tac;
 val perm_full_simp_tac = perm_full_simp_tac NO_DEBUG_tac;
 val supports_tac = supports_tac NO_DEBUG_tac;
--- a/src/HOL/Nominal/nominal_thmdecls.ML	Tue Mar 06 08:09:43 2007 +0100
+++ b/src/HOL/Nominal/nominal_thmdecls.ML	Tue Mar 06 15:28:22 2007 +0100
@@ -58,7 +58,7 @@
 (* the implicational case it is also checked that the variables and permutation fit  *)
 (* together, i.e. are of the right "pt_class", so that a stronger version of the     *)
 (* eqality-lemma can be derived. *)
-exception EQVT_FORM;
+exception EQVT_FORM of string;
 
 val print_data = Data.print o Context.Proof;
 val get_eqvt_thms = Context.Theory #> Data.get #> #eqvt;
@@ -69,6 +69,7 @@
 fun mk_permT T = HOLogic.listT (HOLogic.mk_prodT (T, T));
 
 val perm_bool = thm "perm_bool";
+val perm_fun_def = thm "perm_fun_def";
 
 val NOMINAL_EQVT_DEBUG = ref false;
 
@@ -77,11 +78,13 @@
     then (EVERY [tac, print_tac ("after "^msg)])
     else tac 
 
+fun dynamic_thms thy name = PureThy.get_thms thy (Name name)
+
 fun tactic_eqvt ctx orig_thm pi typi = 
     let 
         val mypi = Thm.cterm_of ctx (Var (pi,typi))
         val mypifree = Thm.cterm_of ctx (Const ("List.rev",typi --> typi) $ Free (fst pi,typi))
-        val perm_pi_simp = PureThy.get_thms ctx (Name "perm_pi_simp")
+        val perm_pi_simp = dynamic_thms ctx "perm_pi_simp"
     in
         EVERY [tactic ("iffI applied",rtac iffI 1), 
                tactic ("simplifies with orig_thm and perm_bool", 
@@ -136,7 +139,7 @@
 	     in 
 		if (Type.of_sort (Sign.tsig_of thy) (ty,[class_name])) 
                 then [(pi,typi)] 
-                else raise EQVT_FORM 
+                else raise EQVT_FORM "Could not find a permutation" 
              end 
         | Abs (_,_,t1) => get_pi_aux t1
         | (t1 $ t2) => get_pi_aux t1 @ get_pi_aux t2
@@ -147,9 +150,29 @@
     (* a singleton-list  *)
     (case (distinct (op =) (get_pi_aux t)) of 
       [(pi,typi)] => (pi,typi)
-    | _ => raise EQVT_FORM)
+    | _ => raise EQVT_FORM "All permutation should be the same")
   end;
 
+fun tactic_eqvt_simple_form thy orig_thm = 
+  let val perm_pi_simp = dynamic_thms thy "perm_pi_simp"
+      val _ = (warning "perm_pi_simp :";map print_thm perm_pi_simp)
+  in
+   tactic ("Try to prove simple eqvt form.",
+    asm_full_simp_tac (HOL_basic_ss addsimps [perm_fun_def,orig_thm]@perm_pi_simp) 1)
+  end
+
+fun derive_simple_statement thy oldlhs orig_thm pi typi typrm =
+let val goal_term =
+  case (strip_comb oldlhs) of   
+    (Const (f,t),args) =>  
+       let val lhs = Const ("Nominal.perm",Type ("fun",[typi,Type ("fun",[t,t])])) $ Var(pi,typi) $ Const (f,t)
+       in  Logic.unvarify (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,Const(f,t)))) end
+   | _ => raise EQVT_FORM "Error deriving simple version."
+  val _ = Display.print_cterm (cterm_of thy goal_term)  
+in
+  Goal.prove_global thy [] [] goal_term (fn _ =>  tactic_eqvt_simple_form thy orig_thm)
+end
+
 (* Either adds a theorem (orig_thm) to or deletes one from the equivaraince *)
 (* lemma list depending on flag. To be added the lemma has to satisfy a     *)
 (* certain form. *) 
@@ -166,20 +189,26 @@
              then 
                (warning ("equivariance lemma of the relational form");
                 [orig_thm, get_derived_thm thy hyp concl orig_thm pi typi])
-             else raise EQVT_FORM
+             else raise EQVT_FORM "Type Implication"
           end
-        (* case: eqvt-lemma is of the equational form *)  
+       (* case: eqvt-lemma is of the equational form *)  
       | (Const ("Trueprop", _) $ (Const ("op =", _) $ 
-            (Const ("Nominal.perm",_) $ Var (pi,typi) $ lhs) $ rhs)) =>
-	      (if (apply_pi lhs (pi,typi)) = rhs 
-               then [orig_thm] 
-               else raise EQVT_FORM)
-      | _ => raise EQVT_FORM)
+            (Const ("Nominal.perm",typrm) $ Var (pi,typi) $ lhs) $ rhs)) =>
+	   let val optth = ([derive_simple_statement thy lhs orig_thm pi typi typrm] 
+		handle (ERROR s) => (warning ("Couldn't prove the simple version:\n"^s);[])
+			| _ => (warning "Couldn't prove the simple version";[])
+		)
+	    in
+	   (if (apply_pi lhs (pi,typi)) = rhs 
+               then orig_thm::optth   
+               else raise EQVT_FORM "Type Equality")
+	   end
+      | _ => raise EQVT_FORM "Type unknown")
   in 
       update_context flag thms_to_be_added context
   end
-  handle EQVT_FORM => 
-      error (string_of_thm orig_thm ^ " does not comply with the form of an equivariance lemma.")
+  handle EQVT_FORM s => 
+      error (string_of_thm orig_thm ^ " does not comply with the form of an equivariance lemma ("^s^").")
 
 (* in cases of bij- and freshness, we just add the lemmas to the *)
 (* data-slot *) 
@@ -193,6 +222,7 @@
 
 fun bij_add_del_aux f   = simple_add_del_aux #bij "bij" f
 fun fresh_add_del_aux f = simple_add_del_aux #fresh "fresh" f
+fun eqvt_force_add_del_aux f = simple_add_del_aux #eqvt "eqvt" f
 
 fun eqvt_map f th (r:Data.T)  = {eqvt = (f th (#eqvt r)), fresh = #fresh r, bij = #bij r};
 fun fresh_map f th (r:Data.T) = {eqvt = #eqvt r, fresh = (f th (#fresh r)), bij = #bij r};
@@ -205,10 +235,14 @@
 val fresh_add = Thm.declaration_attribute (fresh_add_del_aux (fresh_map Drule.add_rule)); 
 val fresh_del = Thm.declaration_attribute (fresh_add_del_aux (fresh_map Drule.del_rule));
 
+val eqvt_force_add  = Thm.declaration_attribute (eqvt_force_add_del_aux (eqvt_map Drule.add_rule)); 
+val eqvt_force_del  = Thm.declaration_attribute (eqvt_force_add_del_aux (eqvt_map Drule.del_rule)); 
+
 val setup =
   Data.init #>
   Attrib.add_attributes 
      [("eqvt",  Attrib.add_del_args eqvt_add eqvt_del,   "equivariance theorem declaration"),
+      ("eqvt_force",  Attrib.add_del_args eqvt_force_add eqvt_force_del,"equivariance theorem declaration (without checking the form of the lemma)"),
       ("fresh", Attrib.add_del_args fresh_add fresh_del, "freshness theorem declaration"),
       ("bij",   Attrib.add_del_args bij_add bij_del,     "bijection theorem declaration")];