src/Doc/more_antiquote.ML
author wenzelm
Sat, 22 Oct 2016 20:09:30 +0200
changeset 64349 26bc905be09d
parent 63554 d7c6a3a01b79
child 67386 998e01d6f8fd
permissions -rw-r--r--
expose results on failure (via mail);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48985
5386df44a037 renamed doc-src to src/Doc;
wenzelm
parents: 43564
diff changeset
     1
(*  Title:      Doc/more_antiquote.ML
28440
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
     3
63026
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
     4
More antiquotations (partly depending on Isabelle/HOL).
28440
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
     5
*)
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
     6
56061
564a7bee8652 modernized setup;
wenzelm
parents: 54889
diff changeset
     7
structure More_Antiquote : sig end =
28440
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
     8
struct
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
     9
63026
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    10
(* class specifications *)
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    11
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    12
local
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    13
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    14
fun class_spec ctxt s =
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    15
  let
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    16
    val thy = Proof_Context.theory_of ctxt;
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    17
    val class = Sign.intern_class thy s;
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    18
  in Thy_Output.output ctxt (Class.pretty_specification thy class) end;
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    19
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    20
in
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    21
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    22
val _ =
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    23
  Theory.setup (Thy_Output.antiquotation @{binding class_spec} (Scan.lift Args.name)
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    24
    (fn {context, ...} => class_spec context));
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    25
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    26
end;
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    27
63554
d7c6a3a01b79 simplified
haftmann
parents: 63160
diff changeset
    28
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    29
(* code theorem antiquotation *)
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    30
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    31
local
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    32
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    33
fun no_vars ctxt thm =
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    34
  let
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    35
    val ctxt' = Variable.set_body false ctxt;
31794
71af1fd6a5e4 renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents: 31156
diff changeset
    36
    val ((_, [thm]), _) = Variable.import true [thm] ctxt';
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    37
  in thm end;
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    38
63554
d7c6a3a01b79 simplified
haftmann
parents: 63160
diff changeset
    39
fun pretty_code_thm ctxt raw_const =
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    40
  let
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 39540
diff changeset
    41
    val thy = Proof_Context.theory_of ctxt;
31156
90fed3d4430f merged module code_unit.ML into code.ML
haftmann
parents: 31143
diff changeset
    42
    val const = Code.check_const thy raw_const;
63160
80a91e0e236e corrected closure scope of static_conv_thingol;
haftmann
parents: 63073
diff changeset
    43
    val { eqngr, ... } = Code_Preproc.obtain true { ctxt = ctxt, consts = [const], terms = [] };
29874
0b92f68124e8 display code theorems with HOL equality
haftmann
parents: 29619
diff changeset
    44
    fun holize thm = @{thm meta_eq_to_obj_eq} OF [thm];
34896
a22b09addd78 adjusted to changes in code equation administration
haftmann
parents: 34072
diff changeset
    45
    val thms = Code_Preproc.cert eqngr const
35246
bcbb5ba7dbbc adjusted to changes in cs b987b803616d
haftmann
parents: 34896
diff changeset
    46
      |> Code.equations_of_cert thy
34896
a22b09addd78 adjusted to changes in code equation administration
haftmann
parents: 34072
diff changeset
    47
      |> snd
54889
4121d64fde90 explicit distinction between empty code equations and no code equations, including convenient declaration attributes
haftmann
parents: 51685
diff changeset
    48
      |> these
35246
bcbb5ba7dbbc adjusted to changes in cs b987b803616d
haftmann
parents: 34896
diff changeset
    49
      |> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
63073
413184c7a2a2 clarified context, notably for internal use of Simplifier;
wenzelm
parents: 63026
diff changeset
    50
      |> map (holize o no_vars ctxt o Axclass.overload ctxt);
63554
d7c6a3a01b79 simplified
haftmann
parents: 63160
diff changeset
    51
  in Thy_Output.output ctxt (map (Thy_Output.pretty_thm ctxt) thms) end;
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    52
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    53
in
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    54
56061
564a7bee8652 modernized setup;
wenzelm
parents: 54889
diff changeset
    55
val _ =
564a7bee8652 modernized setup;
wenzelm
parents: 54889
diff changeset
    56
  Theory.setup (Thy_Output.antiquotation @{binding code_thms} Args.term
63554
d7c6a3a01b79 simplified
haftmann
parents: 63160
diff changeset
    57
    (fn {context, ...} => pretty_code_thm context));
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    58
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    59
end;
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    60
28440
0b9ddfa6458e added more_antiquote.ML
haftmann
parents:
diff changeset
    61
end;