Recursive examples with relation composition
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 10 May 2010 15:54:16 +0200
changeset 2094 56b849d348ae
parent 2093 751d1349329b
child 2096 41c2445fdee4
Recursive examples with relation composition
Nominal/Ex/Ex1rec.thy
Nominal/Ex/ExLetRec.thy
--- a/Nominal/Ex/Ex1rec.thy	Mon May 10 15:45:04 2010 +0200
+++ b/Nominal/Ex/Ex1rec.thy	Mon May 10 15:54:16 2010 +0200
@@ -5,6 +5,7 @@
 atom_decl name
 
 ML {* val _ = cheat_supp_eq := true *}
+ML {* val _ = cheat_equivp := true *}
 
 nominal_datatype lam =
   Var "name"
--- a/Nominal/Ex/ExLetRec.thy	Mon May 10 15:45:04 2010 +0200
+++ b/Nominal/Ex/ExLetRec.thy	Mon May 10 15:54:16 2010 +0200
@@ -7,7 +7,6 @@
 
 atom_decl name
 
-ML {* val _ = cheat_alpha_eqvt := true *}
 ML {* val _ = cheat_equivp := true *}
 
 nominal_datatype trm =
@@ -83,12 +82,6 @@
   apply (simp add: atom_eqvt fresh_star_def)
   done
 
-(*
-declare permute_trm_raw_permute_lts_raw.simps[eqvt]
-declare alpha_gen_eqvt[eqvt]
-equivariance alpha_trm_raw
-*)
-
 end