keep only identifiers public which are explicitly requested or demanded by dependencies
authorhaftmann
Sun, 23 Feb 2014 10:33:43 +0100
changeset 55684 ee49b4f7edc8
parent 55683 5732a55b9232
child 55685 3f8bdc5364a9
keep only identifiers public which are explicitly requested or demanded by dependencies
src/HOL/Tools/Quickcheck/narrowing_generators.ML
src/Tools/Code/code_ml.ML
src/Tools/Code/code_namespace.ML
src/Tools/Code/code_runtime.ML
src/Tools/Code/code_scala.ML
src/Tools/Code/code_target.ML
--- a/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Sun Feb 23 10:33:43 2014 +0100
+++ b/src/HOL/Tools/Quickcheck/narrowing_generators.ML	Sun Feb 23 10:33:43 2014 +0100
@@ -308,7 +308,7 @@
     val ctxt = Proof_Context.init_global thy
     fun evaluator program ((_, vs_ty), t) deps =
       Exn.interruptible_capture (value opts ctxt cookie)
-        (Code_Target.evaluator thy target program deps (vs_ty, t));
+        (Code_Target.evaluator thy target program deps true (vs_ty, t));
   in Exn.release (Code_Thingol.dynamic_value thy (Exn.map_result o postproc) evaluator t) end;
 
 (** counterexample generator **)
--- a/src/Tools/Code/code_ml.ML	Sun Feb 23 10:33:43 2014 +0100
+++ b/src/Tools/Code/code_ml.ML	Sun Feb 23 10:33:43 2014 +0100
@@ -37,7 +37,7 @@
 datatype ml_stmt =
     ML_Exc of string * (typscheme * int)
   | ML_Val of ml_binding
-  | ML_Funs of ml_binding list * Code_Symbol.T list
+  | ML_Funs of (Code_Namespace.export * ml_binding) list * Code_Symbol.T list
   | ML_Datas of (string * (vname list * ((string * vname list) * itype list) list)) list
   | ML_Class of string * (vname * ((class * class) list * (string * itype) list));
 
@@ -260,7 +260,7 @@
             [sig_p]
             (semicolon [p])
           end
-      | print_stmt _ (ML_Funs (binding :: bindings, pseudo_funs)) =
+      | print_stmt _ (ML_Funs ((export, binding) :: exports_bindings, pseudo_funs)) =
           let
             val print_def' = print_def (member (op =) pseudo_funs) false;
             fun print_pseudo_fun sym = concat [
@@ -271,10 +271,11 @@
                 str "();"
               ];
             val (sig_ps, (ps, p)) = (apsnd split_last o split_list)
-              (print_def' "fun" binding :: map (print_def' "and") bindings);
+              (print_def' "fun" binding :: map (print_def' "and" o snd) exports_bindings);
             val pseudo_ps = map print_pseudo_fun pseudo_funs;
           in pair
-            sig_ps
+            (map_filter (fn (export, p) => if Code_Namespace.not_private export then SOME p else NONE)
+              ((export :: map fst exports_bindings) ~~ sig_ps))
             (Pretty.chunks (ps @ semicolon [p] :: pseudo_ps))
           end
      | print_stmt _ (ML_Datas [(tyco, (vs, []))]) =
@@ -605,7 +606,7 @@
             [sig_p]
             (doublesemicolon [p])
           end
-      | print_stmt _ (ML_Funs (binding :: bindings, pseudo_funs)) =
+      | print_stmt _ (ML_Funs ((export, binding) :: exports_bindings, pseudo_funs)) =
           let
             val print_def' = print_def (member (op =) pseudo_funs) false;
             fun print_pseudo_fun sym = concat [
@@ -616,10 +617,11 @@
                 str "();;"
               ];
             val (sig_ps, (ps, p)) = (apsnd split_last o split_list)
-              (print_def' "let rec" binding :: map (print_def' "and") bindings);
+              (print_def' "let rec" binding :: map (print_def' "and" o snd) exports_bindings);
             val pseudo_ps = map print_pseudo_fun pseudo_funs;
           in pair
-            sig_ps
+            (map_filter (fn (export, p) => if Code_Namespace.not_private export then SOME p else NONE)
+              ((export :: map fst exports_bindings) ~~ sig_ps))
             (Pretty.chunks (ps @ doublesemicolon [p] :: pseudo_ps))
           end
      | print_stmt _ (ML_Datas [(tyco, (vs, []))]) =
@@ -667,7 +669,9 @@
                 )
               ];
           in pair
-            (type_decl_p :: map print_classparam_decl classparams)
+           (if Code_Namespace.is_public export
+              then type_decl_p :: map print_classparam_decl classparams
+              else [concat [str "type", print_dicttyp (class, ITyVar v)]])
             (Pretty.chunks (
               doublesemicolon [type_decl_p]
               :: map print_classparam_proj classparams
@@ -733,7 +737,7 @@
       | 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 (sym as Constant const, Code_Thingol.Fun ((tysm as (vs, ty), raw_eqs), _)) =
+    fun ml_binding_of_stmt (sym as Constant const, (export, Code_Thingol.Fun ((tysm as (vs, ty), raw_eqs), _))) =
           let
             val eqs = filter (snd o snd) raw_eqs;
             val (eqs', some_sym) = if null (filter_out (null o snd) vs) then case eqs
@@ -742,49 +746,58 @@
                   else (eqs, SOME (sym, member (op =) (Code_Thingol.add_constsyms t []) sym))
                 | _ => (eqs, NONE)
               else (eqs, NONE)
-          in (ML_Function (const, (tysm, eqs')), some_sym) end
-      | ml_binding_of_stmt (sym as Class_Instance inst, Code_Thingol.Classinst (stmt as { vs, ... })) =
-          (ML_Instance (inst, stmt), if forall (null o snd) vs then SOME (sym, false) else NONE)
+          in ((export, ML_Function (const, (tysm, eqs'))), some_sym) end
+      | ml_binding_of_stmt (sym as Class_Instance inst, (export, Code_Thingol.Classinst (stmt as { vs, ... }))) =
+          ((export, ML_Instance (inst, stmt)),
+            if forall (null o snd) vs then SOME (sym, false) else NONE)
       | ml_binding_of_stmt (sym, _) =
           error ("Binding block containing illegal statement: " ^ 
             Code_Symbol.quote ctxt sym)
-    fun modify_fun (sym, stmt) =
+    fun modify_fun (sym, (export, stmt)) =
       let
-        val (binding, some_value_sym) = ml_binding_of_stmt (sym, stmt);
+        val ((export', binding), some_value_sym) = ml_binding_of_stmt (sym, (export, stmt));
         val ml_stmt = case binding
          of ML_Function (const, ((vs, ty), [])) =>
               ML_Exc (const, ((vs, ty),
                 (length o filter_out (null o snd)) vs + (length o fst o Code_Thingol.unfold_fun) ty))
           | _ => case some_value_sym
-             of NONE => ML_Funs ([binding], [])
-              | SOME (sym, true) => ML_Funs ([binding], [sym])
+             of NONE => ML_Funs ([(export', binding)], [])
+              | SOME (sym, true) => ML_Funs ([(export, binding)], [sym])
               | SOME (sym, false) => ML_Val binding
-      in SOME ml_stmt end;
+      in SOME (export, 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 (Type_Constructor tyco, Code_Thingol.Datatype stmt) => SOME (tyco, stmt) | _ => NONE) stmts)))
-    fun modify_class stmts = single (SOME
-      (ML_Class (the_single (map_filter
-        (fn (Type_Class class, Code_Thingol.Class stmt) => SOME (class, stmt) | _ => NONE) stmts))))
-    fun modify_stmts ([stmt as (_, stmt' as Code_Thingol.Fun _)]) =
+      (Code_Namespace.Opaque, ML_Funs (map_split ml_binding_of_stmt stmts |> (apsnd o map_filter o Option.map) fst)))
+    fun modify_datatypes stmts =
+      map_filter
+        (fn (Type_Constructor tyco, (export, Code_Thingol.Datatype stmt)) => SOME (export, (tyco, stmt)) | _ => NONE) stmts
+      |> split_list
+      |> apfst Code_Namespace.join_exports
+      |> apsnd ML_Datas
+      |> SOME
+      |> single;
+    fun modify_class stmts =
+      the_single (map_filter
+        (fn (Type_Class class, (export, Code_Thingol.Class stmt)) => SOME (export, (class, stmt)) | _ => NONE) stmts)
+      |> apsnd ML_Class
+      |> SOME
+      |> single;
+    fun modify_stmts ([stmt as (_, (_, stmt' as Code_Thingol.Fun _))]) =
           if Code_Thingol.is_case stmt' then [] else [modify_fun stmt]
-      | modify_stmts ((stmts as (_, Code_Thingol.Fun _) :: _)) =
-          modify_funs (filter_out (Code_Thingol.is_case o snd) stmts)
-      | modify_stmts ((stmts as (_, Code_Thingol.Datatypecons _) :: _)) =
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Fun _)) :: _)) =
+          modify_funs (filter_out (Code_Thingol.is_case o snd o snd) stmts)
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Datatypecons _)) :: _)) =
           modify_datatypes stmts
-      | modify_stmts ((stmts as (_, Code_Thingol.Datatype _) :: _)) =
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Datatype _)) :: _)) =
           modify_datatypes stmts
-      | modify_stmts ((stmts as (_, Code_Thingol.Class _) :: _)) =
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Class _)) :: _)) =
           modify_class stmts
-      | modify_stmts ((stmts as (_, Code_Thingol.Classrel _) :: _)) =
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Classrel _)) :: _)) =
           modify_class stmts
-      | modify_stmts ((stmts as (_, Code_Thingol.Classparam _) :: _)) =
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Classparam _)) :: _)) =
           modify_class stmts
-      | modify_stmts ([stmt as (_, Code_Thingol.Classinst _)]) =
+      | modify_stmts ([stmt as (_, (_, Code_Thingol.Classinst _))]) =
           [modify_fun stmt]
-      | modify_stmts ((stmts as (_, Code_Thingol.Classinst _) :: _)) =
+      | modify_stmts ((stmts as (_, (_, Code_Thingol.Classinst _)) :: _)) =
           modify_funs stmts
       | modify_stmts stmts = error ("Illegal mutual dependencies: " ^
           (Library.commas o map (Code_Symbol.quote ctxt o fst)) stmts);
@@ -792,7 +805,9 @@
     Code_Namespace.hierarchical_program ctxt {
       module_name = module_name, reserved = reserved, identifiers = identifiers,
       empty_nsp = (reserved, reserved), namify_module = pair, namify_stmt = namify_stmt,
-      cyclic_modules = false, empty_data = (), memorize_data = K I, modify_stmts = modify_stmts }
+      cyclic_modules = false, class_transitive = false,
+      class_relation_public = true, empty_data = (),
+      memorize_data = K I, modify_stmts = modify_stmts }
   end;
 
 fun serialize_ml print_ml_module print_ml_stmt ctxt
--- a/src/Tools/Code/code_namespace.ML	Sun Feb 23 10:33:43 2014 +0100
+++ b/src/Tools/Code/code_namespace.ML	Sun Feb 23 10:33:43 2014 +0100
@@ -9,6 +9,7 @@
   datatype export = Private | Opaque | Public
   val is_public: export -> bool
   val not_private: export -> bool
+  val join_exports: export list -> export
 
   type flat_program
   val flat_program: Proof.context
@@ -30,8 +31,10 @@
     reserved: Name.context, identifiers: Code_Target.identifier_data,
     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: Code_Symbol.T -> 'b -> 'b,
-    modify_stmts: (Code_Symbol.T * Code_Thingol.stmt) list -> 'a option list }
+    cyclic_modules: bool,
+    class_transitive: bool, class_relation_public: bool,
+    empty_data: 'b, memorize_data: Code_Symbol.T -> 'b -> 'b,
+    modify_stmts: (Code_Symbol.T * (export * Code_Thingol.stmt)) list -> (export * 'a) option list }
       -> Code_Symbol.T list -> Code_Thingol.program
       -> { deresolver: string list -> Code_Symbol.T -> string,
            hierarchical_program: ('a, 'b) hierarchical_program }
@@ -55,6 +58,94 @@
   | not_private Opaque = true
   | not_private _ = false;
 
+fun mark_export Public _ = Public
+  | mark_export _ Public = Public
+  | mark_export Opaque _ = Opaque
+  | mark_export _ Opaque = Opaque
+  | mark_export _ _ = Private;
+
+fun join_exports exports = fold mark_export exports Private;
+
+fun dependent_exports { program = program, class_transitive = class_transitive } =
+  let
+    fun is_datatype_or_class (Code_Symbol.Type_Constructor _) = true
+      | is_datatype_or_class (Code_Symbol.Type_Class _) = true
+      | is_datatype_or_class _ = false;
+    fun is_relevant (Code_Symbol.Class_Relation _) = true
+      | is_relevant sym = is_datatype_or_class sym;
+    val proto_gr = Code_Symbol.Graph.restrict is_relevant program;
+    val gr =
+      proto_gr
+      |> Code_Symbol.Graph.fold
+          (fn (sym, (_, (_, deps))) =>
+            if is_relevant sym
+            then I
+            else
+              Code_Symbol.Graph.new_node (sym, Code_Thingol.NoStmt)
+              #> Code_Symbol.Graph.Keys.fold
+               (fn sym' =>
+                if is_relevant sym'
+                then Code_Symbol.Graph.add_edge (sym, sym')
+                else I) deps) program
+      |> class_transitive ?
+          Code_Symbol.Graph.fold (fn (sym as Code_Symbol.Type_Class _, _) =>
+            fold (curry Code_Symbol.Graph.add_edge sym)
+              ((remove (op =) sym o Code_Symbol.Graph.all_succs proto_gr) [sym]) | _ => I) proto_gr
+    fun deps_of sym =
+      let
+        val succs = Code_Symbol.Graph.Keys.dest o Code_Symbol.Graph.imm_succs gr;
+        val deps1 = succs sym;
+        val deps2 = if class_transitive
+          then []
+          else [] |> fold (union (op =)) (map succs deps1) |> subtract (op =) deps1
+      in (deps1, deps2) end;
+  in
+    { is_datatype_or_class = is_datatype_or_class,
+      deps_of = deps_of }
+  end;
+
+fun mark_exports_aux { program = program, prefix_of = prefix_of, map_export = map_export,
+    is_datatype_or_class = is_datatype_or_class, deps_of = deps_of,
+    class_relation_public = class_relation_public } prefix sym =
+  let
+    val export = (if is_datatype_or_class sym then Opaque else Public);
+    val (dependent_export1, dependent_export2) =
+      case Code_Symbol.Graph.get_node program sym of
+          Code_Thingol.Fun _ => (SOME Opaque, NONE)
+        | Code_Thingol.Classinst _ => (SOME Opaque, NONE)
+        | Code_Thingol.Datatypecons _ => (SOME Public, SOME Opaque)
+        | Code_Thingol.Classparam _ => (SOME Public, SOME Opaque)
+        | Code_Thingol.Classrel _ =>
+           (if class_relation_public
+            then (SOME Public, SOME Opaque)
+            else (SOME Opaque, NONE))
+        | _ => (NONE, NONE);
+    val dependent_exports =
+      case dependent_export1 of
+        SOME export1 => (case dependent_export2 of
+          SOME export2 =>
+            let
+              val (deps1, deps2) = deps_of sym
+            in map (rpair export1) deps1 @ map (rpair export2) deps2 end
+        | NONE => map (rpair export1) (fst (deps_of sym)))
+      | NONE => [];
+  in 
+    map_export prefix sym (mark_export export)
+    #> fold (fn (sym, export) => map_export (prefix_of sym) sym (mark_export export))
+      dependent_exports
+  end;
+
+fun mark_exports { program = program, prefix_of = prefix_of, map_export = map_export,
+    class_transitive = class_transitive, class_relation_public = class_relation_public } =
+  let
+    val { is_datatype_or_class, deps_of } =
+      dependent_exports { program = program, class_transitive = class_transitive };
+  in
+    mark_exports_aux { program = program, prefix_of = prefix_of, map_export = map_export,
+      is_datatype_or_class = is_datatype_or_class, deps_of = deps_of,
+      class_relation_public = class_relation_public }
+  end;
+
 
 (** fundamental module name hierarchy **)
 
@@ -113,13 +204,17 @@
     val sym_priority = has_priority identifiers;
 
     (* distribute statements over hierarchy *)
+    val mark_exports = mark_exports { program = program, prefix_of = fst o prep_sym,
+      map_export = fn module_name => fn sym =>
+        Graph.map_node module_name o apfst o Code_Symbol.Graph.map_node sym o apsnd o apfst,
+        class_transitive = false, class_relation_public = false };
     fun add_stmt sym stmt =
       let
         val (module_name, base) = prep_sym sym;
       in
         Graph.default_node (module_name, (Code_Symbol.Graph.empty, []))
         #> (Graph.map_node module_name o apfst)
-          (Code_Symbol.Graph.new_node (sym, (base, (Public, stmt))))
+          (Code_Symbol.Graph.new_node (sym, (base, (if null exports then Public else Private, stmt))))
       end;
     fun add_dep sym sym' =
       let
@@ -129,9 +224,11 @@
         then (Graph.map_node module_name o apfst) (Code_Symbol.Graph.add_edge (sym, sym'))
         else (Graph.map_node module_name o apsnd)
           (AList.map_default (op =) (module_name', []) (insert (op =) sym'))
+          #> mark_exports module_name' sym'
       end;
     val proto_program = build_proto_program
-      { empty = Graph.empty, add_stmt = add_stmt, add_dep = add_dep } program;
+      { empty = Graph.empty, add_stmt = add_stmt, add_dep = add_dep } program
+      |> fold (fn sym => mark_exports ((fst o prep_sym) sym) sym) exports;
 
     (* name declarations and statement modifications *)
     fun declare sym (base, (_, stmt)) (gr, nsp) = 
@@ -200,7 +297,7 @@
   let
     val some_modules =
       sym_base_nodes
-      |> map (fn (sym, (base, Module content)) => SOME (base, content) | _ => NONE)
+      |> map (fn (_, (base, Module content)) => SOME (base, content) | _ => NONE)
       |> (burrow_options o map o apsnd) f_module;
     val some_export_stmts =
       sym_base_nodes
@@ -214,7 +311,9 @@
   end;
 
 fun hierarchical_program ctxt { module_name, reserved, identifiers, empty_nsp,
-      namify_module, namify_stmt, cyclic_modules, empty_data, memorize_data, modify_stmts }
+      namify_module, namify_stmt, cyclic_modules,
+      class_transitive, class_relation_public,
+      empty_data, memorize_data, modify_stmts }
       exports program =
   let
 
@@ -242,12 +341,17 @@
           o Code_Symbol.lookup identifiers o fst) program;
 
     (* distribute statements over hierarchy *)
+    val mark_exports = mark_exports { program = program, prefix_of = fst o prep_sym,
+      map_export = fn name_fragments => fn sym => fn f =>
+        (map_module name_fragments o apsnd o Code_Symbol.Graph.map_node sym o apsnd)
+          (fn Stmt (export, stmt) => Stmt (f export, stmt)),
+      class_transitive = class_transitive, class_relation_public = class_relation_public };
     fun add_stmt sym stmt =
       let
         val (name_fragments, base) = prep_sym sym;
       in
         (map_module name_fragments o apsnd)
-          (Code_Symbol.Graph.new_node (sym, (base, Stmt (Public, stmt))))
+          (Code_Symbol.Graph.new_node (sym, (base, Stmt (if null exports then Public else Private, stmt))))
       end;
     fun add_edge_acyclic_error error_msg dep gr =
       Code_Symbol.Graph.add_edge_acyclic dep gr
@@ -266,9 +370,13 @@
             ^ Code_Symbol.quote ctxt sym'
             ^ " would result in module dependency cycle") dep
           else Code_Symbol.Graph.add_edge dep;
-      in (map_module name_fragments_common o apsnd) add_edge end;
+      in
+        (map_module name_fragments_common o apsnd) add_edge
+        #> (if is_cross_module then mark_exports name_fragments' sym' else I)
+      end;
     val proto_program = build_proto_program
-      { empty = empty_program, add_stmt = add_stmt, add_dep = add_dep } program;
+      { empty = empty_program, add_stmt = add_stmt, add_dep = add_dep } program
+      |> fold (fn sym => mark_exports ((fst o prep_sym) sym) sym) exports;
 
     (* name declarations, data and statement modifications *)
     fun make_declarations nsps (data, nodes) =
@@ -292,14 +400,9 @@
           let
             val stmts' = modify_stmts syms_stmts
           in stmts' @ replicate (length syms_stmts - length stmts') NONE end;
-        fun modify_stmts'' syms_exports_stmts =
-          syms_exports_stmts
-          |> map (fn (sym, (export, stmt)) => ((sym, stmt), export))
-          |> burrow_fst modify_stmts'
-          |> map (fn (SOME stmt, export) => SOME (export, stmt) | _ => NONE);
         val nodes'' =
           nodes'
-          |> Code_Symbol.Graph.map_strong_conn (map_module_stmts (make_declarations nsps') modify_stmts'');
+          |> Code_Symbol.Graph.map_strong_conn (map_module_stmts (make_declarations nsps') modify_stmts');
         val data' = fold memorize_data stmt_syms data;
       in (data', nodes'') end;
     val (_, hierarchical_program) = make_declarations empty_nsp proto_program;
--- a/src/Tools/Code/code_runtime.ML	Sun Feb 23 10:33:43 2014 +0100
+++ b/src/Tools/Code/code_runtime.ML	Sun Feb 23 10:33:43 2014 +0100
@@ -89,7 +89,7 @@
   in ((Sign.no_frees ctxt o Sign.no_vars ctxt o map_types (K dummyT)) t; t) end;
 
 fun obtain_evaluator thy some_target program consts expr =
-  Code_Target.evaluator thy (the_default target some_target) program consts expr
+  Code_Target.evaluator thy (the_default target some_target) program consts false expr
   |> apfst (fn ml_modules => space_implode "\n\n" (map snd ml_modules));
 
 fun obtain_evaluator' thy some_target program =
--- a/src/Tools/Code/code_scala.ML	Sun Feb 23 10:33:43 2014 +0100
+++ b/src/Tools/Code/code_scala.ML	Sun Feb 23 10:33:43 2014 +0100
@@ -249,7 +249,7 @@
                 val auxs = Name.invent (snd proto_vars) "a" (length tys);
                 val vars = intro_vars auxs proto_vars;
               in
-                privatize export [str "def", constraint (Pretty.block [applify "(" ")"
+                privatize' export [str "def", constraint (Pretty.block [applify "(" ")"
                   (fn (aux, ty) => constraint ((str o lookup_var vars) aux)
                   (print_typ tyvars NOBR ty)) NOBR (add_typarg (deresolve_const classparam))
                   (auxs ~~ tys), str "(implicit ", str implicit_name, str ": ",
@@ -284,7 +284,7 @@
                 val aux_abstr1 = abstract_using aux_dom1;
                 val aux_abstr2 = abstract_using aux_dom2;
               in
-                privatize export ([str "val", print_method classparam, str "="]
+                concat ([str "val", print_method classparam, str "="]
                   @ aux_abstr1 @ aux_abstr2 @| print_app tyvars false (SOME thm) vars NOBR
                     (const, map (IVar o SOME) auxs))
               end;
@@ -333,16 +333,17 @@
         val implicits = filter (is_classinst o Code_Symbol.Graph.get_node program)
           (Code_Symbol.Graph.immediate_succs program sym);
       in union (op =) implicits end;
-    fun modify_stmt (_, Code_Thingol.Fun (_, SOME _)) = NONE
-      | modify_stmt (_, Code_Thingol.Datatypecons _) = NONE
-      | modify_stmt (_, Code_Thingol.Classrel _) = NONE
-      | modify_stmt (_, Code_Thingol.Classparam _) = NONE
-      | modify_stmt (_, stmt) = SOME stmt;
+    fun modify_stmt (_, (_, Code_Thingol.Fun (_, SOME _))) = NONE
+      | modify_stmt (_, (_, Code_Thingol.Datatypecons _)) = NONE
+      | modify_stmt (_, (_, Code_Thingol.Classrel _)) = NONE
+      | modify_stmt (_, (_, Code_Thingol.Classparam _)) = NONE
+      | modify_stmt (_, export_stmt) = SOME export_stmt;
   in
     Code_Namespace.hierarchical_program ctxt
       { module_name = module_name, reserved = reserved, identifiers = identifiers,
         empty_nsp = ((reserved, reserved), reserved), namify_module = namify_module,
-        namify_stmt = namify_stmt, cyclic_modules = true, empty_data = [],
+        namify_stmt = namify_stmt, cyclic_modules = true,
+        class_transitive = true, class_relation_public = false, empty_data = [],
         memorize_data = memorize_implicits, modify_stmts = map modify_stmt } exports program
   end;
 
--- a/src/Tools/Code/code_target.ML	Sun Feb 23 10:33:43 2014 +0100
+++ b/src/Tools/Code/code_target.ML	Sun Feb 23 10:33:43 2014 +0100
@@ -29,7 +29,7 @@
 
   val generatedN: string
   val evaluator: theory -> string -> Code_Thingol.program
-    -> Code_Symbol.T list -> ((string * class list) list * Code_Thingol.itype) * Code_Thingol.iterm
+    -> Code_Symbol.T list -> bool -> ((string * class list) list * Code_Thingol.itype) * Code_Thingol.iterm
     -> (string * string) list * string
 
   type serializer
@@ -426,7 +426,7 @@
     else Isabelle_System.with_tmp_dir "Code_Test" ext_check
   end;
 
-fun evaluation mounted_serializer prepared_program syms ((vs, ty), t) =
+fun evaluation mounted_serializer prepared_program syms all_public ((vs, ty), t) =
   let
     val _ = if Code_Thingol.contains_dict_var t then
       error "Term to be evaluated contains free dictionaries" else ();
@@ -439,7 +439,7 @@
       |> fold (curry (perhaps o try o
           Code_Symbol.Graph.add_edge) Code_Symbol.value) syms;
     val (program_code, deresolve) =
-      produce (mounted_serializer program [Code_Symbol.value]);
+      produce (mounted_serializer program (if all_public then [] else [Code_Symbol.value]));
     val value_name = the (deresolve Code_Symbol.value);
   in (program_code, value_name) end;