Link calls to Raw permutations, FV definition and alpha_definition into the parser.
authorCezary Kaliszyk <kaliszyk@in.tum.de>
Tue, 02 Mar 2010 08:42:10 +0100
changeset 1295 0ecc775e5fce
parent 1294 db1b5cb89aa4
child 1296 790940e90db2
child 1297 0ab16694c3c1
Link calls to Raw permutations, FV definition and alpha_definition into the parser.
Nominal/Parser.thy
Nominal/Test.thy
--- a/Nominal/Parser.thy	Tue Mar 02 06:43:09 2010 +0100
+++ b/Nominal/Parser.thy	Tue Mar 02 08:42:10 2010 +0100
@@ -1,5 +1,5 @@
 theory Parser
-imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp"
+imports "Nominal2_Atoms" "Nominal2_Eqvt" "Nominal2_Supp" "Perm" "Fv"
 begin
 
 atom_decl name
@@ -212,10 +212,16 @@
 ML {* 
 fun nominal_datatype2 dts bn_funs bn_eqs binds lthy =
 let
-  val (((raw_dt_names, (raw_bn_funs, raw_bn_eqs)), raw_binds), lthy') =
+  val (((raw_dt_names, (raw_bn_funs, raw_bn_eqs)), raw_binds), lthy2) =
     raw_nominal_decls dts bn_funs bn_eqs binds lthy
+  val dtinfo = Datatype.the_info (ProofContext.theory_of lthy2) (hd raw_dt_names);
+  val ((raw_perm_def, raw_perm_simps, perms), lthy3) =
+    Local_Theory.theory_result (define_raw_perms dtinfo (length dts)) lthy2;
+  val raw_binds_flat = map (map flat) raw_binds;
+  val (((fv_ts_loc, fv_def_loc), alpha), lthy4) = define_fv_alpha dtinfo raw_binds_flat lthy3;
+
 in
-  ((raw_dt_names, raw_bn_funs, raw_bn_eqs, raw_binds), lthy')
+  ((raw_dt_names, raw_bn_funs, raw_bn_eqs, raw_binds), lthy4)
 end
 *}
 
--- a/Nominal/Test.thy	Tue Mar 02 06:43:09 2010 +0100
+++ b/Nominal/Test.thy	Tue Mar 02 08:42:10 2010 +0100
@@ -11,9 +11,9 @@
 and bp = 
   BP "name" "lam" 
 binder
-  bi::"bp \<Rightarrow> name set"
+  bi::"bp \<Rightarrow> atom set"
 where
-  "bi (BP x t) = {x}"
+  "bi (BP x t) = {atom x}"
 
 typ lam_raw
 term VAR_raw
@@ -21,6 +21,8 @@
 term LET_raw
 term Test.BP_raw
 thm bi_raw.simps
+thm permute_lam_raw_permute_bp_raw.simps
+thm alpha_lam_raw_alpha_bp_raw.intros
 
 print_theorems