CookBook/antiquote_setup.ML
author Christian Urban <urbanc@in.tum.de>
Thu, 12 Feb 2009 16:35:05 +0000
changeset 111 3798baeee55f
parent 106 bdd82350cf22
child 112 a90d0fb24e75
permissions -rw-r--r--
rearranged some functions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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: 103
diff 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
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 57
diff changeset
     6
(* main body *)
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
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
     8
  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
     9
          "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
    10
  in
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
    (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
    12
       [] => 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
    13
     | _  => 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
    14
  end;
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
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
    17
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
fun ml_pat (rhs, pat) =
51
c346c156a7cd completes the recipie on antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    19
  let 
c346c156a7cd completes the recipie on antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    20
     val pat' = implode (map (fn "\\<dots>" => "_" | s => s) (Symbol.explode pat))
c346c156a7cd completes the recipie on antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    21
  in "val " ^ pat' ^ " = " ^ rhs end;
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
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
    24
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
    25
57
065f472c09ab Repaired output of marginal comments in ML antiquotation.
berghofe
parents: 54
diff changeset
    26
val transform_cmts_str =
065f472c09ab Repaired output of marginal comments in ML antiquotation.
berghofe
parents: 54
diff changeset
    27
  Source.of_string #> ML_Lex.source #> Source.exhaust #>
065f472c09ab Repaired output of marginal comments in ML antiquotation.
berghofe
parents: 54
diff changeset
    28
  Chunks.transform_cmts #> implode;
065f472c09ab Repaired output of marginal comments in ML antiquotation.
berghofe
parents: 54
diff changeset
    29
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    30
fun output_ml ml src ctxt ((txt,ovars),pos) =
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  (ML_Context.eval_in (SOME ctxt) false pos (ml ovars txt);
57
065f472c09ab Repaired output of marginal comments in ML antiquotation.
berghofe
parents: 54
diff changeset
    32
   txt |> transform_cmts_str |> space_explode "\n" |>
90
b071a0b88298 made chunks aware of the gray-option
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
    33
   Chunks.output_list (fn _ => fn s => Pretty.str s) src ctxt)
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    35
fun output_ml_aux ml src ctxt (txt,pos) =
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    36
  output_ml (K ml) src ctxt ((txt,()),pos) 
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
103
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    38
fun string_explode str txt =
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    39
  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
    40
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
fun output_ml_response ml src ctxt ((lhs,pat),pos) = 
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  (ML_Context.eval_in (SOME ctxt) false pos (ml (lhs,pat));
51
c346c156a7cd completes the recipie on antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 49
diff changeset
    43
  let 
103
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    44
      val txt = (string_explode "" lhs) @ (string_explode "> " pat)
90
b071a0b88298 made chunks aware of the gray-option
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
    45
  in Chunks.output_list (fn _ => fn s => Pretty.str s) src ctxt txt end)
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
fun output_ml_response_fake ml src ctxt ((lhs,pat),pos) = 
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  (ML_Context.eval_in (SOME ctxt) false pos (ml lhs);
103
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    49
  let val txt = (string_explode "" lhs) @ (string_explode "> " pat)
90
b071a0b88298 made chunks aware of the gray-option
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
    50
  in Chunks.output_list (fn _ => fn s => Pretty.str s) src ctxt txt end)
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
103
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    52
fun output_ml_response_fake_both src node =
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    53
let 
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    54
  fun ouput src ctxt (lhs,pat) = 
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    55
    Chunks.output_list (fn _ => fn s => Pretty.str s) src ctxt 
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    56
     ((string_explode "" lhs) @ (string_explode "> " pat))
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    57
in
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    58
  ThyOutput.args (Scan.lift (Args.name -- Args.name)) ouput src node
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
    59
end
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 53
diff changeset
    60
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 98
diff changeset
    61
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: 98
diff changeset
    62
let 
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 98
diff changeset
    63
  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: 98
diff changeset
    64
   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: 98
diff changeset
    65
   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: 98
diff changeset
    66
in
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 98
diff changeset
    67
  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: 98
diff changeset
    68
      (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: 98
diff changeset
    69
end
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 98
diff changeset
    71
(* 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: 98
diff changeset
    72
(* 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: 98
diff changeset
    73
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: 90
diff changeset
    74
let 
98
0a5c95f4d70c calculated the exact number of goals in the subgoal antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 96
diff changeset
    75
  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: 90
diff changeset
    76
    | 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: 90
diff changeset
    77
    | 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: 90
diff changeset
    78
 
018bfa718982 (re)defined the document antiquotation @subgoal in order to be closer to what is actually printed
Christian Urban <urbanc@in.tum.de>
parents: 90
diff changeset
    79
  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: 90
diff changeset
    80
     (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: 90
diff changeset
    81
          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: 90
diff changeset
    82
        | _ => 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: 90
diff changeset
    83
98
0a5c95f4d70c calculated the exact number of goals in the subgoal antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 96
diff changeset
    84
  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: 98
diff changeset
    85
  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: 96
diff changeset
    86
0a5c95f4d70c calculated the exact number of goals in the subgoal antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 96
diff changeset
    87
  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: 96
diff changeset
    88
  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: 96
diff changeset
    89
  val output  = (case (length As) of
0a5c95f4d70c calculated the exact number of goals in the subgoal antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 96
diff changeset
    90
                      0 => goals 
0a5c95f4d70c calculated the exact number of goals in the subgoal antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 96
diff changeset
    91
                    | 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: 90
diff changeset
    92
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: 90
diff changeset
    93
  ThyOutput.args (Scan.succeed ()) 
98
0a5c95f4d70c calculated the exact number of goals in the subgoal antiquotation
Christian Urban <urbanc@in.tum.de>
parents: 96
diff changeset
    94
   (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: 90
diff changeset
    95
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: 90
diff changeset
    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: 90
diff changeset
    97
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
val _ = ThyOutput.add_commands
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
    99
  [("ML", ThyOutput.args (Scan.lift (OuterParse.position (Args.name --
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
      (Scan.optional (Args.$$$ "for" |-- OuterParse.!!! (Scan.repeat1 Args.name)) [] --
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
   101
       Scan.optional (Args.$$$ "in"  |-- OuterParse.!!! (Scan.repeat1 Args.name)) [])))) 
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
   102
                                                                 (output_ml ml_val_open)),
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 98
diff changeset
   103
   ("ML_file", check_file_exists),
45
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
   ("ML_response", ThyOutput.args (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
      (output_ml_response ml_pat)),
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
   ("ML_response_fake", ThyOutput.args (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
      (output_ml_response_fake ml_val)),
103
fe10da5354a3 further cleanup
Christian Urban <urbanc@in.tum.de>
parents: 102
diff changeset
   108
   ("ML_response_fake_both", output_ml_response_fake_both),
53
0c3580c831a4 removed the @{ML ...} antiquotation in favour of @{ML_open ...x}
Christian Urban <urbanc@in.tum.de>
parents: 51
diff changeset
   109
   ("ML_struct", ThyOutput.args (Scan.lift (OuterParse.position Args.name)) (output_ml_aux ml_struct)),
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: 90
diff changeset
   110
   ("ML_type", ThyOutput.args (Scan.lift (OuterParse.position Args.name)) (output_ml_aux 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: 98
diff changeset
   111
   ("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
   112
   
78aeca00bb54 deleted old antiquote_setup and renamed antiquote_setup_plus to antiquuote_setup
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
end;