src/Doc/Implementation/Syntax.thy
author wenzelm
Wed, 14 Oct 2015 15:10:32 +0200
changeset 61439 2bf52eec4e8a
parent 61416 b9a3324e4e62
child 61458 987533262fc2
permissions -rw-r--r--
more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
57344
wenzelm
parents: 56420
diff changeset
     1
(*:wrap=hard:maxLineLen=78:*)
wenzelm
parents: 56420
diff changeset
     2
30124
b956bf0dc87c basic setup for chapter "Syntax and type-checking";
wenzelm
parents:
diff changeset
     3
theory Syntax
b956bf0dc87c basic setup for chapter "Syntax and type-checking";
wenzelm
parents:
diff changeset
     4
imports Base
b956bf0dc87c basic setup for chapter "Syntax and type-checking";
wenzelm
parents:
diff changeset
     5
begin
b956bf0dc87c basic setup for chapter "Syntax and type-checking";
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
     7
chapter \<open>Concrete syntax and type-checking\<close>
30124
b956bf0dc87c basic setup for chapter "Syntax and type-checking";
wenzelm
parents:
diff changeset
     8
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
     9
text \<open>Pure @{text "\<lambda>"}-calculus as introduced in \chref{ch:logic} is
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    10
  an adequate foundation for logical languages --- in the tradition of
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    11
  \emph{higher-order abstract syntax} --- but end-users require
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    12
  additional means for reading and printing of terms and types.  This
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    13
  important add-on outside the logical core is called \emph{inner
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    14
  syntax} in Isabelle jargon, as opposed to the \emph{outer syntax} of
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
    15
  the theory and proof language @{cite "isabelle-isar-ref"}.
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    16
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
    17
  For example, according to @{cite church40} quantifiers are represented as
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    18
  higher-order constants @{text "All :: ('a \<Rightarrow> bool) \<Rightarrow> bool"} such that @{text
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    19
  "All (\<lambda>x::'a. B x)"} faithfully represents the idea that is displayed in
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    20
  Isabelle as @{text "\<forall>x::'a. B x"} via @{keyword "binder"} notation.
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
    21
  Moreover, type-inference in the style of Hindley-Milner @{cite hindleymilner}
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    22
  (and extensions) enables users to write @{text "\<forall>x. B x"} concisely, when
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    23
  the type @{text "'a"} is already clear from the
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    24
  context.\footnote{Type-inference taken to the extreme can easily confuse
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    25
  users. Beginners often stumble over unexpectedly general types inferred by
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    26
  the system.}
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    27
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
    28
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
    29
  The main inner syntax operations are \emph{read} for
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    30
  parsing together with type-checking, and \emph{pretty} for formatted
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    31
  output.  See also \secref{sec:read-print}.
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    32
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    33
  Furthermore, the input and output syntax layers are sub-divided into
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    34
  separate phases for \emph{concrete syntax} versus \emph{abstract
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    35
  syntax}, see also \secref{sec:parse-unparse} and
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    36
  \secref{sec:term-check}, respectively.  This results in the
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    37
  following decomposition of the main operations:
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    38
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    39
  \begin{itemize}
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    40
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
    41
  \<^item> @{text "read = parse; check"}
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    42
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
    43
  \<^item> @{text "pretty = uncheck; unparse"}
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    44
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    45
  \end{itemize}
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    46
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    47
  For example, some specification package might thus intercept syntax
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    48
  processing at a well-defined stage after @{text "parse"}, to a augment the
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    49
  resulting pre-term before full type-reconstruction is performed by @{text
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    50
  "check"}. Note that the formal status of bound variables, versus free
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    51
  variables, versus constants must not be changed between these phases.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    52
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
    53
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
    54
  In general, @{text check} and @{text uncheck} operate
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    55
  simultaneously on a list of terms. This is particular important for
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    56
  type-checking, to reconstruct types for several terms of the same context
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    57
  and scope. In contrast, @{text parse} and @{text unparse} operate separately
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    58
  on single terms.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    59
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    60
  There are analogous operations to read and print types, with the same
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    61
  sub-division into phases.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    62
\<close>
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
    63
30272
2d612824e642 regenerated document;
wenzelm
parents: 30270
diff changeset
    64
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    65
section \<open>Reading and pretty printing \label{sec:read-print}\<close>
34924
520727474bbe theory data example;
wenzelm
parents: 30272
diff changeset
    66
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    67
text \<open>
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    68
  Read and print operations are roughly dual to each other, such that for the
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    69
  user @{text "s' = pretty (read s)"} looks similar to the original source
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    70
  text @{text "s"}, but the details depend on many side-conditions. There are
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    71
  also explicit options to control the removal of type information in the
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    72
  output. The default configuration routinely looses information, so @{text
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    73
  "t' = read (pretty t)"} might fail, or produce a differently typed term, or
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
    74
  a completely different term in the face of syntactic overloading.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    75
\<close>
34924
520727474bbe theory data example;
wenzelm
parents: 30272
diff changeset
    76
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    77
text %mlref \<open>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    78
  \begin{mldecls}
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    79
  @{index_ML Syntax.read_typs: "Proof.context -> string list -> typ list"} \\
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    80
  @{index_ML Syntax.read_terms: "Proof.context -> string list -> term list"} \\
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    81
  @{index_ML Syntax.read_props: "Proof.context -> string list -> term list"} \\[0.5ex]
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    82
  @{index_ML Syntax.read_typ: "Proof.context -> string -> typ"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    83
  @{index_ML Syntax.read_term: "Proof.context -> string -> term"} \\
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    84
  @{index_ML Syntax.read_prop: "Proof.context -> string -> term"} \\[0.5ex]
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    85
  @{index_ML Syntax.pretty_typ: "Proof.context -> typ -> Pretty.T"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    86
  @{index_ML Syntax.pretty_term: "Proof.context -> term -> Pretty.T"} \\
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    87
  @{index_ML Syntax.string_of_typ: "Proof.context -> typ -> string"} \\
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    88
  @{index_ML Syntax.string_of_term: "Proof.context -> term -> string"} \\
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    89
  \end{mldecls}
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
    90
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    91
  \begin{description}
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    92
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
    93
  \<^descr> @{ML Syntax.read_typs}~@{text "ctxt strs"} parses and checks a
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    94
  simultaneous list of source strings as types of the logic.
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    95
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
    96
  \<^descr> @{ML Syntax.read_terms}~@{text "ctxt strs"} parses and checks a
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
    97
  simultaneous list of source strings as terms of the logic.
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
    98
  Type-reconstruction puts all parsed terms into the same scope: types of
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
    99
  free variables ultimately need to coincide.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   100
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   101
  If particular type-constraints are required for some of the arguments, the
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   102
  read operations needs to be split into its parse and check phases. Then it
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   103
  is possible to use @{ML Type.constraint} on the intermediate pre-terms
57846
wenzelm
parents: 57496
diff changeset
   104
  (\secref{sec:term-check}).
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   105
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   106
  \<^descr> @{ML Syntax.read_props}~@{text "ctxt strs"} parses and checks a
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   107
  simultaneous list of source strings as terms of the logic, with an implicit
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   108
  type-constraint for each argument to enforce type @{typ prop}; this also
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   109
  affects the inner syntax for parsing. The remaining type-reconstruction
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   110
  works as for @{ML Syntax.read_terms}.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   111
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   112
  \<^descr> @{ML Syntax.read_typ}, @{ML Syntax.read_term}, @{ML Syntax.read_prop}
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   113
  are like the simultaneous versions, but operate on a single argument only.
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   114
  This convenient shorthand is adequate in situations where a single item in
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   115
  its own scope is processed. Do not use @{ML "map o Syntax.read_term"} where
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   116
  @{ML Syntax.read_terms} is actually intended!
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   117
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   118
  \<^descr> @{ML Syntax.pretty_typ}~@{text "ctxt T"} and @{ML
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   119
  Syntax.pretty_term}~@{text "ctxt t"} uncheck and pretty-print the given type
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   120
  or term, respectively. Although the uncheck phase acts on a simultaneous
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   121
  list as well, this is rarely used in practice, so only the singleton case is
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   122
  provided as combined pretty operation. There is no distinction of term vs.\
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   123
  proposition.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   124
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   125
  \<^descr> @{ML Syntax.string_of_typ} and @{ML Syntax.string_of_term} are
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   126
  convenient compositions of @{ML Syntax.pretty_typ} and @{ML
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   127
  Syntax.pretty_term} with @{ML Pretty.string_of} for output. The result may
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   128
  be concatenated with other strings, as long as there is no further
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   129
  formatting and line-breaking involved.
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   130
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   131
  \end{description}
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   132
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   133
  @{ML Syntax.read_term}, @{ML Syntax.read_prop}, and @{ML
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   134
  Syntax.string_of_term} are the most important operations in practice.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   135
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
   136
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
   137
  Note that the string values that are passed in and out are
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   138
  annotated by the system, to carry further markup that is relevant for the
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   139
  Prover IDE @{cite "isabelle-jedit"}. User code should neither compose its
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   140
  own input strings, nor try to analyze the output strings. Conceptually this
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   141
  is an abstract datatype, encoded as concrete string for historical reasons.
57345
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   142
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   143
  The standard way to provide the required position markup for input works via
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   144
  the outer syntax parser wrapper @{ML Parse.inner_syntax}, which is already
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   145
  part of @{ML Parse.typ}, @{ML Parse.term}, @{ML Parse.prop}. So a string
8a9639888639 more on syntax phases;
wenzelm
parents: 57344
diff changeset
   146
  obtained from one of the latter may be directly passed to the corresponding
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   147
  read operation: this yields PIDE markup of the input and precise positions
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   148
  for warning and error messages.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   149
\<close>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   150
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   151
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   152
section \<open>Parsing and unparsing \label{sec:parse-unparse}\<close>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   153
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   154
text \<open>
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   155
  Parsing and unparsing converts between actual source text and a certain
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   156
  \emph{pre-term} format, where all bindings and scopes are already resolved
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   157
  faithfully. Thus the names of free variables or constants are determined in
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   158
  the sense of the logical context, but type information might be still
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   159
  missing. Pre-terms support an explicit language of \emph{type constraints}
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   160
  that may be augmented by user code to guide the later \emph{check} phase.
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   161
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   162
  Actual parsing is based on traditional lexical analysis and Earley parsing
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   163
  for arbitrary context-free grammars. The user can specify the grammar
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   164
  declaratively via mixfix annotations. Moreover, there are \emph{syntax
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   165
  translations} that can be augmented by the user, either declaratively via
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   166
  @{command translations} or programmatically via @{command
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   167
  parse_translation}, @{command print_translation} @{cite
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   168
  "isabelle-isar-ref"}. The final scope-resolution is performed by the system,
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   169
  according to name spaces for types, term variables and constants determined
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   170
  by the context.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   171
\<close>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   172
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   173
text %mlref \<open>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   174
  \begin{mldecls}
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   175
  @{index_ML Syntax.parse_typ: "Proof.context -> string -> typ"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   176
  @{index_ML Syntax.parse_term: "Proof.context -> string -> term"} \\
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   177
  @{index_ML Syntax.parse_prop: "Proof.context -> string -> term"} \\[0.5ex]
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   178
  @{index_ML Syntax.unparse_typ: "Proof.context -> typ -> Pretty.T"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   179
  @{index_ML Syntax.unparse_term: "Proof.context -> term -> Pretty.T"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   180
  \end{mldecls}
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   181
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   182
  \begin{description}
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   183
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   184
  \<^descr> @{ML Syntax.parse_typ}~@{text "ctxt str"} parses a source string as
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   185
  pre-type that is ready to be used with subsequent check operations.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   186
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   187
  \<^descr> @{ML Syntax.parse_term}~@{text "ctxt str"} parses a source string as
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   188
  pre-term that is ready to be used with subsequent check operations.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   189
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   190
  \<^descr> @{ML Syntax.parse_prop}~@{text "ctxt str"} parses a source string as
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   191
  pre-term that is ready to be used with subsequent check operations. The
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   192
  inner syntax category is @{typ prop} and a suitable type-constraint is
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   193
  included to ensure that this information is observed in subsequent type
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   194
  reconstruction.
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   195
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   196
  \<^descr> @{ML Syntax.unparse_typ}~@{text "ctxt T"} unparses a type after
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   197
  uncheck operations, to turn it into a pretty tree.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   198
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   199
  \<^descr> @{ML Syntax.unparse_term}~@{text "ctxt T"} unparses a term after
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   200
  uncheck operations, to turn it into a pretty tree. There is no distinction
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   201
  for propositions here.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   202
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   203
  \end{description}
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   204
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   205
  These operations always operate on a single item; use the combinator @{ML
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   206
  map} to apply them to a list.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   207
\<close>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   208
39852
9c977f899ebf tuned chapter arrangement;
wenzelm
parents: 34924
diff changeset
   209
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   210
section \<open>Checking and unchecking \label{sec:term-check}\<close>
34924
520727474bbe theory data example;
wenzelm
parents: 30272
diff changeset
   211
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   212
text \<open>These operations define the transition from pre-terms and
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   213
  fully-annotated terms in the sense of the logical core
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   214
  (\chref{ch:logic}).
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   215
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   216
  The \emph{check} phase is meant to subsume a variety of mechanisms
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   217
  in the manner of ``type-inference'' or ``type-reconstruction'' or
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   218
  ``type-improvement'', not just type-checking in the narrow sense.
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   219
  The \emph{uncheck} phase is roughly dual, it prunes type-information
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   220
  before pretty printing.
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   221
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   222
  A typical add-on for the check/uncheck syntax layer is the @{command
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   223
  abbreviation} mechanism @{cite "isabelle-isar-ref"}. Here the user specifies
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   224
  syntactic definitions that are managed by the system as polymorphic @{text
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   225
  "let"} bindings. These are expanded during the @{text "check"} phase, and
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   226
  contracted during the @{text "uncheck"} phase, without affecting the
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   227
  type-assignment of the given terms.
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   228
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
   229
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 58618
diff changeset
   230
  The precise meaning of type checking depends on the context ---
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   231
  additional check/uncheck modules might be defined in user space.
45258
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   232
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   233
  For example, the @{command class} command defines a context where
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   234
  @{text "check"} treats certain type instances of overloaded
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   235
  constants according to the ``dictionary construction'' of its
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   236
  logical foundation.  This involves ``type improvement''
97f8806c3ed6 some text on inner-syntax;
wenzelm
parents: 42510
diff changeset
   237
  (specialization of slightly too general types) and replacement by
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57846
diff changeset
   238
  certain locale parameters.  See also @{cite "Haftmann-Wenzel:2009"}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   239
\<close>
39852
9c977f899ebf tuned chapter arrangement;
wenzelm
parents: 34924
diff changeset
   240
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   241
text %mlref \<open>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   242
  \begin{mldecls}
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   243
  @{index_ML Syntax.check_typs: "Proof.context -> typ list -> typ list"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   244
  @{index_ML Syntax.check_terms: "Proof.context -> term list -> term list"} \\
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   245
  @{index_ML Syntax.check_props: "Proof.context -> term list -> term list"} \\[0.5ex]
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   246
  @{index_ML Syntax.uncheck_typs: "Proof.context -> typ list -> typ list"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   247
  @{index_ML Syntax.uncheck_terms: "Proof.context -> term list -> term list"} \\
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   248
  \end{mldecls}
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   249
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   250
  \begin{description}
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   251
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   252
  \<^descr> @{ML Syntax.check_typs}~@{text "ctxt Ts"} checks a simultaneous list
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   253
  of pre-types as types of the logic.  Typically, this involves normalization
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   254
  of type synonyms.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   255
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   256
  \<^descr> @{ML Syntax.check_terms}~@{text "ctxt ts"} checks a simultaneous list
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   257
  of pre-terms as terms of the logic. Typically, this involves type-inference
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   258
  and normalization term abbreviations. The types within the given terms are
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   259
  treated in the same way as for @{ML Syntax.check_typs}.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   260
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   261
  Applications sometimes need to check several types and terms together. The
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   262
  standard approach uses @{ML Logic.mk_type} to embed the language of types
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   263
  into that of terms; all arguments are appended into one list of terms that
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   264
  is checked; afterwards the type arguments are recovered with @{ML
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   265
  Logic.dest_type}.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   266
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   267
  \<^descr> @{ML Syntax.check_props}~@{text "ctxt ts"} checks a simultaneous list
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   268
  of pre-terms as terms of the logic, such that all terms are constrained by
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   269
  type @{typ prop}. The remaining check operation works as @{ML
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   270
  Syntax.check_terms} above.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   271
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   272
  \<^descr> @{ML Syntax.uncheck_typs}~@{text "ctxt Ts"} unchecks a simultaneous
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   273
  list of types of the logic, in preparation of pretty printing.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   274
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   275
  \<^descr> @{ML Syntax.uncheck_terms}~@{text "ctxt ts"} unchecks a simultaneous
57346
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   276
  list of terms of the logic, in preparation of pretty printing. There is no
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   277
  distinction for propositions here.
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   278
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   279
  \end{description}
1d6d44a0583f more on syntax phases;
wenzelm
parents: 57345
diff changeset
   280
57496
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   281
  These operations always operate simultaneously on a list; use the combinator
94596c573b38 misc tuning;
wenzelm
parents: 57346
diff changeset
   282
  @{ML singleton} to apply them to a single item.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   283
\<close>
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39865
diff changeset
   284
30124
b956bf0dc87c basic setup for chapter "Syntax and type-checking";
wenzelm
parents:
diff changeset
   285
end