Nominal/nominal_dt_rawfuns.ML
changeset 2293 aecebd5ed424
parent 2292 d134bd4f6d1b
child 2294 72ad4e766acf
--- a/Nominal/nominal_dt_rawfuns.ML	Fri May 21 05:58:23 2010 +0100
+++ b/Nominal/nominal_dt_rawfuns.ML	Fri May 21 11:40:18 2010 +0100
@@ -123,6 +123,17 @@
     NONE => mk_supp arg
   | SOME fv => fv $ arg
 end  
+handle Option  => error "fv_map lookup "
+
+fun mk_fv_binder lthy fv_bn_map args (bn_option, i) = 
+let
+  val arg = nth args i
+in
+  case bn_option of
+    NONE => (setify lthy arg, @{term "{}::atom set"})
+  | SOME bn => (to_set (bn $ arg), the (AList.lookup (op=) fv_bn_map bn) $ arg)
+end  
+handle Option  => error "fv_bn_map lookup "
 
 fun mk_fv_bn_body fv_map fv_bn_map bn_args args i = 
 let
@@ -136,15 +147,7 @@
   | SOME (NONE) => @{term "{}::atom set"}
   | SOME (SOME bn) => the (AList.lookup (op=) fv_bn_map bn) $ arg
 end  
-
-fun mk_fv_binder lthy fv_bn_map args (bn_option, i) = 
-let
-  val arg = nth args i
-in
-  case bn_option of
-    NONE => (setify lthy arg, @{term "{}::atom set"})
-  | SOME bn => (to_set (bn $ arg), the (AList.lookup (op=) fv_bn_map bn) $ arg)
-end  
+handle Option  => error "fv_map/fv_bn_map lookup "
 
 fun mk_fv_rhs lthy fv_map fv_bn_map args (BC (_, binders, bodies)) =
 let
@@ -154,16 +157,13 @@
   fold_union (mk_diff (fold_union t1, fold_union t2)::t3)
 end
 
+(* in case of fv_bn we have to treat the case special, where an
+   "empty" binding clause is given *)
 fun mk_fv_bn_rhs lthy fv_map fv_bn_map bn_args args bclause =
   case bclause of
     BC (_, [], bodies) => fold_union (map (mk_fv_bn_body fv_map fv_bn_map bn_args args) bodies)
-  | BC (_, binders, bodies) => 
-      let
-        val t1 = map (mk_fv_body fv_map args) bodies
-        val (t2, t3) = split_list (map (mk_fv_binder lthy fv_bn_map args) binders)
-      in 
-        fold_union (mk_diff (fold_union t1, fold_union t2)::t3)
-      end
+  | _ => mk_fv_rhs lthy fv_map fv_bn_map args bclause
+
 
 fun mk_fv_eq lthy fv_map fv_bn_map (constr, ty, arg_tys) bclauses = 
 let
@@ -223,14 +223,15 @@
 
   val constrs_info = all_dtyp_constrs_types dt_descr sorts
 
-  val fv_eqs2 = map2 (map2 (mk_fv_eq lthy fv_map fv_bn_map2)) constrs_info bclausesss 
+  val fv_eqs2 = map2 (map2 (mk_fv_eq lthy fv_map fv_bn_map3)) constrs_info bclausesss 
   val fv_bn_eqs2 = map (mk_fv_bn_eqs lthy fv_map fv_bn_map3 constrs_info bclausesss) bn_funs2
   
-  val _ = tracing ("functions to be defined\n " ^ @{make_string} (fv_names @ fv_bn_names2))
-  val _ = tracing ("equations\n " ^ cat_lines (map (Syntax.string_of_term lthy) (flat fv_eqs2 @ flat fv_bn_eqs2)))
+  val _ = tracing ("functions to be defined\n " ^ @{make_string} (t1 @ fv_names @ fv_bn_names2))
+  val _ = tracing ("equations\n " ^ 
+    cat_lines (map (Syntax.string_of_term lthy) (t2 @ flat fv_eqs2 @ flat fv_bn_eqs2)))
 
-  val all_fv_names = map (fn s => (Binding.name s, NONE, NoSyn)) (fv_names @ fv_bn_names2)
-  val all_fv_eqs = map (pair Attrib.empty_binding) (flat fv_eqs2 @ flat fv_bn_eqs2)
+  val all_fv_names = map (fn s => (Binding.name s, NONE, NoSyn)) (t1 @ fv_names @ fv_bn_names2)
+  val all_fv_eqs = map (pair Attrib.empty_binding) (t2 @ flat fv_eqs2 @ flat fv_bn_eqs2)
 
   fun pat_completeness_auto lthy =
     Pat_Completeness.pat_completeness_tac lthy 1