| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 26 Feb 2009 12:16:24 +0000 | |
| changeset 149 | 253ea99c1441 | 
| parent 112 | a90d0fb24e75 | 
| child 165 | 890fbfef6d6b | 
| permissions | -rw-r--r-- | 
| 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: 
103diff
changeset | 1 | (* Auxiliary antiquotations for the tutorial. *) | 
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 2 | |
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | structure AntiquoteSetup: sig end = | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | struct | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | |
| 112 | 6 | (* functions for generating appropriate expressions *) | 
| 7 | ||
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 8 | fun ml_val_open (ys, xs) txt = | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 9 | let fun ml_val_open_aux ys txt = | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 10 |           "fn " ^ (case ys of [] => "_" | _ => enclose "(" ")" (commas ys)) ^ " => (" ^ txt ^ ")";
 | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 11 | in | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 12 | (case xs of | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 13 | [] => ml_val_open_aux ys txt | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 14 |      | _  => ml_val_open_aux ys ("let open " ^ (space_implode " " xs) ^ " in " ^ txt ^ " end"))
 | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 15 | end; | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 16 | |
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 17 | fun ml_val txt = ml_val_open ([],[]) txt; | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 18 | |
| 112 | 19 | fun ml_pat (lhs, pat) = | 
| 51 
c346c156a7cd
completes the recipie on antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
49diff
changeset | 20 | let | 
| 
c346c156a7cd
completes the recipie on antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
49diff
changeset | 21 | val pat' = implode (map (fn "\\<dots>" => "_" | s => s) (Symbol.explode pat)) | 
| 112 | 22 | in "val " ^ pat' ^ " = " ^ lhs end; | 
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 23 | |
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 24 | fun ml_struct txt = "functor DUMMY_FUNCTOR() = struct structure DUMMY = " ^ txt ^ " end"; | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 25 | fun ml_type txt = "val _ = NONE : (" ^ txt ^ ") option";
 | 
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 26 | |
| 112 | 27 | (* eval function *) | 
| 28 | fun eval_fn ctxt pos exp = | |
| 29 | ML_Context.eval_in (SOME ctxt) false pos exp | |
| 57 
065f472c09ab
Repaired output of marginal comments in ML antiquotation.
 berghofe parents: 
54diff
changeset | 30 | |
| 112 | 31 | (* string functions *) | 
| 103 | 32 | fun string_explode str txt = | 
| 33 | map (fn s => str ^ s) (space_explode "\n" txt) | |
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 34 | |
| 112 | 35 | val transform_cmts_str = | 
| 36 | Source.of_string | |
| 37 | #> ML_Lex.source | |
| 38 | #> Source.exhaust | |
| 39 | #> Chunks.transform_cmts | |
| 40 | #> implode | |
| 41 | #> string_explode ""; | |
| 42 | ||
| 43 | (* parser for single and two arguments *) | |
| 44 | val single_arg = Scan.lift (OuterParse.position Args.name) | |
| 45 | val two_args = Scan.lift (OuterParse.position (Args.name -- Args.name)) | |
| 46 | ||
| 47 | (* output function *) | |
| 48 | val output_fn = Chunks.output_list (fn _ => fn s => Pretty.str s) | |
| 49 | ||
| 50 | (* checks and prints open expressions *) | |
| 51 | fun output_ml src node = | |
| 52 | let | |
| 53 | fun output src ctxt ((txt,ovars),pos) = | |
| 54 | (eval_fn ctxt pos (ml_val_open ovars txt); | |
| 55 | output_fn src ctxt (transform_cmts_str txt)) | |
| 56 | ||
| 57 | val parser = Scan.lift (OuterParse.position (Args.name -- | |
| 58 | (Scan.optional (Args.$$$ "for" |-- OuterParse.!!! (Scan.repeat1 Args.name)) [] -- | |
| 59 | Scan.optional (Args.$$$ "in" |-- OuterParse.!!! (Scan.repeat1 Args.name)) []))) | |
| 60 | in | |
| 61 | ThyOutput.args parser output src node | |
| 62 | end | |
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 63 | |
| 112 | 64 | (* checks and prints types and structures *) | 
| 65 | fun output_exp ml src node = | |
| 66 | let | |
| 67 | fun output src ctxt (txt,pos) = | |
| 68 | (eval_fn ctxt pos (ml txt); | |
| 69 | output_fn src ctxt (string_explode "" txt)) | |
| 70 | in | |
| 71 | ThyOutput.args single_arg output src node | |
| 72 | end | |
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 73 | |
| 112 | 74 | (* checks and expression agains a result pattern *) | 
| 75 | fun output_ml_response src node = | |
| 76 | let | |
| 77 | fun output src ctxt ((lhs,pat),pos) = | |
| 78 | (eval_fn ctxt pos (ml_pat (lhs,pat)); | |
| 79 | output_fn src ctxt ((string_explode "" lhs) @ (string_explode "> " pat))) | |
| 80 | in | |
| 81 | ThyOutput.args two_args output src node | |
| 82 | end | |
| 83 | ||
| 84 | (* checks the expressions, but does not check it against a result pattern *) | |
| 85 | fun output_ml_response_fake src node = | |
| 86 | let | |
| 87 | fun output src ctxt ((lhs,pat),pos) = | |
| 88 | (eval_fn ctxt pos (ml_val lhs); | |
| 89 | output_fn src ctxt ((string_explode "" lhs) @ (string_explode "> " pat))) | |
| 90 | in | |
| 91 | ThyOutput.args two_args output src node | |
| 92 | end | |
| 93 | ||
| 94 | (* just prints an expression and a result pattern *) | |
| 103 | 95 | fun output_ml_response_fake_both src node = | 
| 96 | let | |
| 112 | 97 | fun ouput src ctxt ((lhs,pat), _) = | 
| 98 | output_fn src ctxt ((string_explode "" lhs) @ (string_explode "> " pat)) | |
| 103 | 99 | in | 
| 112 | 100 | ThyOutput.args two_args ouput src node | 
| 103 | 101 | end | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
53diff
changeset | 102 | |
| 112 | 103 | (* checks whether a file exists in the Isabelle distribution *) | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 104 | fun check_file_exists src node = | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 105 | let | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 106 | fun check txt = | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 107 |    if File.exists (Path.append (Path.explode ("~~/src")) (Path.explode txt)) then ()
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 108 |    else error ("Source file " ^ (quote txt) ^ " does not exist.")
 | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 109 | in | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 110 | ThyOutput.args (Scan.lift Args.name) | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 111 | (ThyOutput.output (fn _ => fn s => (check s; Pretty.str s))) src node | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 112 | end | 
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 113 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 114 | (* replaces the official subgoal antiquotation with one *) | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 115 | (* that is closer to the actual output *) | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 116 | fun output_goals src node = | 
| 96 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 117 | let | 
| 98 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 118 | fun subgoals 0 = "" | 
| 96 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 119 | | subgoals 1 = "goal (1 subgoal):" | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 120 |     | subgoals n = "goal (" ^ string_of_int n ^ " subgoals):";
 | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 121 | |
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 122 | fun proof_state node = | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 123 | (case Option.map Toplevel.proof_node node of | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 124 | SOME (SOME prf) => ProofNode.current prf | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 125 | | _ => error "No proof state"); | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 126 | |
| 98 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 127 | val state = proof_state node; | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 128 | val goals = Proof.pretty_goals false state; | 
| 98 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 129 | |
| 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 130 |   val {prop, ...} = rep_thm (Proof.get_goal state |> snd |> snd);
 | 
| 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 131 | val (As, B) = Logic.strip_horn prop; | 
| 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 132 | val output = (case (length As) of | 
| 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 133 | 0 => goals | 
| 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 134 | | n => (Pretty.str (subgoals n))::goals) | 
| 96 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 135 | in | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 136 | ThyOutput.args (Scan.succeed ()) | 
| 98 
0a5c95f4d70c
calculated the exact number of goals in the subgoal antiquotation
 Christian Urban <urbanc@in.tum.de> parents: 
96diff
changeset | 137 | (Chunks.output (fn _ => fn _ => Pretty.chunks output)) src node | 
| 96 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 138 | end | 
| 
018bfa718982
(re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
 Christian Urban <urbanc@in.tum.de> parents: 
90diff
changeset | 139 | |
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 140 | val _ = ThyOutput.add_commands | 
| 112 | 141 |   [("ML", output_ml),
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 142 |    ("ML_file", check_file_exists),
 | 
| 112 | 143 |    ("ML_response", output_ml_response),
 | 
| 144 |    ("ML_response_fake", output_ml_response_fake),
 | |
| 103 | 145 |    ("ML_response_fake_both", output_ml_response_fake_both),
 | 
| 112 | 146 |    ("ML_struct", output_exp ml_struct),
 | 
| 147 |    ("ML_type", output_exp ml_type),
 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
98diff
changeset | 148 |    ("subgoals", output_goals)];
 | 
| 45 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 149 | |
| 
78aeca00bb54
deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 150 | end; |