src/Doc/more_antiquote.ML
author wenzelm
Sat, 05 Jan 2019 17:24:33 +0100
changeset 69597 ff784d5a5bfb
parent 67710 cc2db3239932
child 70304 1514efa1e57a
permissions -rw-r--r--
isabelle update -u control_cartouches;
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
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
    13
  Theory.setup (Thy_Output.antiquotation_pretty \<^binding>\<open>class_spec\<close> (Scan.lift Args.name)
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    14
    (fn ctxt => fn s =>
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    15
      let
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    16
        val thy = Proof_Context.theory_of ctxt;
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    17
        val class = Sign.intern_class thy s;
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67469
diff changeset
    18
      in Pretty.chunks (Class.pretty_specification thy class) end));
63026
9a9c2d846d4a fragment of a HOL type class primer
haftmann
parents: 60697
diff changeset
    19
63554
d7c6a3a01b79 simplified
haftmann
parents: 63160
diff changeset
    20
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    21
(* code theorem antiquotation *)
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    22
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    23
fun no_vars ctxt thm =
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    24
  let
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    25
    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
    26
    val ((_, [thm]), _) = Variable.import true [thm] ctxt';
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    27
  in thm end;
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    28
56061
564a7bee8652 modernized setup;
wenzelm
parents: 54889
diff changeset
    29
val _ =
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67710
diff changeset
    30
  Theory.setup (Thy_Output.antiquotation_pretty \<^binding>\<open>code_thms\<close> Args.term
67463
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    31
    (fn ctxt => fn raw_const =>
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    32
      let
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    33
        val thy = Proof_Context.theory_of ctxt;
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    34
        val const = Code.check_const thy raw_const;
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    35
        val { eqngr, ... } = Code_Preproc.obtain true { ctxt = ctxt, consts = [const], terms = [] };
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    36
        val thms = Code_Preproc.cert eqngr const
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    37
          |> Code.equations_of_cert thy
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    38
          |> snd
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    39
          |> these
a5ca98950a91 clarified access to antiquotation options;
wenzelm
parents: 67386
diff changeset
    40
          |> map_filter (fn (_, (some_thm, proper)) => if proper then some_thm else NONE)
67710
cc2db3239932 added HOLogic.mk_obj_eq convenience and eliminated some clones;
wenzelm
parents: 67505
diff changeset
    41
          |> map (HOLogic.mk_obj_eq o no_vars ctxt o Axclass.overload ctxt);
67505
ceb324e34c14 clarified signature: items with \isasep are special;
wenzelm
parents: 67469
diff changeset
    42
      in Pretty.chunks (map (Thy_Output.pretty_thm ctxt) thms) end));
29397
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    43
aab26a65e80f dded code_thm antiquotation
haftmann
parents: 28921
diff changeset
    44
end;