--- a/src/HOL/Nominal/Nominal.thy Mon Apr 23 20:44:12 2007 +0200
+++ b/src/HOL/Nominal/Nominal.thy Tue Apr 24 14:01:23 2007 +0200
@@ -634,6 +634,11 @@
shows "pi1 \<triangleq> pi2 \<Longrightarrow> (rev pi1) \<triangleq> (rev pi2)"
by (simp add: at_prm_rev_eq[OF at])
+lemma at_perm_fresh_fresh:
+ fixes a :: "'x"
+ assumes
+
+
lemma at_ds1:
fixes a :: "'x"
assumes at: "at TYPE('x)"
--- a/src/HOL/Nominal/nominal_fresh_fun.ML Mon Apr 23 20:44:12 2007 +0200
+++ b/src/HOL/Nominal/nominal_fresh_fun.ML Tue Apr 24 14:01:23 2007 +0200
@@ -9,11 +9,13 @@
(* First some functions that could be
in the library *)
-(* A tactical which implements Coq's T;[a,b,c] *)
+(* A tactical which applies a list of int -> tactic to the corresponding subgoals present after the application of another tactic.
+T THENL [A,B,C] is equivalent to T THEN (C 3 THEN B 2 THEN A 1)
+*)
infix 1 THENL
-fun THENL (tac,tacs) =
+fun (op THENL) (tac,tacs) =
tac THEN
(EVERY (map (fn (tac,i) => tac i) (rev tacs ~~ (length tacs downto 1))))