src/HOL/ex/Multiquote.thy
author wenzelm
Thu, 29 Jun 2000 22:31:53 +0200
changeset 9195 29f1e53f9937
parent 8627 44ec33bb5c5b
child 9297 bafe45732b10
permissions -rw-r--r--
fixed is_semicolon (keyword instead of command!);
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
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     4
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     5
Multiple nested quotations and anti-quotations -- basically a
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     6
generalized version of de-Bruijn representation.
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     7
*)
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     8
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
     9
theory Multiquote = Main:;
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    10
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    11
syntax
8578
wenzelm
parents: 8572
diff changeset
    12
  "_quote" :: "'b => ('a => 'b)"	     (".'(_')." [0] 1000)
wenzelm
parents: 8572
diff changeset
    13
  "_antiquote" :: "('a => 'b) => 'b"         ("`_" [1000] 1000);
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    14
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    15
parse_translation {*
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    16
  let
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    17
    fun antiquote_tr i (Const ("_antiquote", _) $ (t as Const ("_antiquote", _) $ _)) =
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    18
          skip_antiquote_tr i t
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    19
      | antiquote_tr i (Const ("_antiquote", _) $ t) =
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    20
          antiquote_tr i t $ Bound i
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    21
      | antiquote_tr i (t $ u) = antiquote_tr i t $ antiquote_tr i u
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    22
      | 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
    23
      | antiquote_tr _ a = a
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    24
    and skip_antiquote_tr i ((c as Const ("_antiquote", _)) $ t) =
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    25
          c $ skip_antiquote_tr i t
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    26
      | skip_antiquote_tr i t = antiquote_tr i t;
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    27
8578
wenzelm
parents: 8572
diff changeset
    28
    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
    29
      | quote_tr ts = raise TERM ("quote_tr", ts);
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    30
  in [("_quote", quote_tr)] end
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    31
*};
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    32
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    33
text {* basic examples *};
8578
wenzelm
parents: 8572
diff changeset
    34
term ".(a + b + c).";
wenzelm
parents: 8572
diff changeset
    35
term ".(a + b + c + `x + `y + 1).";
wenzelm
parents: 8572
diff changeset
    36
term ".(`(f w) + `x).";
wenzelm
parents: 8572
diff changeset
    37
term ".(f `x `y z).";
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    38
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    39
text {* advanced examples *};
8578
wenzelm
parents: 8572
diff changeset
    40
term ".(.(` `x + `y).).";
wenzelm
parents: 8572
diff changeset
    41
term ".(.(` `x + `y). o `f).";
wenzelm
parents: 8572
diff changeset
    42
term ".(`(f o `g)).";
8627
wenzelm
parents: 8578
diff changeset
    43
term ".(.( ` `(f o `g) ).).";
8569
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    44
748a9699f28d added HOL/ex/Multiquote.thy;
wenzelm
parents:
diff changeset
    45
end;