# HG changeset patch # User Christian Urban # Date 1269326600 -3600 # Node ID 7cebb576fae3fb71cef5b0d546f72e31a1770f3f # Parent b6da798cef68e323df7b171aaf004499a2861a52# Parent d804729d6cf4ba9fc45210cf2a9f2670e3b53f00 merged diff -r b6da798cef68 -r 7cebb576fae3 Nominal/Abs.thy --- a/Nominal/Abs.thy Tue Mar 23 07:39:10 2010 +0100 +++ b/Nominal/Abs.thy Tue Mar 23 07:43:20 2010 +0100 @@ -673,25 +673,5 @@ apply(simp) done -lemma alpha_gen_compose_eqvt: - fixes pia - assumes b: "(g d, t) \gen (\x1 x2. R x1 x2 \ R (pi \ x1) (pi \ x2)) f pia (g e, s)" - and c: "\y. pi \ (g y) = g (pi \ y)" - and a: "\x. pi \ (f x) = f (pi \ x)" - shows "(g (pi \ d), pi \ t) \gen R f (pi \ pia) (g (pi \ e), pi \ s)" - using b - apply - - apply(simp add: alpha_gen.simps) - apply(erule conjE)+ - apply(rule conjI) - apply(rule_tac ?p1="- pi" in permute_eq_iff[THEN iffD1]) - apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt c[symmetric]) - apply(rule conjI) - apply(rule_tac ?p1="- pi" in fresh_star_permute_iff[THEN iffD1]) - apply(simp add: a[symmetric] atom_eqvt Diff_eqvt insert_eqvt set_eqvt empty_eqvt c[symmetric]) - apply(subst permute_eqvt[symmetric]) - apply(simp) - sorry - end diff -r b6da798cef68 -r 7cebb576fae3 Nominal/Term5.thy diff -r b6da798cef68 -r 7cebb576fae3 Nominal/Test.thy --- a/Nominal/Test.thy Tue Mar 23 07:39:10 2010 +0100 +++ b/Nominal/Test.thy Tue Mar 23 07:43:20 2010 +0100 @@ -2,8 +2,6 @@ imports "Parser" "../Attic/Prove" begin -text {* example 1, equivalent to example 2 from Terms *} - atom_decl name ML {* val _ = recursive := false *} @@ -89,6 +87,8 @@ then show "P c lm" by simp qed +text {* example 1, equivalent to example 2 from Terms *} + nominal_datatype lam = VAR "name" | APP "lam" "lam"