merged
authorhaftmann
Thu, 02 Sep 2010 08:35:16 +0200
changeset 39025 582546cc49a1
parent 38978 4bf80c23320e (current diff)
parent 39024 30d5dd2f30b6 (diff)
child 39026 962d12bc546c
merged
--- a/src/HOL/Boogie/Tools/boogie_vcs.ML	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/HOL/Boogie/Tools/boogie_vcs.ML	Thu Sep 02 08:35:16 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/Imperative_HOL/Heap_Monad.thy	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/HOL/Imperative_HOL/Heap_Monad.thy	Thu Sep 02 08:35:16 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/Tools/Predicate_Compile/predicate_compile_core.ML	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/HOL/Tools/SMT/smt_monomorph.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_parser.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/HOL/Tools/SMT/z3_proof_reconstruction.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/General/graph.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/General/table.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/Isar/code.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/Thy/thy_info.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/context.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/proofterm.ML	Thu Sep 02 08:35:16 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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Pure/sorts.ML	Thu Sep 02 08:35:16 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_namespace.ML	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Tools/Code/code_namespace.ML	Thu Sep 02 08:35:16 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,44 @@
         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;
+        val modify_stmts' = filter (member (op =) stmt_names)
+          #> AList.make (snd o Graph.get_node nodes)
+          #> split_list
+          ##> map (fn Stmt stmt => stmt)
+          #> (fn (names, stmts) => names ~~ modify_stmts (names ~~ stmts));
+        val stmtss' = maps modify_stmts' (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	Wed Sep 01 23:03:31 2010 +0200
+++ b/src/Tools/Code/code_scala.ML	Thu Sep 02 08:35:16 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,
@@ -364,7 +367,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