Use fs typeclass in showing finite support + some cheat cleaning.
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Fri, 19 Mar 2010 14:54:30 +0100
changeset 1547 57f7af5d7564
parent 1544 c6849a634582
child 1548 708cd99307a2
Use fs typeclass in showing finite support + some cheat cleaning.
Nominal/Fv.thy
Nominal/LFex.thy
Nominal/Parser.thy
Nominal/TySch.thy
--- a/Nominal/Fv.thy	Fri Mar 19 12:28:35 2010 +0100
+++ b/Nominal/Fv.thy	Fri Mar 19 14:54:30 2010 +0100
@@ -908,7 +908,7 @@
 fun fs_tac induct supports = ind_tac induct THEN_ALL_NEW (
   rtac @{thm supports_finite} THEN' resolve_tac supports) THEN_ALL_NEW
   asm_full_simp_tac (HOL_ss addsimps @{thms supp_atom supp_atom_image supp_fset_to_set
-    supp_fmap_atom finite_insert finite.emptyI finite_Un})
+    supp_fmap_atom finite_insert finite.emptyI finite_Un finite_supp})
 *}
 
 ML {*
--- a/Nominal/LFex.thy	Fri Mar 19 12:28:35 2010 +0100
+++ b/Nominal/LFex.thy	Fri Mar 19 14:54:30 2010 +0100
@@ -8,8 +8,6 @@
 ML {* val _ = cheat_fv_rsp := false *}
 ML {* val _ = cheat_const_rsp := false *}
 ML {* val _ = cheat_equivp := false *}
-ML {* val _ = cheat_fv_eqvt := false *}
-ML {* val _ = cheat_alpha_eqvt := false *}
 
 nominal_datatype kind =
     Type
--- a/Nominal/Parser.thy	Fri Mar 19 12:28:35 2010 +0100
+++ b/Nominal/Parser.thy	Fri Mar 19 14:54:30 2010 +0100
@@ -271,9 +271,9 @@
 
 ML {* val cheat_equivp = Unsynchronized.ref true *}
 
-(* Fixes for these 2 are known *)
-ML {* val cheat_fv_eqvt = Unsynchronized.ref true *} (* The tactic works, building the goal needs fixing *)
-ML {* val cheat_alpha_eqvt = Unsynchronized.ref true *} (* The tactic works, building the goal needs fixing *)
+(* These 2 are not needed any more *)
+ML {* val cheat_fv_eqvt = Unsynchronized.ref false *}
+ML {* val cheat_alpha_eqvt = Unsynchronized.ref false *}
 
 
 ML {*
@@ -408,8 +408,8 @@
   val (_, lthy20) = Local_Theory.note ((Binding.empty,
     [Attrib.internal (fn _ => Nominal_ThmDecls.eqvt_add)]), q_eqvt) lthy19;
   val _ = tracing "Finite Support";
-  val supports = map (prove_supports lthy20 q_perm) consts handle _ => []
-  val fin_supp = HOLogic.conj_elims (prove_fs lthy20 q_induct supports q_tys) handle _ => []
+  val supports = map (prove_supports lthy20 q_perm) consts;
+  val fin_supp = HOLogic.conj_elims (prove_fs lthy20 q_induct supports q_tys);
   val thy3 = Local_Theory.exit_global lthy20;
   val lthy21 = Theory_Target.instantiation (qty_full_names, [], @{sort fs}) thy3;
   fun tac _ = Class.intro_classes_tac [] THEN (ALLGOALS (resolve_tac fin_supp))
--- a/Nominal/TySch.thy	Fri Mar 19 12:28:35 2010 +0100
+++ b/Nominal/TySch.thy	Fri Mar 19 14:54:30 2010 +0100
@@ -7,8 +7,6 @@
 ML {* val _ = cheat_fv_rsp := false *}
 ML {* val _ = cheat_const_rsp := false *}
 ML {* val _ = cheat_equivp := false *}
-ML {* val _ = cheat_fv_eqvt := false *}
-ML {* val _ = cheat_alpha_eqvt := false *}
 
 nominal_datatype t =
   Var "name"