explicit namings for generated code
authorhaftmann
Fri, 24 Oct 2008 10:41:13 +0200
changeset 28679 d7384e8e99b3
parent 28678 d93980a6c3cb
child 28680 f1c76cf10915
explicit namings for generated code
doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy
doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex
doc-src/more_antiquote.ML
--- a/doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy	Thu Oct 23 16:07:03 2008 +0200
+++ b/doc-src/IsarAdvanced/Codegen/Thy/Adaption.thy	Fri Oct 24 10:41:13 2008 +0200
@@ -2,7 +2,7 @@
 imports Setup
 begin
 
-setup %invisible {* Code_Target.extend_target ("\<SML>", ("SML", I)) *}
+setup %invisible {* Code_Target.extend_target ("\<SML>", ("SML", K I)) *}
 
 section {* Adaption to target languages \label{sec:adaption} *}
 
--- a/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex	Thu Oct 23 16:07:03 2008 +0200
+++ b/doc-src/IsarAdvanced/Codegen/Thy/document/Adaption.tex	Fri Oct 24 10:41:13 2008 +0200
@@ -26,7 +26,7 @@
 %
 \isataginvisible
 \isacommand{setup}\isamarkupfalse%
-\ {\isacharverbatimopen}\ Code{\isacharunderscore}Target{\isachardot}extend{\isacharunderscore}target\ {\isacharparenleft}{\isachardoublequote}{\isasymSML}{\isachardoublequote}{\isacharcomma}\ {\isacharparenleft}{\isachardoublequote}SML{\isachardoublequote}{\isacharcomma}\ I{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}%
+\ {\isacharverbatimopen}\ Code{\isacharunderscore}Target{\isachardot}extend{\isacharunderscore}target\ {\isacharparenleft}{\isachardoublequote}{\isasymSML}{\isachardoublequote}{\isacharcomma}\ {\isacharparenleft}{\isachardoublequote}SML{\isachardoublequote}{\isacharcomma}\ K\ I{\isacharparenright}{\isacharparenright}\ {\isacharverbatimclose}%
 \endisataginvisible
 {\isafoldinvisible}%
 %
--- a/doc-src/more_antiquote.ML	Thu Oct 23 16:07:03 2008 +0200
+++ b/doc-src/more_antiquote.ML	Fri Oct 24 10:41:13 2008 +0200
@@ -56,19 +56,19 @@
   val parse_const_terms = Scan.repeat1 Args.term
     >> (fn ts => fn thy => map (Code_Unit.check_const thy) ts);
   val parse_consts = Scan.lift (Args.parens (Args.$$$ "consts")) |-- parse_const_terms
-    >> (fn mk_cs => fn thy => map (Code_Name.const thy) (mk_cs thy));
+    >> (fn mk_cs => fn thy => fn naming => map (the o Code_Thingol.lookup_const naming) (mk_cs thy));
   val parse_types = Scan.lift (Args.parens (Args.$$$ "types") |-- Scan.repeat1 Args.name)
-    >> (fn tycos => fn thy => map (Code_Name.tyco thy o Sign.intern_type thy) tycos);
+    >> (fn tycos => fn thy => fn naming => map (the o Code_Thingol.lookup_tyco naming o Sign.intern_type thy) tycos);
   val parse_classes = Scan.lift (Args.parens (Args.$$$ "classes") |-- Scan.repeat1 Args.name)
-    >> (fn classes => fn thy => map (Code_Name.class thy o Sign.intern_class thy) classes);
+    >> (fn classes => fn thy => fn naming => map (the o Code_Thingol.lookup_class naming o Sign.intern_class thy) classes);
   val parse_instances = Scan.lift (Args.parens (Args.$$$ "instances") |-- Scan.repeat1 (Args.name --| Args.$$$ "::" -- Args.name))
-    >> (fn insts => fn thy => map (Code_Name.instance thy o apsnd (Sign.intern_type thy) o apfst (Sign.intern_class thy) o swap) insts);
+    >> (fn insts => fn thy => fn naming => map (the o Code_Thingol.lookup_instance naming o apsnd (Sign.intern_type thy) o apfst (Sign.intern_class thy) o swap) insts);
   val parse_names = parse_consts || parse_types || parse_classes || parse_instances; 
 
   fun code_stmts src ctxt ((mk_cs, mk_stmtss), target) =
     Code_Target.code_of (ProofContext.theory_of ctxt)
       target "Example" (mk_cs (ProofContext.theory_of ctxt))
-        (maps (fn f => f (ProofContext.theory_of ctxt)) mk_stmtss)
+        (fn naming => maps (fn f => f (ProofContext.theory_of ctxt) naming) mk_stmtss)
     |> split_lines
     |> verbatim_lines;