src/HOL/ex/Multiquote.thy
author oheimb
Thu, 01 Feb 2001 20:51:48 +0100
changeset 11025 a70b796d9af8
parent 11020 646c929b6293
child 11586 d8a7f6318457
permissions -rw-r--r--
converted to Isar therory, adding attributes complete_split and split_format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/ex/Multiquote.thy
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
9297
wenzelm
parents: 8627
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     5
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     6
Multiple nested quotations and anti-quotations -- basically a
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     7
generalized version of de-Bruijn representation.
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     8
*)
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     9
10357
wenzelm
parents: 9297
diff changeset
    10
header {* Multiple nested quotations and anti-quotations *}
wenzelm
parents: 9297
diff changeset
    11
9297
wenzelm
parents: 8627
diff changeset
    12
theory Multiquote = Main:
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    13
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    14
syntax
8578
wenzelm
parents: 8572
diff changeset
    15
  "_quote" :: "'b => ('a => 'b)"	     (".'(_')." [0] 1000)
10838
9423817dee84 use \<acute>;
wenzelm
parents: 10834
diff changeset
    16
  "_antiquote" :: "('a => 'b) => 'b"         ("\<acute>_" [1000] 1000)
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    17
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    18
parse_translation {*
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    19
  let
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    20
    fun antiquote_tr i (Const ("_antiquote", _) $ (t as Const ("_antiquote", _) $ _)) =
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    21
          skip_antiquote_tr i t
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    22
      | antiquote_tr i (Const ("_antiquote", _) $ t) =
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    23
          antiquote_tr i t $ Bound i
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    24
      | antiquote_tr i (t $ u) = antiquote_tr i t $ antiquote_tr i u
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    25
      | antiquote_tr i (Abs (x, T, t)) = Abs (x, T, antiquote_tr (i + 1) t)
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    26
      | antiquote_tr _ a = a
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    27
    and skip_antiquote_tr i ((c as Const ("_antiquote", _)) $ t) =
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    28
          c $ skip_antiquote_tr i t
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    29
      | skip_antiquote_tr i t = antiquote_tr i t;
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    30
8578
wenzelm
parents: 8572
diff changeset
    31
    fun quote_tr [t] = Abs ("s", dummyT, antiquote_tr 0 (Term.incr_boundvars 1 t))
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    32
      | quote_tr ts = raise TERM ("quote_tr", ts);
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    33
  in [("_quote", quote_tr)] end
9297
wenzelm
parents: 8627
diff changeset
    34
*}
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    35
9297
wenzelm
parents: 8627
diff changeset
    36
text {* basic examples *}
wenzelm
parents: 8627
diff changeset
    37
term ".(a + b + c)."
10838
9423817dee84 use \<acute>;
wenzelm
parents: 10834
diff changeset
    38
term ".(a + b + c + \<acute>x + \<acute>y + 1)."
9423817dee84 use \<acute>;
wenzelm
parents: 10834
diff changeset
    39
term ".(\<acute>(f w) + \<acute>x)."
9423817dee84 use \<acute>;
wenzelm
parents: 10834
diff changeset
    40
term ".(f \<acute>x \<acute>y z)."
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    41
9297
wenzelm
parents: 8627
diff changeset
    42
text {* advanced examples *}
10838
9423817dee84 use \<acute>;
wenzelm
parents: 10834
diff changeset
    43
term ".(.(\<acute>\<acute>x + \<acute>y).)."
11020
wenzelm
parents: 10838
diff changeset
    44
term ".(.(\<acute>\<acute>x + \<acute>y). \<circ> \<acute>f)."
wenzelm
parents: 10838
diff changeset
    45
term ".(\<acute>(f \<circ> \<acute>g))."
wenzelm
parents: 10838
diff changeset
    46
term ".(.( \<acute>\<acute>(f \<circ> \<acute>g) ).)."
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    47
9297
wenzelm
parents: 8627
diff changeset
    48
end