| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 19 Nov 2009 14:11:50 +0100 | |
| changeset 394 | 0019ebf76e10 | 
| parent 346 | 0fea8b7a14a1 | 
| child 396 | a2e49e0771b3 | 
| permissions | -rw-r--r-- | 
| 23 | 1 | theory Base | 
| 264 | 2 | imports Main LaTeXsugar | 
| 23 | 3 | uses | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 4 |   ("output_tutorial.ML")
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 5 |   ("antiquote_setup.ML")
 | 
| 23 | 6 | begin | 
| 7 | ||
| 394 | 8 | notation (latex output) | 
| 9 |   Cons ("_ # _" [66,65] 65)
 | |
| 10 | ||
| 11 | ||
| 317 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 12 | (* rebinding of writeln and tracing so that it can *) | 
| 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 13 | (* be compared in intiquotations *) | 
| 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 14 | ML {* 
 | 
| 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 15 | fun writeln s = (Output.writeln s; s) | 
| 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 16 | fun tracing s = (Output.tracing s; s) | 
| 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 17 | *} | 
| 
d69214e47ef9
added an experimental antiquotation to replace eventually ML_response_fake
 Christian Urban <urbanc@in.tum.de> parents: 
316diff
changeset | 18 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 19 | (* re-definition of various ML antiquotations *) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 20 | (* to have a special tag for text enclosed in ML; *) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 21 | (* they also write the code into a separate file *) | 
| 301 
2728e8daebc0
replaced "writeln" with "tracing"
 Christian Urban <urbanc@in.tum.de> parents: 
264diff
changeset | 22 | |
| 64 
9a6e5e0c4906
deleted old files and added code to give a special tag to the command ML
 Christian Urban <urbanc@in.tum.de> parents: 
57diff
changeset | 23 | ML {*
 | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 24 | val (filename, setup_filename) = Attrib.config_string "filename" "File_Code.ML" | 
| 310 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 25 | *} | 
| 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 26 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 27 | setup {* setup_filename *}
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 28 | |
| 310 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 29 | ML {*
 | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 30 | fun write_file txt thy = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 31 | let | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 32 | val stream = Config.get_thy thy filename | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 33 | |> TextIO.openAppend | 
| 310 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 34 | in | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 35 | TextIO.output (stream, txt); | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 36 | TextIO.flushOut stream; (* needed ?*) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 37 | TextIO.closeOut stream | 
| 310 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 38 | end | 
| 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 39 | *} | 
| 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 40 | |
| 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 41 | ML {*
 | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 42 | fun write_file_ml_blk txt thy = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 43 | let | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 44 |   val pre  = implode ["\n", "ML ", "{", "*", "\n"]
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 45 | val post = implode ["\n", "*", "}", "\n"] | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 46 | val _ = write_file (enclose pre post txt) thy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 47 | in | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 48 | thy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 49 | end | 
| 310 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 50 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 51 | fun write_file_setup_blk txt thy = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 52 | let | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 53 |   val pre  = implode ["\n", "setup ", "{", "*", "\n"]
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 54 | val post = implode ["\n", "*", "}", "\n"] | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 55 | val _ = write_file (enclose pre post txt) thy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 56 | in | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 57 | thy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 58 | end | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 59 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 60 | fun write_file_lsetup_blk txt lthy = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 61 | let | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 62 |   val pre  = implode ["\n", "local_setup ", "{", "*", "\n"]
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 63 | val post = implode ["\n", "*", "}", "\n"] | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 64 | val _ = write_file (enclose pre post txt) (ProofContext.theory_of lthy) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 65 | in | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 66 | lthy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 67 | end | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 68 | |
| 310 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 69 | *} | 
| 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 70 | |
| 
007922777ff1
added some rudimentary inrastructure for producing the ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
302diff
changeset | 71 | ML {*
 | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 72 | fun open_file name thy = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 73 | let | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 74 |   val _ = tracing ("Open File: " ^ name)
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 75 | val _ = TextIO.openOut name | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 76 | in | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 77 | Config.put_thy filename name thy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 78 | end | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 79 | *} | 
| 64 
9a6e5e0c4906
deleted old files and added code to give a special tag to the command ML
 Christian Urban <urbanc@in.tum.de> parents: 
57diff
changeset | 80 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 81 | ML {*
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 82 | fun open_file_with_prelude name txts thy = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 83 | let | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 84 | val thy' = open_file name thy | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 85 | val _ = write_file (cat_lines txts) thy' | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 86 | in | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 87 | thy' | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 88 | end | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 89 | *} | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 90 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 91 | ML {*
 | 
| 260 
5accec94b6df
updated to lates Isabelle changes
 Christian Urban <urbanc@in.tum.de> parents: 
256diff
changeset | 92 | fun propagate_env (context as Context.Proof lthy) = | 
| 394 | 93 | Context.Proof (Local_Theory.map_contexts (ML_Env.inherit context) lthy) | 
| 264 | 94 | | propagate_env context = context | 
| 210 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 95 | |
| 260 
5accec94b6df
updated to lates Isabelle changes
 Christian Urban <urbanc@in.tum.de> parents: 
256diff
changeset | 96 | fun propagate_env_prf prf = Proof.map_contexts | 
| 264 | 97 | (Context.proof_map (ML_Env.inherit (Context.Proof (Proof.context_of prf)))) prf | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 98 | *} | 
| 210 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 99 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 100 | ML {*
 | 
| 64 
9a6e5e0c4906
deleted old files and added code to give a special tag to the command ML
 Christian Urban <urbanc@in.tum.de> parents: 
57diff
changeset | 101 | val _ = | 
| 
9a6e5e0c4906
deleted old files and added code to give a special tag to the command ML
 Christian Urban <urbanc@in.tum.de> parents: 
57diff
changeset | 102 | OuterSyntax.command "ML" "eval ML text within theory" | 
| 106 
bdd82350cf22
renamed in the pdf all instances of cookbook to tutorial (in order to sound more serious)
 Christian Urban <urbanc@in.tum.de> parents: 
80diff
changeset | 103 | (OuterKeyword.tag "TutorialML" OuterKeyword.thy_decl) | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 104 | (OuterParse.position OuterParse.text >> | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 105 | (fn (txt, pos) => | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 106 | Toplevel.generic_theory | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 107 | (ML_Context.exec | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 108 | (fn () => ML_Context.eval true pos txt) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 109 | #> propagate_env #> Context.map_theory (write_file_ml_blk txt)))) | 
| 210 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 110 | |
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 111 | val _ = | 
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 112 | OuterSyntax.command "ML_prf" "ML text within proof" | 
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 113 | (OuterKeyword.tag "TutorialML" OuterKeyword.prf_decl) | 
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 114 | (OuterParse.ML_source >> (fn (txt, pos) => | 
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 115 | Toplevel.proof (Proof.map_context (Context.proof_map | 
| 260 
5accec94b6df
updated to lates Isabelle changes
 Christian Urban <urbanc@in.tum.de> parents: 
256diff
changeset | 116 | (ML_Context.exec (fn () => ML_Context.eval true pos txt))) #> propagate_env_prf))) | 
| 210 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 117 | |
| 224 
647cab4a72c2
finished the heavy duty stuff for the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 118 | val _ = | 
| 
647cab4a72c2
finished the heavy duty stuff for the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 119 | OuterSyntax.command "ML_val" "diagnostic ML text" | 
| 
647cab4a72c2
finished the heavy duty stuff for the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 120 | (OuterKeyword.tag "TutorialML" OuterKeyword.diag) | 
| 264 | 121 | (OuterParse.ML_source >> IsarCmd.ml_diag true) | 
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 122 | *} | 
| 224 
647cab4a72c2
finished the heavy duty stuff for the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 123 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 124 | ML {*
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 125 | val _ = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 126 | OuterSyntax.command "setup" "ML theory setup" | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 127 | (OuterKeyword.tag_ml OuterKeyword.thy_decl) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 128 | (OuterParse.ML_source >> (fn (txt, pos) => | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 129 | (Toplevel.theory (IsarCmd.global_setup (txt, pos) #> write_file_setup_blk txt)))) | 
| 64 
9a6e5e0c4906
deleted old files and added code to give a special tag to the command ML
 Christian Urban <urbanc@in.tum.de> parents: 
57diff
changeset | 130 | *} | 
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
224diff
changeset | 131 | |
| 346 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 132 | ML {* Context.proof_map *}
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 133 | ML {* IsarCmd.local_setup ; Context.map_proof (write_file_lsetup_blk "") *}
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 134 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 135 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 136 | ML {*
 | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 137 | val _ = | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 138 | OuterSyntax.local_theory "local_setup" "ML local theory setup" | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 139 | (OuterKeyword.tag_ml OuterKeyword.thy_decl) | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 140 | (OuterParse.ML_source >> (fn (txt, pos) => | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 141 | IsarCmd.local_setup (txt, pos) #> write_file_lsetup_blk txt)); | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 142 | *} | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 143 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 144 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 145 | use "output_tutorial.ML" | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 146 | use "antiquote_setup.ML" | 
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 147 | |
| 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 Christian Urban <urbanc@in.tum.de> parents: 
328diff
changeset | 148 | |
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
224diff
changeset | 149 | end |