merged
authorhaftmann
Thu, 02 Sep 2010 11:42:50 +0200
changeset 39033 e8b68ec3bb9c
parent 39016 caad9d509bc4 (current diff)
parent 39032 548e933b90ad (diff)
child 39034 ebeb48fd653b
merged
src/HOL/HOL.thy
--- a/src/HOL/Boogie/Tools/boogie_vcs.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_vcs.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -348,6 +348,6 @@
   let
     fun rewr (_, (t, _)) = vc_of_term (f thy t)
       |> (fn vc => (vc, (t, thm_of thy vc)))
-  in VCs.map (Option.map (fn (vcs, _) => (Symtab.map rewr vcs, g))) thy end
+  in VCs.map (Option.map (fn (vcs, _) => (Symtab.map (K rewr) vcs, g))) thy end
 
 end
--- a/src/HOL/HOL.thy	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/HOL.thy	Thu Sep 02 11:42:50 2010 +0200
@@ -1931,6 +1931,10 @@
 code_reserved Scala
   Boolean
 
+code_modulename SML Pure HOL
+code_modulename OCaml Pure HOL
+code_modulename Haskell Pure HOL
+
 text {* using built-in Haskell equality *}
 
 code_class equal
--- a/src/HOL/Imperative_HOL/Heap_Monad.thy	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Thu Sep 02 11:42:50 2010 +0200
@@ -578,7 +578,7 @@
             (map o pairself) imp_monad_bind pats),
               imp_monad_bind t0);
 
-  in (Graph.map_nodes o map_terms_stmt) imp_monad_bind end;
+  in (Graph.map o K o map_terms_stmt) imp_monad_bind end;
 
 in
 
--- a/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Library/Sum_Of_Squares/sum_of_squares.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -124,10 +124,10 @@
 fun vector_cmul c (v:vector) =
  let val n = dim v
  in if c =/ rat_0 then vector_0 n
-    else (n,FuncUtil.Intfunc.map (fn x => c */ x) (snd v))
+    else (n,FuncUtil.Intfunc.map (fn _ => fn x => c */ x) (snd v))
  end;
 
-fun vector_neg (v:vector) = (fst v,FuncUtil.Intfunc.map Rat.neg (snd v)) :vector;
+fun vector_neg (v:vector) = (fst v,FuncUtil.Intfunc.map (K Rat.neg) (snd v)) :vector;
 
 fun vector_add (v1:vector) (v2:vector) =
  let val m = dim v1
@@ -167,11 +167,11 @@
 fun matrix_cmul c (m:matrix) =
  let val (i,j) = dimensions m
  in if c =/ rat_0 then matrix_0 (i,j)
-    else ((i,j),FuncUtil.Intpairfunc.map (fn x => c */ x) (snd m))
+    else ((i,j),FuncUtil.Intpairfunc.map (fn _ => fn x => c */ x) (snd m))
  end;
 
 fun matrix_neg (m:matrix) =
-  (dimensions m, FuncUtil.Intpairfunc.map Rat.neg (snd m)) :matrix;
+  (dimensions m, FuncUtil.Intpairfunc.map (K Rat.neg) (snd m)) :matrix;
 
 fun matrix_add (m1:matrix) (m2:matrix) =
  let val d1 = dimensions m1
@@ -229,14 +229,14 @@
 
 fun monomial_pow m k =
   if k = 0 then monomial_1
-  else FuncUtil.Ctermfunc.map (fn x => k * x) m;
+  else FuncUtil.Ctermfunc.map (fn _ => fn x => k * x) m;
 
 fun monomial_divides m1 m2 =
   FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => FuncUtil.Ctermfunc.tryapplyd m2 x 0 >= k andalso a) m1 true;;
 
 fun monomial_div m1 m2 =
  let val m = FuncUtil.Ctermfunc.combine Integer.add
-   (fn x => x = 0) m1 (FuncUtil.Ctermfunc.map (fn x => ~ x) m2)
+   (fn x => x = 0) m1 (FuncUtil.Ctermfunc.map (fn _ => fn x => ~ x) m2)
  in if FuncUtil.Ctermfunc.fold (fn (x, k) => fn a => k >= 0 andalso a) m true then m
   else error "monomial_div: non-divisible"
  end;
@@ -270,9 +270,9 @@
 
 fun poly_cmul c p =
   if c =/ rat_0 then poly_0
-  else FuncUtil.Monomialfunc.map (fn x => c */ x) p;
+  else FuncUtil.Monomialfunc.map (fn _ => fn x => c */ x) p;
 
-fun poly_neg p = FuncUtil.Monomialfunc.map Rat.neg p;;
+fun poly_neg p = FuncUtil.Monomialfunc.map (K Rat.neg) p;;
 
 fun poly_add p1 p2 =
   FuncUtil.Monomialfunc.combine (curry op +/) (fn x => x =/ rat_0) p1 p2;
@@ -282,7 +282,7 @@
 fun poly_cmmul (c,m) p =
  if c =/ rat_0 then poly_0
  else if FuncUtil.Ctermfunc.is_empty m
-      then FuncUtil.Monomialfunc.map (fn d => c */ d) p
+      then FuncUtil.Monomialfunc.map (fn _ => fn d => c */ d) p
       else FuncUtil.Monomialfunc.fold (fn (m', d) => fn a => (FuncUtil.Monomialfunc.update (monomial_mul m m', c */ d) a)) p poly_0;
 
 fun poly_mul p1 p2 =
@@ -596,13 +596,13 @@
  let
   val cd1 = fold_rev (common_denominator FuncUtil.Intpairfunc.fold) mats (rat_1)
   val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj)  (rat_1)
-  val mats' = map (FuncUtil.Intpairfunc.map (fn x => cd1 */ x)) mats
+  val mats' = map (FuncUtil.Intpairfunc.map (fn _ => fn x => cd1 */ x)) mats
   val obj' = vector_cmul cd2 obj
   val max1 = fold_rev (maximal_element FuncUtil.Intpairfunc.fold) mats' (rat_0)
   val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') (rat_0)
   val scal1 = pow2 (20 - trunc(Math.ln (float_of_rat max1) / Math.ln 2.0))
   val scal2 = pow2 (20 - trunc(Math.ln (float_of_rat max2) / Math.ln 2.0))
-  val mats'' = map (FuncUtil.Intpairfunc.map (fn x => x */ scal1)) mats'
+  val mats'' = map (FuncUtil.Intpairfunc.map (fn _ => fn x => x */ scal1)) mats'
   val obj'' = vector_cmul scal2 obj'
  in solver obj'' mats''
   end
@@ -629,13 +629,13 @@
  let
   val cd1 = fold_rev (common_denominator Inttriplefunc.fold) mats (rat_1)
   val cd2 = common_denominator FuncUtil.Intfunc.fold (snd obj)  (rat_1)
-  val mats' = map (Inttriplefunc.map (fn x => cd1 */ x)) mats
+  val mats' = map (Inttriplefunc.map (fn _ => fn x => cd1 */ x)) mats
   val obj' = vector_cmul cd2 obj
   val max1 = fold_rev (maximal_element Inttriplefunc.fold) mats' (rat_0)
   val max2 = maximal_element FuncUtil.Intfunc.fold (snd obj') (rat_0)
   val scal1 = pow2 (20 - int_of_float(Math.ln (float_of_rat max1) / Math.ln 2.0))
   val scal2 = pow2 (20 - int_of_float(Math.ln (float_of_rat max2) / Math.ln 2.0))
-  val mats'' = map (Inttriplefunc.map (fn x => x */ scal1)) mats'
+  val mats'' = map (Inttriplefunc.map (fn _ => fn x => x */ scal1)) mats'
   val obj'' = vector_cmul scal2 obj'
  in solver obj'' mats''
   end
@@ -655,7 +655,7 @@
 (* Stuff for "equations" ((int*int*int)->num functions).                         *)
 
 fun tri_equation_cmul c eq =
-  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn d => c */ d) eq;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn _ => fn d => c */ d) eq;
 
 fun tri_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
 
@@ -686,7 +686,7 @@
      in if b =/ rat_0 then e else
         tri_equation_add e (tri_equation_cmul (Rat.neg b // a) eq)
      end
-   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map elim dun)) (map elim oeqs)
+   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map (K elim) dun)) (map elim oeqs)
    end)
   handle Failure _ => eliminate vs dun eqs)
 in
@@ -724,7 +724,7 @@
          in if b =/ rat_0 then e
             else tri_equation_add e (tri_equation_cmul (Rat.neg b // a) eq)
          end
-    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map elim dun))
+    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map (K elim) dun))
                  (map elim oeqs)
     end
 in fn eqs =>
@@ -744,7 +744,7 @@
             (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
   val ass =
     Inttriplefunc.combine (curry op +/) (K false)
-    (Inttriplefunc.map (tri_equation_eval vfn) assigs) vfn
+    (Inttriplefunc.map (K (tri_equation_eval vfn)) assigs) vfn
  in if forall (fn e => tri_equation_eval ass e =/ rat_0) eqs
     then Inttriplefunc.delete_safe one ass else raise Sanity
  end;
@@ -762,7 +762,7 @@
 (* Usual operations on equation-parametrized poly.                           *)
 
 fun tri_epoly_cmul c l =
-  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (tri_equation_cmul c) l;;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (K (tri_equation_cmul c)) l;;
 
 val tri_epoly_neg = tri_epoly_cmul (Rat.rat_of_int ~1);
 
@@ -773,7 +773,7 @@
 (* Stuff for "equations" ((int*int)->num functions).                         *)
 
 fun pi_equation_cmul c eq =
-  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn d => c */ d) eq;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (fn _ => fn d => c */ d) eq;
 
 fun pi_equation_add eq1 eq2 = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0) eq1 eq2;
 
@@ -804,7 +804,7 @@
      in if b =/ rat_0 then e else
         pi_equation_add e (pi_equation_cmul (Rat.neg b // a) eq)
      end
-   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map elim dun)) (map elim oeqs)
+   in eliminate vs (Inttriplefunc.update (v,eq') (Inttriplefunc.map (K elim) dun)) (map elim oeqs)
    end
   handle Failure _ => eliminate vs dun eqs
 in
@@ -842,7 +842,7 @@
          in if b =/ rat_0 then e
             else pi_equation_add e (pi_equation_cmul (Rat.neg b // a) eq)
          end
-    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map elim dun))
+    in eliminate (Inttriplefunc.update(v, eq') (Inttriplefunc.map (K elim) dun))
                  (map elim oeqs)
     end
 in fn eqs =>
@@ -862,7 +862,7 @@
             (Inttriplefunc.onefunc(one, Rat.rat_of_int ~1))
   val ass =
     Inttriplefunc.combine (curry op +/) (K false)
-    (Inttriplefunc.map (pi_equation_eval vfn) assigs) vfn
+    (Inttriplefunc.map (K (pi_equation_eval vfn)) assigs) vfn
  in if forall (fn e => pi_equation_eval ass e =/ rat_0) eqs
     then Inttriplefunc.delete_safe one ass else raise Sanity
  end;
@@ -880,7 +880,7 @@
 (* Usual operations on equation-parametrized poly.                           *)
 
 fun pi_epoly_cmul c l =
-  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (pi_equation_cmul c) l;;
+  if c =/ rat_0 then Inttriplefunc.empty else Inttriplefunc.map (K (pi_equation_cmul c)) l;;
 
 val pi_epoly_neg = pi_epoly_cmul (Rat.rat_of_int ~1);
 
@@ -1035,7 +1035,7 @@
 val bmatrix_add = Inttriplefunc.combine (curry op +/) (fn x => x =/ rat_0);
 fun bmatrix_cmul c bm =
   if c =/ rat_0 then Inttriplefunc.empty
-  else Inttriplefunc.map (fn x => c */ x) bm;
+  else Inttriplefunc.map (fn _ => fn x => c */ x) bm;
 
 val bmatrix_neg = bmatrix_cmul (Rat.rat_of_int ~1);
 fun bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);;
--- a/src/HOL/Library/positivstellensatz.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Library/positivstellensatz.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -549,7 +549,7 @@
 (* A linear arithmetic prover *)
 local
   val linear_add = FuncUtil.Ctermfunc.combine (curry op +/) (fn z => z =/ Rat.zero)
-  fun linear_cmul c = FuncUtil.Ctermfunc.map (fn x => c */ x)
+  fun linear_cmul c = FuncUtil.Ctermfunc.map (fn _ => fn x => c */ x)
   val one_tm = @{cterm "1::real"}
   fun contradictory p (e,_) = ((FuncUtil.Ctermfunc.is_empty e) andalso not(p Rat.zero)) orelse
      ((eq_set (op aconvc) (FuncUtil.Ctermfunc.dom e, [one_tm])) andalso
--- a/src/HOL/Multivariate_Analysis/normarith.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/normarith.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -32,16 +32,16 @@
                       (Thm.dest_arg t) acc
       | _ => augment_norm true t acc
 
- val cterm_lincomb_neg = FuncUtil.Ctermfunc.map Rat.neg
+ val cterm_lincomb_neg = FuncUtil.Ctermfunc.map (K Rat.neg)
  fun cterm_lincomb_cmul c t =
-    if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn x => x */ c) t
+    if c =/ Rat.zero then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn _ => fn x => x */ c) t
  fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
  fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r)
  fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r)
 
- val int_lincomb_neg = FuncUtil.Intfunc.map Rat.neg
+ val int_lincomb_neg = FuncUtil.Intfunc.map (K Rat.neg)
  fun int_lincomb_cmul c t =
-    if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn x => x */ c) t
+    if c =/ Rat.zero then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn _ => fn x => x */ c) t
  fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +/) (fn x => x =/ Rat.zero) l r
  fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r)
  fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r)
--- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -310,7 +310,7 @@
 val predfun_neg_intro_of = #neg_intro ooo the_predfun_data
 
 val intros_graph_of =
-  Graph.map_nodes (#intros o rep_pred_data) o PredData.get o ProofContext.theory_of
+  Graph.map (K (#intros o rep_pred_data)) o PredData.get o ProofContext.theory_of
 
 (* diagnostic display functions *)
 
--- a/src/HOL/Tools/SMT/smt_monomorph.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_monomorph.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -133,7 +133,7 @@
     fun complete (vT as (v, _)) subst =
       subst
       |> not (Vartab.defined subst v) ? Vartab.update vT
-      |> Vartab.map (apsnd (Term.map_atyps (replace vT)))
+      |> Vartab.map (K (apsnd (Term.map_atyps (replace vT))))
 
     fun cert (ix, (S, T)) = pairself (Thm.ctyp_of thy) (TVar (ix, S), T)
 
--- a/src/HOL/Tools/SMT/z3_proof_parser.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_parser.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -191,7 +191,7 @@
     fun cert @{term True} = @{cterm "~False"}
       | cert t = certify ctxt' t
 
-  in (typs, Symtab.map cert terms, Inttab.empty, Inttab.empty, [], ctxt') end
+  in (typs, Symtab.map (K cert) terms, Inttab.empty, Inttab.empty, [], ctxt') end
 
 fun fresh_name n (typs, terms, exprs, steps, vars, ctxt) =
   let val (n', ctxt') = yield_singleton Variable.variant_fixes n ctxt
--- a/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -823,7 +823,7 @@
 
     fun result (p, (cx, _)) = (thm_of p, cx)
   in
-    (fn (idx, ptab) => result (lookup idx (ctxt, Inttab.map Unproved ptab)))
+    (fn (idx, ptab) => result (lookup idx (ctxt, Inttab.map (K Unproved) ptab)))
   end
 
 fun reconstruct (output, {typs, terms, unfolds, assms}) ctxt =
--- a/src/Pure/General/graph.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/General/graph.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -21,7 +21,7 @@
   val maximals: 'a T -> key list
   val subgraph: (key -> bool) -> 'a T -> 'a T
   val get_entry: 'a T -> key -> 'a * (key list * key list)            (*exception UNDEF*)
-  val map_nodes: ('a -> 'b) -> 'a T -> 'b T
+  val map: (key -> 'a -> 'b) -> 'a T -> 'b T
   val get_node: 'a T -> key -> 'a                                     (*exception UNDEF*)
   val map_node: key -> ('a -> 'a) -> 'a T -> 'a T
   val map_node_yield: key -> ('a -> 'b * 'a) -> 'a T -> 'b * 'a T
@@ -114,7 +114,7 @@
 
 (* nodes *)
 
-fun map_nodes f (Graph tab) = Graph (Table.map (fn (i, ps) => (f i, ps)) tab);
+fun map_nodes f (Graph tab) = Graph (Table.map (apfst o f) tab);
 
 fun get_node G = #1 o get_entry G;
 
@@ -161,7 +161,7 @@
 fun del_nodes xs (Graph tab) =
   Graph (tab
     |> fold Table.delete xs
-    |> Table.map (fn (i, (preds, succs)) =>
+    |> Table.map (fn _ => fn (i, (preds, succs)) =>
       (i, (fold remove_key xs preds, fold remove_key xs succs))));
 
 fun del_node x (G as Graph tab) =
@@ -206,13 +206,13 @@
 fun join f (G1 as Graph tab1, G2 as Graph tab2) =
   let fun join_node key ((i1, edges1), (i2, _)) = (f key (i1, i2), edges1) in
     if pointer_eq (G1, G2) then G1
-    else fold add_edge (edges G2) (Graph (Table.join join_node (tab1, Table.map no_edges tab2)))
+    else fold add_edge (edges G2) (Graph (Table.join join_node (tab1, Table.map (K no_edges) tab2)))
   end;
 
 fun gen_merge add eq (G1 as Graph tab1, G2 as Graph tab2) =
   let fun eq_node ((i1, _), (i2, _)) = eq (i1, i2) in
     if pointer_eq (G1, G2) then G1
-    else fold add (edges G2) (Graph (Table.merge eq_node (tab1, Table.map no_edges tab2)))
+    else fold add (edges G2) (Graph (Table.merge eq_node (tab1, Table.map (K no_edges) tab2)))
   end;
 
 fun merge eq GG = gen_merge add_edge eq GG;
@@ -286,6 +286,7 @@
 
 
 (*final declarations of this structure!*)
+val map = map_nodes;
 val fold = fold_graph;
 
 end;
--- a/src/Pure/General/table.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/General/table.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -20,8 +20,7 @@
   exception UNDEF of key
   val empty: 'a table
   val is_empty: 'a table -> bool
-  val map: ('a -> 'b) -> 'a table -> 'b table
-  val map': (key -> 'a -> 'b) -> 'a table -> 'b table
+  val map: (key -> 'a -> 'b) -> 'a table -> 'b table
   val fold: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a
   val fold_rev: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a
   val dest: 'a table -> (key * 'a) list
@@ -403,8 +402,7 @@
 
 
 (*final declarations of this structure!*)
-fun map f = map_table (K f);
-val map' = map_table;
+val map = map_table;
 val fold = fold_table;
 val fold_rev = fold_rev_table;
 
--- a/src/Pure/Isar/code.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/Isar/code.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -282,7 +282,7 @@
   then NONE
   else thy
     |> (Code_Data.map o apfst)
-        ((map_functions o Symtab.map) (fn ((changed, current), history) =>
+        ((map_functions o Symtab.map) (fn _ => fn ((changed, current), history) =>
           ((false, current),
             if changed then (serial (), current) :: history else history))
         #> map_history_concluded (K true))
--- a/src/Pure/Thy/thy_info.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/Thy/thy_info.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -346,6 +346,6 @@
 
 (* finish all theories *)
 
-fun finish () = change_thys (Graph.map_nodes (fn (_, entry) => (NONE, entry)));
+fun finish () = change_thys (Graph.map (fn _ => fn (_, entry) => (NONE, entry)));
 
 end;
--- a/src/Pure/context.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/context.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -137,7 +137,7 @@
     val _ = CRITICAL (fn () => Unsynchronized.change kinds (Datatab.update (k, kind)));
   in k end;
 
-val extend_data = Datatab.map' invoke_extend;
+val extend_data = Datatab.map invoke_extend;
 
 fun merge_data pp (data1, data2) =
   Datatab.keys (Datatab.merge (K true) (data1, data2))
@@ -476,7 +476,7 @@
   | NONE => raise Fail "Invalid proof data identifier");
 
 fun init_data thy =
-  Datatab.map' (fn k => fn _ => invoke_init k thy) (! kinds);
+  Datatab.map (fn k => fn _ => invoke_init k thy) (! kinds);
 
 fun init_new_data data thy =
   Datatab.merge (K true) (data, init_data thy);
--- a/src/Pure/proofterm.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/proofterm.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -515,7 +515,7 @@
   | remove_types t = t;
 
 fun remove_types_env (Envir.Envir {maxidx, tenv, tyenv}) =
-  Envir.Envir {maxidx = maxidx, tenv = Vartab.map (apsnd remove_types) tenv, tyenv = tyenv};
+  Envir.Envir {maxidx = maxidx, tenv = Vartab.map (K (apsnd remove_types)) tenv, tyenv = tyenv};
 
 fun norm_proof' env prf = norm_proof (remove_types_env env) prf;
 
--- a/src/Pure/sorts.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Pure/sorts.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -321,7 +321,7 @@
         | NONE => NONE)
       else NONE;
     val classes' = classes |> Graph.subgraph P;
-    val arities' = arities |> Symtab.map' (map_filter o restrict_arity);
+    val arities' = arities |> Symtab.map (map_filter o restrict_arity);
   in (restrict_sort, rebuild_arities pp (make_algebra (classes', arities'))) end;
 
 
--- a/src/Tools/Code/code_ml.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Tools/Code/code_ml.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -707,73 +707,39 @@
 
 (** SML/OCaml generic part **)
 
-local
-
-datatype ml_node =
-    Dummy of string
-  | Stmt of string * ml_stmt
-  | Module of string * ((Name.context * Name.context) * ml_node Graph.T);
-
-in
-
-fun ml_node_of_program labelled_name reserved module_alias program =
+fun ml_program_of_program labelled_name reserved module_alias program =
   let
-    val reserved = Name.make_context reserved;
-    val empty_module = ((reserved, reserved), Graph.empty);
-    fun map_node [] f = f
-      | map_node (m::ms) f =
-          Graph.default_node (m, Module (m, empty_module))
-          #> Graph.map_node m (fn (Module (module_name, (nsp, nodes))) =>
-               Module (module_name, (nsp, map_node ms f nodes)));
-    fun map_nsp_yield [] f (nsp, nodes) =
-          let
-            val (x, nsp') = f nsp
-          in (x, (nsp', nodes)) end
-      | map_nsp_yield (m::ms) f (nsp, nodes) =
-          let
-            val (x, nodes') =
-              nodes
-              |> Graph.default_node (m, Module (m, empty_module))
-              |> Graph.map_node_yield m (fn Module (d_module_name, nsp_nodes) => 
-                  let
-                    val (x, nsp_nodes') = map_nsp_yield ms f nsp_nodes
-                  in (x, Module (d_module_name, nsp_nodes')) end)
-          in (x, (nsp, nodes')) end;
-    fun map_nsp_fun_yield f (nsp_fun, nsp_typ) =
+    fun namify_const upper base (nsp_const, nsp_type) =
+      let
+        val (base', nsp_const') = yield_singleton Name.variants
+          (if upper then first_upper base else base) nsp_const
+      in (base', (nsp_const', nsp_type)) end;
+    fun namify_type base (nsp_const, nsp_type) =
       let
-        val (x, nsp_fun') = f nsp_fun
-      in (x, (nsp_fun', nsp_typ)) end;
-    fun map_nsp_typ_yield f (nsp_fun, nsp_typ) =
-      let
-        val (x, nsp_typ') = f nsp_typ
-      in (x, (nsp_fun, nsp_typ')) end;
-    val mk_name_module = mk_name_module reserved NONE module_alias program;
-    fun mk_name_stmt upper name nsp =
-      let
-        val (_, base) = dest_name name;
-        val base' = if upper then first_upper base else base;
-        val ([base''], nsp') = Name.variants [base'] nsp;
-      in (base'', nsp') end;
-    fun deps_of names =
-      []
-      |> fold (fold (insert (op =)) o Graph.imm_succs program) names
-      |> subtract (op =) names
-      |> filter_out (Code_Thingol.is_case o Graph.get_node program);
+        val (base', nsp_type') = yield_singleton Name.variants base nsp_type
+      in (base', (nsp_const, nsp_type')) end;
+    fun namify_stmt (Code_Thingol.Fun _) = namify_const false
+      | namify_stmt (Code_Thingol.Datatype _) = namify_type
+      | namify_stmt (Code_Thingol.Datatypecons _) = namify_const true
+      | namify_stmt (Code_Thingol.Class _) = namify_type
+      | namify_stmt (Code_Thingol.Classrel _) = namify_const false
+      | namify_stmt (Code_Thingol.Classparam _) = namify_const false
+      | namify_stmt (Code_Thingol.Classinst _) = namify_const false;
     fun ml_binding_of_stmt (name, Code_Thingol.Fun (_, ((tysm as (vs, ty), raw_eqs), _))) =
           let
             val eqs = filter (snd o snd) raw_eqs;
-            val (eqs', is_value) = if null (filter_out (null o snd) vs) then case eqs
+            val (eqs', some_value_name) = if null (filter_out (null o snd) vs) then case eqs
                of [(([], t), some_thm)] => if (not o null o fst o Code_Thingol.unfold_fun) ty
                   then ([(([IVar (SOME "x")], t `$ IVar (SOME "x")), some_thm)], NONE)
                   else (eqs, SOME (name, member (op =) (Code_Thingol.add_constnames t []) name))
                 | _ => (eqs, NONE)
               else (eqs, NONE)
-          in (ML_Function (name, (tysm, eqs')), is_value) end
+          in (ML_Function (name, (tysm, eqs')), some_value_name) end
       | ml_binding_of_stmt (name, Code_Thingol.Classinst (stmt as ((_, (_, vs)), _))) =
           (ML_Instance (name, stmt), if forall (null o snd) vs then SOME (name, false) else NONE)
       | ml_binding_of_stmt (name, _) =
           error ("Binding block containing illegal statement: " ^ labelled_name name)
-    fun add_fun (name, stmt) =
+    fun modify_fun (name, stmt) =
       let
         val (binding, some_value_name) = ml_binding_of_stmt (name, stmt);
         val ml_stmt = case binding
@@ -784,121 +750,40 @@
              of NONE => ML_Funs ([binding], [])
               | SOME (name, true) => ML_Funs ([binding], [name])
               | SOME (name, false) => ML_Val binding
-      in
-        map_nsp_fun_yield (mk_name_stmt false name)
-        #>> (fn name' => ([name'], ml_stmt))
-      end;
-    fun add_funs stmts =
-      let
-        val ml_stmt = ML_Funs (map_split ml_binding_of_stmt stmts |> (apsnd o map_filter o Option.map) fst);
-      in
-        fold_map (fn (name, _) => map_nsp_fun_yield (mk_name_stmt false name)) stmts
-        #>> rpair ml_stmt
-      end;
-    fun add_datatypes stmts =
-      fold_map
-        (fn (name, Code_Thingol.Datatype (_, stmt)) =>
-              map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt))
-          | (name, Code_Thingol.Datatypecons _) =>
-              map_nsp_fun_yield (mk_name_stmt true name) #>> rpair NONE
-          | (name, _) =>
-              error ("Datatype block containing illegal statement: " ^ labelled_name name)
-        ) stmts
-      #>> (split_list #> apsnd (map_filter I
-        #> (fn [] => error ("Datatype block without data statement: "
-                  ^ (Library.commas o map (labelled_name o fst)) stmts)
-             | stmts => ML_Datas stmts)));
-    fun add_class stmts =
-      fold_map
-        (fn (name, Code_Thingol.Class (_, stmt)) =>
-              map_nsp_typ_yield (mk_name_stmt false name) #>> rpair (SOME (name, stmt))
-          | (name, Code_Thingol.Classrel _) =>
-              map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE
-          | (name, Code_Thingol.Classparam _) =>
-              map_nsp_fun_yield (mk_name_stmt false name) #>> rpair NONE
-          | (name, _) =>
-              error ("Class block containing illegal statement: " ^ labelled_name name)
-        ) stmts
-      #>> (split_list #> apsnd (map_filter I
-        #> (fn [] => error ("Class block without class statement: "
-                  ^ (Library.commas o map (labelled_name o fst)) stmts)
-             | [stmt] => ML_Class stmt)));
-    fun add_stmts ([stmt as (name, Code_Thingol.Fun _)]) =
-          add_fun stmt
-      | add_stmts ((stmts as (_, Code_Thingol.Fun _)::_)) =
-          add_funs stmts
-      | add_stmts ((stmts as (_, Code_Thingol.Datatypecons _)::_)) =
-          add_datatypes stmts
-      | add_stmts ((stmts as (_, Code_Thingol.Datatype _)::_)) =
-          add_datatypes stmts
-      | add_stmts ((stmts as (_, Code_Thingol.Class _)::_)) =
-          add_class stmts
-      | add_stmts ((stmts as (_, Code_Thingol.Classrel _)::_)) =
-          add_class stmts
-      | add_stmts ((stmts as (_, Code_Thingol.Classparam _)::_)) =
-          add_class stmts
-      | add_stmts ([stmt as (_, Code_Thingol.Classinst _)]) =
-          add_fun stmt
-      | add_stmts ((stmts as (_, Code_Thingol.Classinst _)::_)) =
-          add_funs stmts
-      | add_stmts stmts = error ("Illegal mutual dependencies: " ^
+      in SOME ml_stmt end;
+    fun modify_funs stmts = single (SOME
+      (ML_Funs (map_split ml_binding_of_stmt stmts |> (apsnd o map_filter o Option.map) fst)))
+    fun modify_datatypes stmts = single (SOME
+      (ML_Datas (map_filter
+        (fn (name, Code_Thingol.Datatype (_, stmt)) => SOME (name, stmt) | _ => NONE) stmts)))
+    fun modify_class stmts = single (SOME
+      (ML_Class (the_single (map_filter
+        (fn (name, Code_Thingol.Class (_, stmt)) => SOME (name, stmt) | _ => NONE) stmts))))
+    fun modify_stmts ([stmt as (name, Code_Thingol.Fun _)]) =
+          [modify_fun stmt]
+      | modify_stmts ((stmts as (_, Code_Thingol.Fun _)::_)) =
+          modify_funs stmts
+      | modify_stmts ((stmts as (_, Code_Thingol.Datatypecons _)::_)) =
+          modify_datatypes stmts
+      | modify_stmts ((stmts as (_, Code_Thingol.Datatype _)::_)) =
+          modify_datatypes stmts
+      | modify_stmts ((stmts as (_, Code_Thingol.Class _)::_)) =
+          modify_class stmts
+      | modify_stmts ((stmts as (_, Code_Thingol.Classrel _)::_)) =
+          modify_class stmts
+      | modify_stmts ((stmts as (_, Code_Thingol.Classparam _)::_)) =
+          modify_class stmts
+      | modify_stmts ([stmt as (_, Code_Thingol.Classinst _)]) =
+          [modify_fun stmt]
+      | modify_stmts ((stmts as (_, Code_Thingol.Classinst _)::_)) =
+          modify_funs stmts
+      | modify_stmts stmts = error ("Illegal mutual dependencies: " ^
           (Library.commas o map (labelled_name o fst)) stmts);
-    fun add_stmts' stmts nsp_nodes =
-      let
-        val names as (name :: names') = map fst stmts;
-        val deps = deps_of names;
-        val (module_names, _) = (split_list o map dest_name) names;
-        val module_name = (the_single o distinct (op =) o map mk_name_module) module_names
-          handle Empty =>
-            error ("Different namespace prefixes for mutual dependencies:\n"
-              ^ Library.commas (map labelled_name names)
-              ^ "\n"
-              ^ Library.commas module_names);
-        val module_name_path = Long_Name.explode module_name;
-        fun add_dep name name' =
-          let
-            val module_name' = (mk_name_module o fst o dest_name) name';
-          in if module_name = module_name' then
-            map_node module_name_path (Graph.add_edge (name, name'))
-          else let
-            val (common, (diff1 :: _, diff2 :: _)) = chop_prefix (op =)
-              (module_name_path, Long_Name.explode module_name');
-          in
-            map_node common
-              (fn node => Graph.add_edge_acyclic (diff1, diff2) node
-                handle Graph.CYCLES _ => error ("Dependency "
-                  ^ quote name ^ " -> " ^ quote name'
-                  ^ " would result in module dependency cycle"))
-          end end;
-      in
-        nsp_nodes
-        |> map_nsp_yield module_name_path (add_stmts stmts)
-        |-> (fn (base' :: bases', stmt') =>
-           apsnd (map_node module_name_path (Graph.new_node (name, (Stmt (base', stmt')))
-              #> fold2 (fn name' => fn base' =>
-                   Graph.new_node (name', (Dummy base'))) names' bases')))
-        |> apsnd (fold (fn name => fold (add_dep name) deps) names)
-        |> apsnd (fold_product (curry (map_node module_name_path o Graph.add_edge)) names names)
-      end;
-    val stmts = map (AList.make (Graph.get_node program)) (rev (Graph.strong_conn program))
-      |> filter_out (fn [(_, stmt)] => Code_Thingol.is_case stmt | _ => false);
-    val (_, nodes) = fold add_stmts' stmts empty_module;
-    fun deresolver prefix name = 
-      let
-        val module_name = (fst o dest_name) name;
-        val module_name' = (Long_Name.explode o mk_name_module) module_name;
-        val (_, (_, remainder)) = chop_prefix (op =) (prefix, module_name');
-        val stmt_name =
-          nodes
-          |> fold (fn name => fn node => case Graph.get_node node name
-              of Module (_, (_, node)) => node) module_name'
-          |> (fn node => case Graph.get_node node name of Stmt (stmt_name, _) => stmt_name
-               | Dummy stmt_name => stmt_name);
-      in
-        Long_Name.implode (remainder @ [stmt_name])
-      end handle Graph.UNDEF _ =>
-        error ("Unknown statement name: " ^ labelled_name name);
-  in (deresolver, nodes) end;
+  in
+    Code_Namespace.hierarchical_program labelled_name { module_alias = module_alias, reserved = reserved,
+      empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt,
+      cyclic_modules = false, empty_data = (), memorize_data = K I, modify_stmts = modify_stmts } program
+  end;
 
 fun serialize_ml target print_module print_stmt with_signatures { labelled_name,
   reserved_syms, includes, module_alias, class_syntax, tyco_syntax,
@@ -906,35 +791,36 @@
   let
     val is_cons = Code_Thingol.is_cons program;
     val is_presentation = not (null presentation_names);
-    val (deresolver, nodes) = ml_node_of_program labelled_name
-      reserved_syms module_alias program;
-    val reserved = make_vars reserved_syms;
-    fun print_node prefix (Dummy _) =
+    val { deresolver, hierarchical_program = ml_program } =
+      ml_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program;
+    fun print_node _ (_, Code_Namespace.Dummy) =
           NONE
-      | print_node prefix (Stmt (_, stmt)) = if is_presentation andalso
-          (null o filter (member (op =) presentation_names) o stmt_names_of) stmt
+      | print_node prefix_fragments (_, Code_Namespace.Stmt stmt) =
+          if is_presentation andalso
+            (null o filter (member (op =) presentation_names) o stmt_names_of) stmt
           then NONE
-          else SOME (print_stmt labelled_name tyco_syntax const_syntax reserved is_cons (deresolver prefix) stmt)
-      | print_node prefix (Module (module_name, (_, nodes))) =
+          else SOME (print_stmt labelled_name tyco_syntax const_syntax (make_vars reserved_syms)
+            is_cons (deresolver prefix_fragments) stmt)
+      | print_node prefix_fragments (name_fragment, Code_Namespace.Module (_, nodes)) =
           let
-            val (decls, body) = print_nodes (prefix @ [module_name]) nodes;
+            val (decls, body) = print_nodes (prefix_fragments @ [name_fragment]) nodes;
             val p = if is_presentation then Pretty.chunks2 body
-              else print_module module_name (if with_signatures then SOME decls else NONE) body;
+              else print_module name_fragment
+                (if with_signatures then SOME decls else NONE) body;
           in SOME ([], p) end
-    and print_nodes prefix nodes = (map_filter (print_node prefix o Graph.get_node nodes)
+    and print_nodes prefix_fragments nodes = (map_filter
+        (fn name => print_node prefix_fragments (name, snd (Graph.get_node nodes name)))
         o rev o flat o Graph.strong_conn) nodes
       |> split_list
       |> (fn (decls, body) => (flat decls, body))
     val names' = map (try (deresolver [])) names;
-    val p = Pretty.chunks2 (map snd includes @ snd (print_nodes [] nodes));
+    val p = Pretty.chunks2 (map snd includes @ snd (print_nodes [] ml_program));
     fun write width NONE = writeln_pretty width
       | write width (SOME p) = File.write p o string_of_pretty width;
   in
     Code_Target.serialization write (fn width => (rpair names' o string_of_pretty width)) p
   end;
 
-end; (*local*)
-
 val serializer_sml : Code_Target.serializer =
   Code_Target.parse_args (Scan.optional (Args.$$$ "no_signatures" >> K false) true
   >> (fn with_signatures => serialize_ml target_SML
--- a/src/Tools/Code/code_namespace.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Tools/Code/code_namespace.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -6,17 +6,18 @@
 
 signature CODE_NAMESPACE =
 sig
-  datatype 'a node =
+  datatype ('a, 'b) node =
       Dummy
-    | Stmt of Code_Thingol.stmt
-    | Module of ('a * (string * 'a node) Graph.T);
+    | Stmt of 'a
+    | Module of ('b * (string * ('a, 'b) node) Graph.T);
   val hierarchical_program: (string -> string) -> { module_alias: string -> string option,
-    reserved: Name.context, empty_nsp: 'b, namify_module: string -> 'b -> string * 'b,
-    namify_stmt: Code_Thingol.stmt -> string -> 'b -> string * 'b,
-    cyclic_modules: bool, empty_data: 'a, memorize_data: string -> 'a -> 'a }
+    reserved: Name.context, empty_nsp: 'c, namify_module: string -> 'c -> string * 'c,
+    namify_stmt: Code_Thingol.stmt -> string -> 'c -> string * 'c,
+    cyclic_modules: bool, empty_data: 'b, memorize_data: string -> 'b -> 'b,
+    modify_stmts: (string * Code_Thingol.stmt) list -> 'a option list }
       -> Code_Thingol.program
       -> { deresolver: string list -> string -> string,
-           hierarchical_program: (string * 'a node) Graph.T }
+           hierarchical_program: (string * ('a, 'b) node) Graph.T }
 end;
 
 structure Code_Namespace : CODE_NAMESPACE =
@@ -24,13 +25,20 @@
 
 (* hierarchical program structure *)
 
-datatype 'a node =
+datatype ('a, 'b) node =
     Dummy
-  | Stmt of Code_Thingol.stmt
-  | Module of ('a * (string * 'a node) Graph.T);
+  | Stmt of 'a
+  | Module of ('b * (string * ('a, 'b) node) Graph.T);
+
+fun map_module_content f (Module content) = Module (f content);
+
+fun map_module [] = I
+  | map_module (name_fragment :: name_fragments) =
+      apsnd o Graph.map_node name_fragment o apsnd o map_module_content
+        o map_module name_fragments;
 
 fun hierarchical_program labelled_name { module_alias, reserved, empty_nsp,
-      namify_module, namify_stmt, cyclic_modules, empty_data, memorize_data } program =
+      namify_module, namify_stmt, cyclic_modules, empty_data, memorize_data, modify_stmts } program =
   let
 
     (* building module name hierarchy *)
@@ -45,11 +53,6 @@
 
     (* building empty module hierarchy *)
     val empty_module = (empty_data, Graph.empty);
-    fun map_module f (Module content) = Module (f content);
-    fun change_module [] = I
-      | change_module (name_fragment :: name_fragments) =
-          apsnd o Graph.map_node name_fragment o apsnd o map_module
-            o change_module name_fragments;
     fun ensure_module name_fragment (data, nodes) =
       if can (Graph.get_node nodes) name_fragment then (data, nodes)
       else (data,
@@ -57,7 +60,7 @@
     fun allocate_module [] = I
       | allocate_module (name_fragment :: name_fragments) =
           ensure_module name_fragment
-          #> (apsnd o Graph.map_node name_fragment o apsnd o map_module o allocate_module) name_fragments;
+          #> (apsnd o Graph.map_node name_fragment o apsnd o map_module_content o allocate_module) name_fragments;
     val empty_program = Symtab.fold (fn (_, fragments) => allocate_module fragments)
       fragments_tab empty_module;
 
@@ -66,8 +69,7 @@
       let
         val (name_fragments, base) = dest_name name;
       in
-        change_module name_fragments (fn (data, nodes) =>
-          (memorize_data name data, Graph.new_node (name, (base, Stmt stmt)) nodes))
+        (map_module name_fragments o apsnd) (Graph.new_node (name, (base, Stmt stmt)))
       end;
     fun add_dependency name name' =
       let
@@ -75,42 +77,47 @@
         val (name_fragments', base') = dest_name name';
         val (name_fragments_common, (diff, diff')) =
           chop_prefix (op =) (name_fragments, name_fragments');
-        val (is_module, dep) = if null diff then (false, (name, name'))
-          else (true, (hd diff, hd diff'))
+        val is_module = not (null diff andalso null diff');
+        val dep = pairself hd (diff @ [name], diff' @ [name']);
         val add_edge = if is_module andalso not cyclic_modules
           then (fn node => Graph.add_edge_acyclic dep node
             handle Graph.CYCLES _ => error ("Dependency "
               ^ quote name ^ " -> " ^ quote name'
               ^ " would result in module dependency cycle"))
           else Graph.add_edge dep
-      in (change_module name_fragments_common o apsnd) add_edge end;
+      in (map_module name_fragments_common o apsnd) add_edge end;
     val proto_program = empty_program
       |> Graph.fold (fn (name, (stmt, _)) => add_stmt name stmt) program
       |> Graph.fold (fn (name, (_, (_, names))) => fold (add_dependency name) names) program;
 
-    (* name declarations *)
+    (* name declarations, data and statement modifications *)
     fun make_declarations nsps (data, nodes) =
       let
         val (module_fragments, stmt_names) = List.partition
           (fn name_fragment => case Graph.get_node nodes name_fragment
             of (_, Module _) => true | _ => false) (Graph.keys nodes);
-        fun modify_stmt (Stmt (Code_Thingol.Datatypecons _)) = Dummy
-          | modify_stmt (Stmt (Code_Thingol.Classrel _)) = Dummy
-          | modify_stmt (Stmt (Code_Thingol.Classparam _)) = Dummy
-          | modify_stmt stmt = stmt;
-        fun declare namify modify name (nsps, nodes) =
+        fun declare namify name (nsps, nodes) =
           let
             val (base, node) = Graph.get_node nodes name;
             val (base', nsps') = namify node base nsps;
-            val nodes' = Graph.map_node name (K (base', modify node)) nodes;
+            val nodes' = Graph.map_node name (K (base', node)) nodes;
           in (nsps', nodes') end;
         val (nsps', nodes') = (nsps, nodes)
-          |> fold (declare (K namify_module) I) module_fragments
-          |> fold (declare (namify_stmt o (fn Stmt stmt => stmt)) modify_stmt) stmt_names;
-        val nodes'' = nodes'
-          |> fold (fn name_fragment => (Graph.map_node name_fragment
-              o apsnd o map_module) (make_declarations nsps')) module_fragments;
-      in (data, nodes'') end;
+          |> fold (declare (K namify_module)) module_fragments
+          |> fold (declare (namify_stmt o (fn Stmt stmt => stmt))) stmt_names;
+        fun zip_fillup xs ys = xs ~~ ys @ replicate (length xs - length ys) NONE;
+        fun select_names names = case filter (member (op =) stmt_names) names
+         of [] => NONE
+          | xs => SOME xs;
+        val modify_stmts' = AList.make (snd o Graph.get_node nodes)
+          #> split_list
+          ##> map (fn Stmt stmt => stmt)
+          #> (fn (names, stmts) => zip_fillup names (modify_stmts (names ~~ stmts)));
+        val stmtss' = (maps modify_stmts' o map_filter select_names o Graph.strong_conn) nodes;
+        val nodes'' = Graph.map (fn name => apsnd (fn Module content => Module (make_declarations nsps' content)
+            | _ => case AList.lookup (op =) stmtss' name of SOME (SOME stmt) => Stmt stmt | _ => Dummy)) nodes';
+        val data' = fold memorize_data stmt_names data;
+      in (data', nodes'') end;
     val (_, hierarchical_program) = make_declarations empty_nsp proto_program;
 
     (* deresolving *)
--- a/src/Tools/Code/code_scala.ML	Thu Sep 02 09:13:28 2010 +0200
+++ b/src/Tools/Code/code_scala.ML	Thu Sep 02 11:42:50 2010 +0200
@@ -193,8 +193,7 @@
                 str "match", str "{"], str "}")
               (map print_clause eqs)
           end;
-    val print_method = str o Library.enclose "`" "`" o space_implode "+"
-      o Long_Name.explode o deresolve_full;
+    val print_method = str o Library.enclose "`" "`" o deresolve_full;
     fun print_stmt (name, Code_Thingol.Fun (_, (((vs, ty), raw_eqs), _))) =
           print_def name (vs, ty) (filter (snd o snd) raw_eqs)
       | print_stmt (name, Code_Thingol.Datatype (_, (vs, cos))) =
@@ -318,10 +317,14 @@
         val implicits = filter (is_classinst o Graph.get_node program)
           (Graph.imm_succs program name);
       in union (op =) implicits end;
+    fun modify_stmt (_, Code_Thingol.Datatypecons _) = NONE
+      | modify_stmt (_, Code_Thingol.Classrel _) = NONE
+      | modify_stmt (_, Code_Thingol.Classparam _) = NONE
+      | modify_stmt (_, stmt) = SOME stmt;
   in
     Code_Namespace.hierarchical_program labelled_name { module_alias = module_alias, reserved = reserved,
       empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module, namify_stmt = namify_stmt,
-      cyclic_modules = true, empty_data = [], memorize_data = memorize_implicits } program
+      cyclic_modules = true, empty_data = [], memorize_data = memorize_implicits, modify_stmts = map modify_stmt } program
   end;
 
 fun serialize_scala { labelled_name, reserved_syms, includes,
@@ -330,9 +333,8 @@
   let
 
     (* build program *)
-    val reserved = fold (insert (op =) o fst) includes reserved_syms;
     val { deresolver, hierarchical_program = sca_program } =
-      scala_program_of_program labelled_name (Name.make_context reserved) module_alias program;
+      scala_program_of_program labelled_name (Name.make_context reserved_syms) module_alias program;
 
     (* print statements *)
     fun lookup_constr tyco constr = case Graph.get_node program tyco
@@ -352,7 +354,7 @@
      of Code_Thingol.Datatypecons (_, tyco) => null (lookup_constr tyco c)
       | _ => false;
     val print_stmt = print_scala_stmt labelled_name tyco_syntax const_syntax
-      (make_vars reserved) args_num is_singleton_constr;
+      (make_vars reserved_syms) args_num is_singleton_constr;
 
     (* print nodes *)
     fun print_module base implicit_ps p = Pretty.chunks2
@@ -364,7 +366,7 @@
       let
         val s = deresolver prefix_fragments implicit;
       in if length (Long_Name.explode s) = 1 then NONE else SOME (str s) end;
-    fun print_node _ (_, Dummy) = NONE
+    fun print_node _ (_, Code_Namespace.Dummy) = NONE
       | print_node prefix_fragments (name, Code_Namespace.Stmt stmt) =
           if null presentation_names
           orelse member (op =) presentation_names name