src/HOL/ex/Antiquote.thy
author wenzelm
Sat, 05 Jan 2019 17:24:33 +0100
changeset 69597 ff784d5a5bfb
parent 61933 cf58b5b794b2
permissions -rw-r--r--
isabelle update -u control_cartouches;
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
    Author:     Markus Wenzel, TU Muenchen
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     3
*)
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
     4
59031
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
     5
section \<open>Antiquotations\<close>
10357
wenzelm
parents: 9297
diff changeset
     6
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 32960
diff changeset
     7
theory Antiquote
1a0c129bb2e0 modernized translations;
wenzelm
parents: 32960
diff changeset
     8
imports Main
1a0c129bb2e0 modernized translations;
wenzelm
parents: 32960
diff changeset
     9
begin
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    10
59031
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    11
text \<open>A simple example on quote / antiquote in higher-order abstract syntax.\<close>
11586
wenzelm
parents: 10357
diff changeset
    12
56233
797060c19f5c more symbols;
wenzelm
parents: 44603
diff changeset
    13
definition var :: "'a \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat"  ("VAR _" [1000] 999)
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 32960
diff changeset
    14
  where "var x env = env x"
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    15
56233
797060c19f5c more symbols;
wenzelm
parents: 44603
diff changeset
    16
definition Expr :: "(('a \<Rightarrow> nat) \<Rightarrow> nat) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> nat"
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 32960
diff changeset
    17
  where "Expr exp env = exp env"
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    18
44603
a6f9a70d655d tuned document;
wenzelm
parents: 42284
diff changeset
    19
syntax
56233
797060c19f5c more symbols;
wenzelm
parents: 44603
diff changeset
    20
  "_Expr" :: "'a \<Rightarrow> 'a"  ("EXPR _" [1000] 999)
44603
a6f9a70d655d tuned document;
wenzelm
parents: 42284
diff changeset
    21
59031
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    22
parse_translation
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    23
  \<open>[Syntax_Trans.quote_antiquote_tr
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61933
diff changeset
    24
    \<^syntax_const>\<open>_Expr\<close> \<^const_syntax>\<open>var\<close> \<^const_syntax>\<open>Expr\<close>]\<close>
35113
1a0c129bb2e0 modernized translations;
wenzelm
parents: 32960
diff changeset
    25
59031
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    26
print_translation
4c3bb56b8ce7 misc tuning and modernization;
wenzelm
parents: 58889
diff changeset
    27
  \<open>[Syntax_Trans.quote_antiquote_tr'
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 61933
diff changeset
    28
    \<^syntax_const>\<open>_Expr\<close> \<^const_syntax>\<open>var\<close> \<^const_syntax>\<open>Expr\<close>]\<close>
8559
fd3753188232 ex/Antiquote.thy made new-style theory;
wenzelm
parents: 5368
diff changeset
    29
9297
wenzelm
parents: 8559
diff changeset
    30
term "EXPR (a + b + c)"
wenzelm
parents: 8559
diff changeset
    31
term "EXPR (a + b + c + VAR x + VAR y + 1)"
wenzelm
parents: 8559
diff changeset
    32
term "EXPR (VAR (f w) + VAR x)"
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    33
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 59031
diff changeset
    34
term "Expr (\<lambda>env. env x)"    \<comment> \<open>improper\<close>
10357
wenzelm
parents: 9297
diff changeset
    35
term "Expr (\<lambda>env. f env)"
61933
cf58b5b794b2 isabelle update_cartouches -c -t;
wenzelm
parents: 59031
diff changeset
    36
term "Expr (\<lambda>env. f env + env x)"    \<comment> \<open>improper\<close>
10357
wenzelm
parents: 9297
diff changeset
    37
term "Expr (\<lambda>env. f env y z)"
wenzelm
parents: 9297
diff changeset
    38
term "Expr (\<lambda>env. f env + g y env)"
wenzelm
parents: 9297
diff changeset
    39
term "Expr (\<lambda>env. f env + g env y + h a env z)"
5368
7c8d1c7c876d added Antiquote example;
wenzelm
parents:
diff changeset
    40
9297
wenzelm
parents: 8559
diff changeset
    41
end