src/HOL/ex/Antiquote.thy
author wenzelm
Wed, 14 Jun 2000 17:59:53 +0200
changeset 9066 b1e874e38dab
parent 8559 fd3753188232
child 9297 bafe45732b10
permissions -rw-r--r--
theorems [cases type: bool] = case_split;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/ex/Antiquote.thy
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     4
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     5
Simple quote / antiquote example.
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     6
*)
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     7
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
     8
theory Antiquote = Main:;
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     9
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    10
syntax
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    11
  "_Expr" :: "'a => 'a"				("EXPR _" [1000] 999);
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    12
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    13
constdefs
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    14
  var :: "'a => ('a => nat) => nat"		("VAR _" [1000] 999)
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    15
  "var x env == env x"
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    16
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    17
  Expr :: "(('a => nat) => nat) => ('a => nat) => nat"
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    18
  "Expr exp env == exp env";
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    19
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    20
parse_translation {* [Syntax.quote_antiquote_tr "_Expr" "var" "Expr"] *};
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    21
print_translation {* [Syntax.quote_antiquote_tr' "_Expr" "var" "Expr"] *};
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    22
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    23
term "EXPR (a + b + c)";
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    24
term "EXPR (a + b + c + VAR x + VAR y + 1)";
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    25
term "EXPR (VAR (f w) + VAR x)";
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    26
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    27
term "Expr (%env. env x)";				(*improper*)
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    28
term "Expr (%env. f env)";
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    29
term "Expr (%env. f env + env x)";			(*improper*)
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    30
term "Expr (%env. f env y z)";
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    31
term "Expr (%env. f env + g y env)";
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    32
term "Expr (%env. f env + g env y + h a env z)";
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    33
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    34
end;
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    35