Nominal-General/Nominal2_Eqvt.thy
changeset 1952 27cdc0a3a763
parent 1947 51f411b1197d
child 1953 186d8486dfd5
--- a/Nominal-General/Nominal2_Eqvt.thy	Wed Apr 21 12:25:52 2010 +0200
+++ b/Nominal-General/Nominal2_Eqvt.thy	Mon Apr 26 10:01:13 2010 +0200
@@ -6,7 +6,7 @@
     (contains many, but not all such lemmas).
 *)
 theory Nominal2_Eqvt
-imports Nominal2_Base Nominal2_Atoms
+imports Nominal2_Base 
 uses ("nominal_thmdecls.ML")
      ("nominal_permeq.ML")
      ("nominal_eqvt.ML")
@@ -249,7 +249,7 @@
   imp_eqvt [folded induct_implies_def]
 
   (* nominal *)
-  supp_eqvt fresh_eqvt
+  supp_eqvt fresh_eqvt add_perm_eqvt
 
   (* datatypes *)
   permute_prod.simps append_eqvt rev_eqvt set_eqvt
@@ -259,8 +259,7 @@
   empty_eqvt UNIV_eqvt union_eqvt inter_eqvt mem_eqvt
   Diff_eqvt Compl_eqvt insert_eqvt Collect_eqvt image_eqvt
 
-  atom_eqvt add_perm_eqvt
-
+ 
 lemmas [eqvt_raw] =
   permute_eqvt_raw[THEN eq_reflection] (* the normal version of this lemma loops *)
 
@@ -288,22 +287,13 @@
 use "nominal_permeq.ML"
 setup Nominal_Permeq.setup
 
-ML {*
-val test1 = Scan.optional (Scan.lift (Args.add -- Args.colon) |-- Attrib.thms) [];
-val test2 = Scan.optional (Scan.lift ((Args.$$$ "exclude") -- Args.colon) |-- 
-  (Scan.repeat (Args.const true))) []
-*}
-
 method_setup perm_simp =
- {* test1 -- test2 >> 
-    (fn (thms, consts) => fn ctxt => SIMPLE_METHOD (HEADGOAL (Nominal_Permeq.eqvt_tac ctxt [] consts))) *}
- {* pushes permutations inside *}
+ {* Nominal_Permeq.args_parser >> Nominal_Permeq.perm_simp_meth *}
+ {* pushes permutations inside. *}
 
 method_setup perm_strict_simp =
- {* test1 -- test2 >> 
-    (fn (thms, consts) => fn ctxt => SIMPLE_METHOD (HEADGOAL 
-       (Nominal_Permeq.eqvt_strict_tac ctxt thms consts))) *}
- {* pushes permutations inside, raises an error if it cannot solve all permutations *}
+ {* Nominal_Permeq.args_parser >> Nominal_Permeq.perm_strict_simp_meth *}
+ {* pushes permutations inside, raises an error if it cannot solve all permutations. *}
 
 declare [[trace_eqvt = true]]