src/Doc/Implementation/Prelim.thy
author wenzelm
Tue, 20 Oct 2015 23:53:40 +0200
changeset 61493 0debd22f0c0e
parent 61477 e467ae7aa808
child 61503 28e788ca2c5d
permissions -rw-r--r--
isabelle update_cartouches -t;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
29755
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29581
diff changeset
     1
theory Prelim
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29581
diff changeset
     2
imports Base
d66b34e46bdf observe usual theory naming conventions;
wenzelm
parents: 29581
diff changeset
     3
begin
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
     4
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
     5
chapter \<open>Preliminaries\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
     6
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
     7
section \<open>Contexts \label{sec:context}\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
     8
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
     9
text \<open>
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    10
  A logical context represents the background that is required for
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    11
  formulating statements and composing proofs.  It acts as a medium to
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    12
  produce formal content, depending on earlier material (declarations,
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    13
  results etc.).
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    14
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    15
  For example, derivations within the Isabelle/Pure logic can be
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    16
  described as a judgment \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<phi>\<close>, which means that a
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    17
  proposition \<open>\<phi>\<close> is derivable from hypotheses \<open>\<Gamma>\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    18
  within the theory \<open>\<Theta>\<close>.  There are logical reasons for
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    19
  keeping \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> separate: theories can be
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    20
  liberal about supporting type constructors and schematic
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    21
  polymorphism of constants and axioms, while the inner calculus of
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    22
  \<open>\<Gamma> \<turnstile> \<phi>\<close> is strictly limited to Simple Type Theory (with
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    23
  fixed type variables in the assumptions).
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    24
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    25
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    26
  Contexts and derivations are linked by the following key
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    27
  principles:
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    28
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    29
  \<^item> Transfer: monotonicity of derivations admits results to be
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    30
  transferred into a \<^emph>\<open>larger\<close> context, i.e.\ \<open>\<Gamma> \<turnstile>\<^sub>\<Theta>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    31
  \<phi>\<close> implies \<open>\<Gamma>' \<turnstile>\<^sub>\<Theta>\<^sub>' \<phi>\<close> for contexts \<open>\<Theta>'
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    32
  \<supseteq> \<Theta>\<close> and \<open>\<Gamma>' \<supseteq> \<Gamma>\<close>.
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    33
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    34
  \<^item> Export: discharge of hypotheses admits results to be exported
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    35
  into a \<^emph>\<open>smaller\<close> context, i.e.\ \<open>\<Gamma>' \<turnstile>\<^sub>\<Theta> \<phi>\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    36
  implies \<open>\<Gamma> \<turnstile>\<^sub>\<Theta> \<Delta> \<Longrightarrow> \<phi>\<close> where \<open>\<Gamma>' \<supseteq> \<Gamma>\<close> and
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    37
  \<open>\<Delta> = \<Gamma>' - \<Gamma>\<close>.  Note that \<open>\<Theta>\<close> remains unchanged here,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    38
  only the \<open>\<Gamma>\<close> part is affected.
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    39
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    40
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    41
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    42
  By modeling the main characteristics of the primitive
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    43
  \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> above, and abstracting over any
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    44
  particular logical content, we arrive at the fundamental notions of
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
    45
  \<^emph>\<open>theory context\<close> and \<^emph>\<open>proof context\<close> in Isabelle/Isar.
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
    46
  These implement a certain policy to manage arbitrary \<^emph>\<open>context
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
    47
  data\<close>.  There is a strongly-typed mechanism to declare new kinds of
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    48
  data at compile time.
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    49
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    50
  The internal bootstrap process of Isabelle/Pure eventually reaches a
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    51
  stage where certain data slots provide the logical content of \<open>\<Theta>\<close> and \<open>\<Gamma>\<close> sketched above, but this does not stop there!
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    52
  Various additional data slots support all kinds of mechanisms that
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    53
  are not necessarily part of the core logic.
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    54
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    55
  For example, there would be data for canonical introduction and
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    56
  elimination rules for arbitrary operators (depending on the
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    57
  object-logic and application), which enables users to perform
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    58
  standard proof steps implicitly (cf.\ the \<open>rule\<close> method
58555
7975676c08c0 prefer @{cite} antiquotation;
wenzelm
parents: 57421
diff changeset
    59
  @{cite "isabelle-isar-ref"}).
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    60
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    61
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    62
  Thus Isabelle/Isar is able to bring forth more and more
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    63
  concepts successively.  In particular, an object-logic like
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    64
  Isabelle/HOL continues the Isabelle/Pure setup by adding specific
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    65
  components for automated reasoning (classical reasoner, tableau
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    66
  prover, structured induction etc.) and derived specification
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    67
  mechanisms (inductive predicates, recursive functions etc.).  All of
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    68
  this is ultimately based on the generic data management by theory
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    69
  and proof contexts introduced here.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    70
\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    71
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    72
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
    73
subsection \<open>Theory context \label{sec:context-theory}\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    74
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
    75
text \<open>A \<^emph>\<open>theory\<close> is a data container with explicit name and
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
    76
  unique identifier.  Theories are related by a (nominal) sub-theory
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    77
  relation, which corresponds to the dependency graph of the original
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
    78
  construction; each theory is derived from a certain sub-graph of
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
    79
  ancestor theories.  To this end, the system maintains a set of
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
    80
  symbolic ``identification stamps'' within each theory.
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    81
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    82
  The \<open>begin\<close> operation starts a new theory by importing several
61046
6b97896d4946 tuned documentation -- merge is implicitly performed by the system;
wenzelm
parents: 60948
diff changeset
    83
  parent theories (with merged contents) and entering a special mode of
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    84
  nameless incremental updates, until the final \<open>end\<close> operation is
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
    85
  performed.
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
    86
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    87
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
    88
  The example in \figref{fig:ex-theory} below shows a theory
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    89
  graph derived from \<open>Pure\<close>, with theory \<open>Length\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    90
  importing \<open>Nat\<close> and \<open>List\<close>.  The body of \<open>Length\<close> consists of a sequence of updates, resulting in locally a
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52421
diff changeset
    91
  linear sub-theory relation for each intermediate step.
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
    92
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
    93
  \begin{figure}[htb]
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
    94
  \begin{center}
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
    95
  \begin{tabular}{rcccl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    96
        &            & \<open>Pure\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    97
        &            & \<open>\<down>\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
    98
        &            & \<open>FOL\<close> \\
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
    99
        & $\swarrow$ &              & $\searrow$ & \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   100
  \<open>Nat\<close> &    &              &            & \<open>List\<close> \\
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
   101
        & $\searrow$ &              & $\swarrow$ \\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   102
        &            & \<open>Length\<close> \\
26864
1417e704d724 replaced macros by antiquotations;
wenzelm
parents: 24137
diff changeset
   103
        &            & \multicolumn{3}{l}{~~@{keyword "begin"}} \\
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
   104
        &            & $\vdots$~~ \\
26864
1417e704d724 replaced macros by antiquotations;
wenzelm
parents: 24137
diff changeset
   105
        &            & \multicolumn{3}{l}{~~@{command "end"}} \\
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
   106
  \end{tabular}
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   107
  \caption{A theory definition depending on ancestors}\label{fig:ex-theory}
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   108
  \end{center}
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   109
  \end{figure}
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   110
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   111
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   112
  Derived formal entities may retain a reference to the
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52421
diff changeset
   113
  background theory in order to indicate the formal context from which
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52421
diff changeset
   114
  they were produced.  This provides an immutable certificate of the
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   115
  background theory.\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
   116
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   117
text %mlref \<open>
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   118
  \begin{mldecls}
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   119
  @{index_ML_type theory} \\
60948
b710a5087116 prefer theory_id operations;
wenzelm
parents: 59902
diff changeset
   120
  @{index_ML Context.eq_thy: "theory * theory -> bool"} \\
b710a5087116 prefer theory_id operations;
wenzelm
parents: 59902
diff changeset
   121
  @{index_ML Context.subthy: "theory * theory -> bool"} \\
48930
10b89c127153 updated Theory.begin_theory;
wenzelm
parents: 47174
diff changeset
   122
  @{index_ML Theory.begin_theory: "string * Position.T -> theory list -> theory"} \\
39837
eacb7825025d cover some more theory operations;
wenzelm
parents: 39833
diff changeset
   123
  @{index_ML Theory.parents_of: "theory -> theory list"} \\
eacb7825025d cover some more theory operations;
wenzelm
parents: 39833
diff changeset
   124
  @{index_ML Theory.ancestors_of: "theory -> theory list"} \\
20547
wenzelm
parents: 20530
diff changeset
   125
  \end{mldecls}
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   126
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   127
  \<^descr> Type @{ML_type theory} represents theory contexts.
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   128
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   129
  \<^descr> @{ML "Context.eq_thy"}~\<open>(thy\<^sub>1, thy\<^sub>2)\<close> check strict
39837
eacb7825025d cover some more theory operations;
wenzelm
parents: 39833
diff changeset
   130
  identity of two theories.
eacb7825025d cover some more theory operations;
wenzelm
parents: 39833
diff changeset
   131
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   132
  \<^descr> @{ML "Context.subthy"}~\<open>(thy\<^sub>1, thy\<^sub>2)\<close> compares theories
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   133
  according to the intrinsic graph structure of the construction.
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   134
  This sub-theory relation is a nominal approximation of inclusion
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   135
  (\<open>\<subseteq>\<close>) of the corresponding content (according to the
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   136
  semantics of the ML modules that implement the data).
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   137
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   138
  \<^descr> @{ML "Theory.begin_theory"}~\<open>name parents\<close> constructs
39825
f9066b94bf07 eliminated fancy \ML logo for the sake of simpler source text (less dependence on LaTeX);
wenzelm
parents: 39821
diff changeset
   139
  a new theory based on the given parents.  This ML function is
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   140
  normally not invoked directly.
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   141
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   142
  \<^descr> @{ML "Theory.parents_of"}~\<open>thy\<close> returns the direct
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   143
  ancestors of \<open>thy\<close>.
39837
eacb7825025d cover some more theory operations;
wenzelm
parents: 39833
diff changeset
   144
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   145
  \<^descr> @{ML "Theory.ancestors_of"}~\<open>thy\<close> returns all
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   146
  ancestors of \<open>thy\<close> (not including \<open>thy\<close> itself).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   147
\<close>
20430
wenzelm
parents: 20429
diff changeset
   148
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   149
text %mlantiq \<open>
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   150
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   151
  @{ML_antiquotation_def "theory"} & : & \<open>ML_antiquotation\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   152
  @{ML_antiquotation_def "theory_context"} & : & \<open>ML_antiquotation\<close> \\
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   153
  \end{matharray}
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   154
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 53015
diff changeset
   155
  @{rail \<open>
42510
b9c106763325 use @{rail} antiquotation (with some nested markup);
wenzelm
parents: 42401
diff changeset
   156
  @@{ML_antiquotation theory} nameref?
51686
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
   157
  ;
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
   158
  @@{ML_antiquotation theory_context} nameref
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 53015
diff changeset
   159
  \<close>}
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   160
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   161
  \<^descr> \<open>@{theory}\<close> refers to the background theory of the
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   162
  current context --- as abstract value.
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   163
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   164
  \<^descr> \<open>@{theory A}\<close> refers to an explicitly named ancestor
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   165
  theory \<open>A\<close> of the background theory of the current context
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   166
  --- as abstract value.
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   167
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   168
  \<^descr> \<open>@{theory_context A}\<close> is similar to \<open>@{theory
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   169
  A}\<close>, but presents the result as initial @{ML_type Proof.context}
51686
532e0ac5a66d added ML antiquotation @{theory_context};
wenzelm
parents: 48992
diff changeset
   170
  (see also @{ML Proof_Context.init_global}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   171
\<close>
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   172
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
   173
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   174
subsection \<open>Proof context \label{sec:context-proof}\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
   175
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   176
text \<open>A proof context is a container for pure data that refers to
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   177
  the theory from which it is derived. The \<open>init\<close> operation
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52421
diff changeset
   178
  creates a proof context from a given theory. There is an explicit
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   179
  \<open>transfer\<close> operation to force resynchronization with updates
52788
da1fdbfebd39 type theory is purely value-oriented;
wenzelm
parents: 52421
diff changeset
   180
  to the background theory -- this is rarely required in practice.
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
   181
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   182
  Entities derived in a proof context need to record logical
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   183
  requirements explicitly, since there is no separate context
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   184
  identification or symbolic inclusion as for theories.  For example,
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   185
  hypotheses used in primitive derivations (cf.\ \secref{sec:thms})
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   186
  are recorded separately within the sequent \<open>\<Gamma> \<turnstile> \<phi>\<close>, just to
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   187
  make double sure.  Results could still leak into an alien proof
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   188
  context due to programming errors, but Isabelle/Isar includes some
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   189
  extra validity checks in critical positions, notably at the end of a
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   190
  sub-proof.
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
   191
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   192
  Proof contexts may be manipulated arbitrarily, although the common
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   193
  discipline is to follow block structure as a mental model: a given
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   194
  context is extended consecutively, and results are exported back
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   195
  into the original context.  Note that an Isar proof state models
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   196
  block-structured reasoning explicitly, using a stack of proof
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   197
  contexts internally.  For various technical reasons, the background
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   198
  theory of an Isar proof state must not be changed while the proof is
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   199
  still under construction!
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   200
\<close>
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
   201
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   202
text %mlref \<open>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   203
  \begin{mldecls}
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   204
  @{index_ML_type Proof.context} \\
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   205
  @{index_ML Proof_Context.init_global: "theory -> Proof.context"} \\
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   206
  @{index_ML Proof_Context.theory_of: "Proof.context -> theory"} \\
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   207
  @{index_ML Proof_Context.transfer: "theory -> Proof.context -> Proof.context"} \\
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   208
  \end{mldecls}
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   209
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   210
  \<^descr> Type @{ML_type Proof.context} represents proof contexts.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   211
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   212
  \<^descr> @{ML Proof_Context.init_global}~\<open>thy\<close> produces a proof
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   213
  context derived from \<open>thy\<close>, initializing all data.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   214
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   215
  \<^descr> @{ML Proof_Context.theory_of}~\<open>ctxt\<close> selects the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   216
  background theory from \<open>ctxt\<close>.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   217
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   218
  \<^descr> @{ML Proof_Context.transfer}~\<open>thy ctxt\<close> promotes the
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   219
  background theory of \<open>ctxt\<close> to the super theory \<open>thy\<close>.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   220
\<close>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   221
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   222
text %mlantiq \<open>
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   223
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   224
  @{ML_antiquotation_def "context"} & : & \<open>ML_antiquotation\<close> \\
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   225
  \end{matharray}
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   226
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   227
  \<^descr> \<open>@{context}\<close> refers to \<^emph>\<open>the\<close> context at
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   228
  compile-time --- as abstract value.  Independently of (local) theory
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   229
  or proof mode, this always produces a meaningful result.
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   230
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   231
  This is probably the most common antiquotation in interactive
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   232
  experimentation with ML inside Isar.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   233
\<close>
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   234
20430
wenzelm
parents: 20429
diff changeset
   235
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   236
subsection \<open>Generic contexts \label{sec:generic-context}\<close>
20429
116255c9209b more on contexts;
wenzelm
parents: 20214
diff changeset
   237
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   238
text \<open>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   239
  A generic context is the disjoint sum of either a theory or proof
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   240
  context.  Occasionally, this enables uniform treatment of generic
20450
wenzelm
parents: 20449
diff changeset
   241
  context data, typically extra-logical information.  Operations on
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   242
  generic contexts include the usual injections, partial selections,
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   243
  and combinators for lifting operations on either component of the
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   244
  disjoint sum.
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   245
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   246
  Moreover, there are total operations \<open>theory_of\<close> and \<open>proof_of\<close> to convert a generic context into either kind: a theory
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   247
  can always be selected from the sum, while a proof context might
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   248
  have to be constructed by an ad-hoc \<open>init\<close> operation, which
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   249
  incurs a small runtime overhead.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   250
\<close>
20430
wenzelm
parents: 20429
diff changeset
   251
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   252
text %mlref \<open>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   253
  \begin{mldecls}
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   254
  @{index_ML_type Context.generic} \\
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   255
  @{index_ML Context.theory_of: "Context.generic -> theory"} \\
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   256
  @{index_ML Context.proof_of: "Context.generic -> Proof.context"} \\
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   257
  \end{mldecls}
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   258
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   259
  \<^descr> Type @{ML_type Context.generic} is the direct sum of @{ML_type
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   260
  "theory"} and @{ML_type "Proof.context"}, with the datatype
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   261
  constructors @{ML "Context.Theory"} and @{ML "Context.Proof"}.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   262
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   263
  \<^descr> @{ML Context.theory_of}~\<open>context\<close> always produces a
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   264
  theory from the generic \<open>context\<close>, using @{ML
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   265
  "Proof_Context.theory_of"} as required.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   266
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   267
  \<^descr> @{ML Context.proof_of}~\<open>context\<close> always produces a
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   268
  proof context from the generic \<open>context\<close>, using @{ML
42361
23f352990944 modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   269
  "Proof_Context.init_global"} as required (note that this re-initializes the
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   270
  context data with each invocation).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   271
\<close>
20437
wenzelm
parents: 20430
diff changeset
   272
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   273
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   274
subsection \<open>Context data \label{sec:context-data}\<close>
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   275
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   276
text \<open>The main purpose of theory and proof contexts is to manage
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   277
  arbitrary (pure) data.  New data types can be declared incrementally
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   278
  at compile time.  There are separate declaration mechanisms for any
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   279
  of the three kinds of contexts: theory, proof, generic.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   280
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   281
  \paragraph{Theory data} declarations need to implement the following
57421
94081154306d misc tuning;
wenzelm
parents: 56420
diff changeset
   282
  ML signature:
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   283
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   284
  \<^medskip>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   285
  \begin{tabular}{ll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   286
  \<open>\<type> T\<close> & representing type \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   287
  \<open>\<val> empty: T\<close> & empty default value \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   288
  \<open>\<val> extend: T \<rightarrow> T\<close> & re-initialize on import \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   289
  \<open>\<val> merge: T \<times> T \<rightarrow> T\<close> & join on import \\
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   290
  \end{tabular}
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   291
  \<^medskip>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   292
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   293
  The \<open>empty\<close> value acts as initial default for \<^emph>\<open>any\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   294
  theory that does not declare actual data content; \<open>extend\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   295
  is acts like a unitary version of \<open>merge\<close>.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   296
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   297
  Implementing \<open>merge\<close> can be tricky.  The general idea is
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   298
  that \<open>merge (data\<^sub>1, data\<^sub>2)\<close> inserts those parts of \<open>data\<^sub>2\<close> into \<open>data\<^sub>1\<close> that are not yet present, while
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   299
  keeping the general order of things.  The @{ML Library.merge}
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   300
  function on plain lists may serve as canonical template.
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   301
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   302
  Particularly note that shared parts of the data must not be
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   303
  duplicated by naive concatenation, or a theory graph that is like a
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   304
  chain of diamonds would cause an exponential blowup!
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   305
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   306
  \paragraph{Proof context data} declarations need to implement the
57421
94081154306d misc tuning;
wenzelm
parents: 56420
diff changeset
   307
  following ML signature:
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   308
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   309
  \<^medskip>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   310
  \begin{tabular}{ll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   311
  \<open>\<type> T\<close> & representing type \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   312
  \<open>\<val> init: theory \<rightarrow> T\<close> & produce initial value \\
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   313
  \end{tabular}
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   314
  \<^medskip>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   315
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   316
  The \<open>init\<close> operation is supposed to produce a pure value
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   317
  from the given background theory and should be somehow
34921
008126f730a0 formal markup of type aliases;
wenzelm
parents: 33524
diff changeset
   318
  ``immediate''.  Whenever a proof context is initialized, which
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   319
  happens frequently, the the system invokes the \<open>init\<close>
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   320
  operation of \<^emph>\<open>all\<close> theory data slots ever declared.  This also
39821
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   321
  means that one needs to be economic about the total number of proof
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   322
  data declarations in the system, i.e.\ each ML module should declare
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   323
  at most one, sometimes two data slots for its internal use.
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   324
  Repeated data declarations to simulate a record type should be
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   325
  avoided!
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   326
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   327
  \paragraph{Generic data} provides a hybrid interface for both theory
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   328
  and proof data.  The \<open>init\<close> operation for proof contexts is
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   329
  predefined to select the current data value from the background
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   330
  theory.
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   331
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   332
  \<^bigskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   333
  Any of the above data declarations over type \<open>T\<close>
39821
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   334
  result in an ML structure with the following signature:
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   335
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   336
  \<^medskip>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   337
  \begin{tabular}{ll}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   338
  \<open>get: context \<rightarrow> T\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   339
  \<open>put: T \<rightarrow> context \<rightarrow> context\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   340
  \<open>map: (T \<rightarrow> T) \<rightarrow> context \<rightarrow> context\<close> \\
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   341
  \end{tabular}
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   342
  \<^medskip>
20449
f8a7a8236c68 more stuff;
wenzelm
parents: 20447
diff changeset
   343
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   344
  These other operations provide exclusive access for the particular
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   345
  kind of context (theory, proof, or generic context).  This interface
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   346
  observes the ML discipline for types and scopes: there is no other
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   347
  way to access the corresponding data slot of a context.  By keeping
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   348
  these operations private, an Isabelle/ML module may maintain
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   349
  abstract values authentically.\<close>
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   350
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   351
text %mlref \<open>
20450
wenzelm
parents: 20449
diff changeset
   352
  \begin{mldecls}
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   353
  @{index_ML_functor Theory_Data} \\
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   354
  @{index_ML_functor Proof_Data} \\
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   355
  @{index_ML_functor Generic_Data} \\
20450
wenzelm
parents: 20449
diff changeset
   356
  \end{mldecls}
wenzelm
parents: 20449
diff changeset
   357
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   358
  \<^descr> @{ML_functor Theory_Data}\<open>(spec)\<close> declares data for
20450
wenzelm
parents: 20449
diff changeset
   359
  type @{ML_type theory} according to the specification provided as
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   360
  argument structure.  The resulting structure provides data init and
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   361
  access operations as described above.
20450
wenzelm
parents: 20449
diff changeset
   362
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   363
  \<^descr> @{ML_functor Proof_Data}\<open>(spec)\<close> is analogous to
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   364
  @{ML_functor Theory_Data} for type @{ML_type Proof.context}.
20450
wenzelm
parents: 20449
diff changeset
   365
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   366
  \<^descr> @{ML_functor Generic_Data}\<open>(spec)\<close> is analogous to
33524
a08e6c1cbc04 updated functor Theory_Data, Proof_Data, Generic_Data;
wenzelm
parents: 33174
diff changeset
   367
  @{ML_functor Theory_Data} for type @{ML_type Context.generic}.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   368
\<close>
20450
wenzelm
parents: 20449
diff changeset
   369
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   370
text %mlex \<open>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   371
  The following artificial example demonstrates theory
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   372
  data: we maintain a set of terms that are supposed to be wellformed
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   373
  wrt.\ the enclosing theory.  The public interface is as follows:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   374
\<close>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   375
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   376
ML \<open>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   377
  signature WELLFORMED_TERMS =
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   378
  sig
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   379
    val get: theory -> term list
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   380
    val add: term -> theory -> theory
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   381
  end;
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   382
\<close>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   383
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   384
text \<open>The implementation uses private theory data internally, and
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   385
  only exposes an operation that involves explicit argument checking
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   386
  wrt.\ the given theory.\<close>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   387
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   388
ML \<open>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   389
  structure Wellformed_Terms: WELLFORMED_TERMS =
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   390
  struct
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   391
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   392
  structure Terms = Theory_Data
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   393
  (
39687
4e9b6ada3a21 modernized structure Ord_List;
wenzelm
parents: 37533
diff changeset
   394
    type T = term Ord_List.T;
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   395
    val empty = [];
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   396
    val extend = I;
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   397
    fun merge (ts1, ts2) =
39687
4e9b6ada3a21 modernized structure Ord_List;
wenzelm
parents: 37533
diff changeset
   398
      Ord_List.union Term_Ord.fast_term_ord ts1 ts2;
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   399
  );
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   400
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   401
  val get = Terms.get;
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   402
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   403
  fun add raw_t thy =
39821
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   404
    let
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   405
      val t = Sign.cert_term thy raw_t;
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   406
    in
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   407
      Terms.map (Ord_List.insert Term_Ord.fast_term_ord t) thy
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   408
    end;
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   409
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   410
  end;
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   411
\<close>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   412
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   413
text \<open>Type @{ML_type "term Ord_List.T"} is used for reasonably
39864
wenzelm
parents: 39863
diff changeset
   414
  efficient representation of a set of terms: all operations are
wenzelm
parents: 39863
diff changeset
   415
  linear in the number of stored elements.  Here we assume that users
wenzelm
parents: 39863
diff changeset
   416
  of this module do not care about the declaration order, since that
wenzelm
parents: 39863
diff changeset
   417
  data structure forces its own arrangement of elements.
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   418
40153
wenzelm
parents: 40126
diff changeset
   419
  Observe how the @{ML_text merge} operation joins the data slots of
39687
4e9b6ada3a21 modernized structure Ord_List;
wenzelm
parents: 37533
diff changeset
   420
  the two constituents: @{ML Ord_List.union} prevents duplication of
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   421
  common data from different branches, thus avoiding the danger of
39821
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   422
  exponential blowup.  Plain list append etc.\ must never be used for
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   423
  theory data merges!
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   424
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   425
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   426
  Our intended invariant is achieved as follows:
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   427
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   428
  \<^enum> @{ML Wellformed_Terms.add} only admits terms that have passed
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   429
  the @{ML Sign.cert_term} check of the given theory at that point.
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   430
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   431
  \<^enum> Wellformedness in the sense of @{ML Sign.cert_term} is
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   432
  monotonic wrt.\ the sub-theory relation.  So our data can move
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   433
  upwards in the hierarchy (via extension or merges), and maintain
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   434
  wellformedness without further checks.
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   435
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   436
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   437
  Note that all basic operations of the inference kernel (which
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   438
  includes @{ML Sign.cert_term}) observe this monotonicity principle,
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   439
  but other user-space tools don't.  For example, fully-featured
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   440
  type-inference via @{ML Syntax.check_term} (cf.\
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   441
  \secref{sec:term-check}) is not necessarily monotonic wrt.\ the
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   442
  background theory, since constraints of term constants can be
39821
bf164c153d10 minor tuning and updating;
wenzelm
parents: 39687
diff changeset
   443
  modified by later declarations, for example.
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   444
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   445
  In most cases, user-space context data does not have to take such
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   446
  invariants too seriously.  The situation is different in the
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   447
  implementation of the inference kernel itself, which uses the very
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   448
  same data mechanisms for types, constants, axioms etc.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   449
\<close>
34928
c4cd36411983 moved examples to proper place;
wenzelm
parents: 34927
diff changeset
   450
20447
5b75f1c4d7d6 more on contexts;
wenzelm
parents: 20437
diff changeset
   451
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   452
subsection \<open>Configuration options \label{sec:config-options}\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   453
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   454
text \<open>A \<^emph>\<open>configuration option\<close> is a named optional value of
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   455
  some basic type (Boolean, integer, string) that is stored in the
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   456
  context.  It is a simple application of general context data
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   457
  (\secref{sec:context-data}) that is sufficiently common to justify
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   458
  customized setup, which includes some concrete declarations for
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   459
  end-users using existing notation for attributes (cf.\
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   460
  \secref{sec:attributes}).
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   461
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   462
  For example, the predefined configuration option @{attribute
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   463
  show_types} controls output of explicit type constraints for
39876
1ff9bce085bd preliminary material on "Concrete syntax and type-checking";
wenzelm
parents: 39866
diff changeset
   464
  variables in printed terms (cf.\ \secref{sec:read-print}).  Its
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   465
  value can be modified within Isar text like this:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   466
\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   467
59902
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   468
experiment
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   469
begin
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   470
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   471
declare [[show_types = false]]
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   472
  -- \<open>declaration within (local) theory context\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   473
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40628
diff changeset
   474
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40628
diff changeset
   475
begin
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   476
  note [[show_types = true]]
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   477
    -- \<open>declaration within proof (forward mode)\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   478
  term x
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   479
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   480
  have "x = x"
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   481
    using [[show_types = false]]
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   482
      -- \<open>declaration within proof (backward mode)\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   483
    ..
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40628
diff changeset
   484
end
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   485
59902
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   486
end
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   487
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   488
text \<open>Configuration options that are not set explicitly hold a
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   489
  default value that can depend on the application context.  This
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   490
  allows to retrieve the value from another slot within the context,
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   491
  or fall back on a global preference mechanism, for example.
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   492
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   493
  The operations to declare configuration options and get/map their
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   494
  values are modeled as direct replacements for historic global
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   495
  references, only that the context is made explicit.  This allows
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   496
  easy configuration of tools, without relying on the execution order
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   497
  as required for old-style mutable references.\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   498
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   499
text %mlref \<open>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   500
  \begin{mldecls}
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   501
  @{index_ML Config.get: "Proof.context -> 'a Config.T -> 'a"} \\
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   502
  @{index_ML Config.map: "'a Config.T -> ('a -> 'a) -> Proof.context -> Proof.context"} \\
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   503
  @{index_ML Attrib.setup_config_bool: "binding -> (Context.generic -> bool) ->
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   504
  bool Config.T"} \\
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   505
  @{index_ML Attrib.setup_config_int: "binding -> (Context.generic -> int) ->
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   506
  int Config.T"} \\
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   507
  @{index_ML Attrib.setup_config_real: "binding -> (Context.generic -> real) ->
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   508
  real Config.T"} \\
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   509
  @{index_ML Attrib.setup_config_string: "binding -> (Context.generic -> string) ->
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   510
  string Config.T"} \\
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   511
  \end{mldecls}
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   512
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   513
  \<^descr> @{ML Config.get}~\<open>ctxt config\<close> gets the value of
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   514
  \<open>config\<close> in the given context.
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   515
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   516
  \<^descr> @{ML Config.map}~\<open>config f ctxt\<close> updates the context
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   517
  by updating the value of \<open>config\<close>.
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   518
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   519
  \<^descr> \<open>config =\<close>~@{ML Attrib.setup_config_bool}~\<open>name
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   520
  default\<close> creates a named configuration option of type @{ML_type
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   521
  bool}, with the given \<open>default\<close> depending on the application
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   522
  context.  The resulting \<open>config\<close> can be used to get/map its
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   523
  value in a given context.  There is an implicit update of the
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   524
  background theory that registers the option as attribute with some
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   525
  concrete syntax.
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   526
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   527
  \<^descr> @{ML Attrib.config_int}, @{ML Attrib.config_real}, and @{ML
40291
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40241
diff changeset
   528
  Attrib.config_string} work like @{ML Attrib.config_bool}, but for
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40241
diff changeset
   529
  types @{ML_type int} and @{ML_type string}, respectively.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   530
\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   531
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   532
text %mlex \<open>The following example shows how to declare and use a
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   533
  Boolean configuration option called \<open>my_flag\<close> with constant
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   534
  default value @{ML false}.\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   535
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   536
ML \<open>
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   537
  val my_flag =
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   538
    Attrib.setup_config_bool @{binding my_flag} (K false)
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   539
\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   540
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   541
text \<open>Now the user can refer to @{attribute my_flag} in
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 39876
diff changeset
   542
  declarations, while ML tools can retrieve the current value from the
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   543
  context via @{ML Config.get}.\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   544
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   545
ML_val \<open>@{assert} (Config.get @{context} my_flag = false)\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   546
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   547
declare [[my_flag = true]]
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   548
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   549
ML_val \<open>@{assert} (Config.get @{context} my_flag = true)\<close>
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   550
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40628
diff changeset
   551
notepad
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40628
diff changeset
   552
begin
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   553
  {
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   554
    note [[my_flag = false]]
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   555
    ML_val \<open>@{assert} (Config.get @{context} my_flag = false)\<close>
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   556
  }
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   557
  ML_val \<open>@{assert} (Config.get @{context} my_flag = true)\<close>
40964
482a8334ee9e prefer 'notepad' over 'example_proof';
wenzelm
parents: 40628
diff changeset
   558
end
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   559
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   560
text \<open>Here is another example involving ML type @{ML_type real}
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   561
  (floating-point numbers).\<close>
40291
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40241
diff changeset
   562
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   563
ML \<open>
42616
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   564
  val airspeed_velocity =
92715b528e78 added Attrib.setup_config_XXX conveniences, with implicit setup of the background theory;
wenzelm
parents: 42510
diff changeset
   565
    Attrib.setup_config_real @{binding airspeed_velocity} (K 0.0)
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   566
\<close>
40291
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40241
diff changeset
   567
40296
ac4d75f86d97 syntax category "real" subsumes plain "int";
wenzelm
parents: 40291
diff changeset
   568
declare [[airspeed_velocity = 10]]
40291
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40241
diff changeset
   569
declare [[airspeed_velocity = 9.9]]
012ed4426fda support for real valued configuration options;
wenzelm
parents: 40241
diff changeset
   570
39865
a724b90f951e more on "Configuration options";
wenzelm
parents: 39864
diff changeset
   571
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   572
section \<open>Names \label{sec:names}\<close>
20451
27ea2ba48fa3 misc cleanup;
wenzelm
parents: 20450
diff changeset
   573
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   574
text \<open>In principle, a name is just a string, but there are various
34925
38a44d813a3c more details on Isabelle symbols;
wenzelm
parents: 34924
diff changeset
   575
  conventions for representing additional structure.  For example,
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   576
  ``\<open>Foo.bar.baz\<close>'' is considered as a long name consisting of
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   577
  qualifier \<open>Foo.bar\<close> and base name \<open>baz\<close>.  The
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   578
  individual constituents of a name may have further substructure,
58723
33be43d70147 more antiquotations;
wenzelm
parents: 58668
diff changeset
   579
  e.g.\ the string ``@{verbatim \<alpha>}'' encodes as a single
52421
6d93140a206c clarified strings of symbols, including ML string literals;
wenzelm
parents: 51686
diff changeset
   580
  symbol (\secref{sec:symbols}).
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   581
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   582
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   583
  Subsequently, we shall introduce specific categories of
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   584
  names.  Roughly speaking these correspond to logical entities as
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   585
  follows:
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   586
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   587
  \<^item> Basic names (\secref{sec:basic-name}): free and bound
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   588
  variables.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   589
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   590
  \<^item> Indexed names (\secref{sec:indexname}): schematic variables.
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   591
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   592
  \<^item> Long names (\secref{sec:long-name}): constants of any kind
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   593
  (type constructors, term constants, other concepts defined in user
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   594
  space).  Such entities are typically managed via name spaces
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   595
  (\secref{sec:name-space}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   596
\<close>
20437
wenzelm
parents: 20430
diff changeset
   597
wenzelm
parents: 20430
diff changeset
   598
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   599
subsection \<open>Basic names \label{sec:basic-name}\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   600
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   601
text \<open>
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   602
  A \<^emph>\<open>basic name\<close> essentially consists of a single Isabelle
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   603
  identifier.  There are conventions to mark separate classes of basic
29761
2b658e50683a minor tuning and typographic fixes;
wenzelm
parents: 29758
diff changeset
   604
  names, by attaching a suffix of underscores: one underscore means
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   605
  \<^emph>\<open>internal name\<close>, two underscores means \<^emph>\<open>Skolem name\<close>,
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   606
  three underscores means \<^emph>\<open>internal Skolem name\<close>.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   607
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   608
  For example, the basic name \<open>foo\<close> has the internal version
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   609
  \<open>foo_\<close>, with Skolem versions \<open>foo__\<close> and \<open>foo___\<close>, respectively.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   610
20488
wenzelm
parents: 20479
diff changeset
   611
  These special versions provide copies of the basic name space, apart
wenzelm
parents: 20479
diff changeset
   612
  from anything that normally appears in the user text.  For example,
wenzelm
parents: 20479
diff changeset
   613
  system generated variables in Isar proof contexts are usually marked
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   614
  as internal, which prevents mysterious names like \<open>xaa\<close> to
34926
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   615
  appear in human-readable text.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   616
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   617
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   618
  Manipulating binding scopes often requires on-the-fly
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   619
  renamings.  A \<^emph>\<open>name context\<close> contains a collection of already
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   620
  used names.  The \<open>declare\<close> operation adds names to the
20488
wenzelm
parents: 20479
diff changeset
   621
  context.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   622
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   623
  The \<open>invents\<close> operation derives a number of fresh names from
20488
wenzelm
parents: 20479
diff changeset
   624
  a given starting point.  For example, the first three names derived
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   625
  from \<open>a\<close> are \<open>a\<close>, \<open>b\<close>, \<open>c\<close>.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   626
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   627
  The \<open>variants\<close> operation produces fresh names by
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   628
  incrementing tentative names as base-26 numbers (with digits \<open>a..z\<close>) until all clashes are resolved.  For example, name \<open>foo\<close> results in variants \<open>fooa\<close>, \<open>foob\<close>, \<open>fooc\<close>, \dots, \<open>fooaa\<close>, \<open>fooab\<close> etc.; each renaming
20488
wenzelm
parents: 20479
diff changeset
   629
  step picks the next unused variant from this sequence.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   630
\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   631
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   632
text %mlref \<open>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   633
  \begin{mldecls}
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   634
  @{index_ML Name.internal: "string -> string"} \\
20547
wenzelm
parents: 20530
diff changeset
   635
  @{index_ML Name.skolem: "string -> string"} \\
wenzelm
parents: 20530
diff changeset
   636
  \end{mldecls}
wenzelm
parents: 20530
diff changeset
   637
  \begin{mldecls}
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   638
  @{index_ML_type Name.context} \\
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   639
  @{index_ML Name.context: Name.context} \\
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   640
  @{index_ML Name.declare: "string -> Name.context -> Name.context"} \\
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   641
  @{index_ML Name.invent: "Name.context -> string -> int -> string list"} \\
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42616
diff changeset
   642
  @{index_ML Name.variant: "string -> Name.context -> string * Name.context"} \\
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   643
  \end{mldecls}
34926
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   644
  \begin{mldecls}
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   645
  @{index_ML Variable.names_of: "Proof.context -> Name.context"} \\
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   646
  \end{mldecls}
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   647
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   648
  \<^descr> @{ML Name.internal}~\<open>name\<close> produces an internal name
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   649
  by adding one underscore.
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   650
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   651
  \<^descr> @{ML Name.skolem}~\<open>name\<close> produces a Skolem name by
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   652
  adding two underscores.
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   653
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   654
  \<^descr> Type @{ML_type Name.context} represents the context of already
39864
wenzelm
parents: 39863
diff changeset
   655
  used names; the initial value is @{ML "Name.context"}.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   656
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   657
  \<^descr> @{ML Name.declare}~\<open>name\<close> enters a used name into the
20488
wenzelm
parents: 20479
diff changeset
   658
  context.
20437
wenzelm
parents: 20430
diff changeset
   659
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   660
  \<^descr> @{ML Name.invent}~\<open>context name n\<close> produces \<open>n\<close> fresh names derived from \<open>name\<close>.
20488
wenzelm
parents: 20479
diff changeset
   661
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   662
  \<^descr> @{ML Name.variant}~\<open>name context\<close> produces a fresh
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   663
  variant of \<open>name\<close>; the result is declared to the context.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   664
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   665
  \<^descr> @{ML Variable.names_of}~\<open>ctxt\<close> retrieves the context
34926
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   666
  of declared type and term variable names.  Projecting a proof
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   667
  context down to a primitive name context is occasionally useful when
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   668
  invoking lower-level operations.  Regular management of ``fresh
55837
154855d9a564 clarified names of antiquotations and markup;
wenzelm
parents: 55112
diff changeset
   669
  variables'' is done by suitable operations of structure @{ML_structure
34926
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   670
  Variable}, which is also able to provide an official status of
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   671
  ``locally fixed variable'' within the logical environment (cf.\
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   672
  \secref{sec:variables}).
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   673
\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   674
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   675
text %mlex \<open>The following simple examples demonstrate how to produce
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   676
  fresh names from the initial @{ML Name.context}.\<close>
39857
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   677
59902
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   678
ML_val \<open>
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   679
  val list1 = Name.invent Name.context "a" 5;
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   680
  @{assert} (list1 = ["a", "b", "c", "d", "e"]);
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   681
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   682
  val list2 =
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42616
diff changeset
   683
    #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] Name.context);
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   684
  @{assert} (list2 = ["x", "xa", "a", "aa", "'a", "'aa"]);
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   685
\<close>
39857
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   686
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   687
text \<open>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   688
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   689
  The same works relatively to the formal context as follows.\<close>
39857
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   690
59902
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   691
experiment fixes a b c :: 'a
39857
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   692
begin
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   693
59902
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   694
ML_val \<open>
39857
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   695
  val names = Variable.names_of @{context};
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   696
43329
84472e198515 tuned signature: Name.invent and Name.invent_names;
wenzelm
parents: 43326
diff changeset
   697
  val list1 = Name.invent names "a" 5;
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   698
  @{assert} (list1 = ["d", "e", "f", "g", "h"]);
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   699
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   700
  val list2 =
43326
47cf4bc789aa simplified Name.variant -- discontinued builtin fold_map;
wenzelm
parents: 42616
diff changeset
   701
    #1 (fold_map Name.variant ["x", "x", "a", "a", "'a", "'a"] names);
39866
5ec01d5acd0c more robust examples: explicit @{assert} instead of unchecked output;
wenzelm
parents: 39865
diff changeset
   702
  @{assert} (list2 = ["x", "xa", "aa", "ab", "'aa", "'ab"]);
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   703
\<close>
39857
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   704
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   705
end
ea93e088398d more examples;
wenzelm
parents: 39846
diff changeset
   706
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   707
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   708
subsection \<open>Indexed names \label{sec:indexname}\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   709
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   710
text \<open>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   711
  An \<^emph>\<open>indexed name\<close> (or \<open>indexname\<close>) is a pair of a basic
20488
wenzelm
parents: 20479
diff changeset
   712
  name and a natural number.  This representation allows efficient
wenzelm
parents: 20479
diff changeset
   713
  renaming by incrementing the second component only.  The canonical
wenzelm
parents: 20479
diff changeset
   714
  way to rename two collections of indexnames apart from each other is
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   715
  this: determine the maximum index \<open>maxidx\<close> of the first
20488
wenzelm
parents: 20479
diff changeset
   716
  collection, then increment all indexes of the second collection by
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   717
  \<open>maxidx + 1\<close>; the maximum index of an empty collection is
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   718
  \<open>-1\<close>.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   719
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   720
  Occasionally, basic names are injected into the same pair type of
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   721
  indexed names: then \<open>(x, -1)\<close> is used to encode the basic
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   722
  name \<open>x\<close>.
20488
wenzelm
parents: 20479
diff changeset
   723
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   724
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   725
  Isabelle syntax observes the following rules for
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   726
  representing an indexname \<open>(x, i)\<close> as a packed string:
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   727
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   728
  \<^item> \<open>?x\<close> if \<open>x\<close> does not end with a digit and \<open>i = 0\<close>,
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   729
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   730
  \<^item> \<open>?xi\<close> if \<open>x\<close> does not end with a digit,
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   731
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   732
  \<^item> \<open>?x.i\<close> otherwise.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   733
20470
c839b38a1f32 more on variables;
wenzelm
parents: 20452
diff changeset
   734
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   735
  Indexnames may acquire large index numbers after several maxidx
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   736
  shifts have been applied.  Results are usually normalized towards
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   737
  \<open>0\<close> at certain checkpoints, notably at the end of a proof.
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   738
  This works by producing variants of the corresponding basic name
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   739
  components.  For example, the collection \<open>?x1, ?x7, ?x42\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   740
  becomes \<open>?x, ?xa, ?xb\<close>.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   741
\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   742
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   743
text %mlref \<open>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   744
  \begin{mldecls}
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   745
  @{index_ML_type indexname: "string * int"} \\
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   746
  \end{mldecls}
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   747
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   748
  \<^descr> Type @{ML_type indexname} represents indexed names.  This is
39864
wenzelm
parents: 39863
diff changeset
   749
  an abbreviation for @{ML_type "string * int"}.  The second component
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   750
  is usually non-negative, except for situations where \<open>(x,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   751
  -1)\<close> is used to inject basic names into this type.  Other negative
34926
19294b07e445 Variable.names_of;
wenzelm
parents: 34925
diff changeset
   752
  indexes should not be used.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   753
\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   754
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   755
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   756
subsection \<open>Long names \label{sec:long-name}\<close>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   757
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   758
text \<open>A \<^emph>\<open>long name\<close> consists of a sequence of non-empty name
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   759
  components.  The packed representation uses a dot as separator, as
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   760
  in ``\<open>A.b.c\<close>''.  The last component is called \<^emph>\<open>base
61477
e467ae7aa808 more control symbols;
wenzelm
parents: 61458
diff changeset
   761
  name\<close>, the remaining prefix is called \<^emph>\<open>qualifier\<close> (which may be
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   762
  empty).  The qualifier can be understood as the access path to the
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   763
  named entity while passing through some nested block-structure,
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   764
  although our free-form long names do not really enforce any strict
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   765
  discipline.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   766
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   767
  For example, an item named ``\<open>A.b.c\<close>'' may be understood as
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   768
  a local entity \<open>c\<close>, within a local structure \<open>b\<close>,
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   769
  within a global structure \<open>A\<close>.  In practice, long names
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   770
  usually represent 1--3 levels of qualification.  User ML code should
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   771
  not make any assumptions about the particular structure of long
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   772
  names!
20437
wenzelm
parents: 20430
diff changeset
   773
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   774
  The empty name is commonly used as an indication of unnamed
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   775
  entities, or entities that are not entered into the corresponding
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   776
  name space, whenever this makes any sense.  The basic operations on
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   777
  long names map empty names again to empty names.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   778
\<close>
20437
wenzelm
parents: 20430
diff changeset
   779
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   780
text %mlref \<open>
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   781
  \begin{mldecls}
30365
790129514c2d adapted to structure Long_Name;
wenzelm
parents: 30281
diff changeset
   782
  @{index_ML Long_Name.base_name: "string -> string"} \\
790129514c2d adapted to structure Long_Name;
wenzelm
parents: 30281
diff changeset
   783
  @{index_ML Long_Name.qualifier: "string -> string"} \\
790129514c2d adapted to structure Long_Name;
wenzelm
parents: 30281
diff changeset
   784
  @{index_ML Long_Name.append: "string -> string -> string"} \\
790129514c2d adapted to structure Long_Name;
wenzelm
parents: 30281
diff changeset
   785
  @{index_ML Long_Name.implode: "string list -> string"} \\
790129514c2d adapted to structure Long_Name;
wenzelm
parents: 30281
diff changeset
   786
  @{index_ML Long_Name.explode: "string -> string list"} \\
20547
wenzelm
parents: 20530
diff changeset
   787
  \end{mldecls}
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   788
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   789
  \<^descr> @{ML Long_Name.base_name}~\<open>name\<close> returns the base name
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   790
  of a long name.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   791
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   792
  \<^descr> @{ML Long_Name.qualifier}~\<open>name\<close> returns the qualifier
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   793
  of a long name.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   794
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   795
  \<^descr> @{ML Long_Name.append}~\<open>name\<^sub>1 name\<^sub>2\<close> appends two long
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   796
  names.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   797
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   798
  \<^descr> @{ML Long_Name.implode}~\<open>names\<close> and @{ML
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   799
  Long_Name.explode}~\<open>name\<close> convert between the packed string
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   800
  representation and the explicit list form of long names.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   801
\<close>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   802
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   803
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   804
subsection \<open>Name spaces \label{sec:name-space}\<close>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   805
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   806
text \<open>A \<open>name space\<close> manages a collection of long names,
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   807
  together with a mapping between partially qualified external names
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   808
  and fully qualified internal names (in both directions).  Note that
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   809
  the corresponding \<open>intern\<close> and \<open>extern\<close> operations
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   810
  are mostly used for parsing and printing only!  The \<open>declare\<close> operation augments a name space according to the accesses
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   811
  determined by a given binding, and a naming policy from the context.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   812
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   813
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   814
  A \<open>binding\<close> specifies details about the prospective
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   815
  long name of a newly introduced formal entity.  It consists of a
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   816
  base name, prefixes for qualification (separate ones for system
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   817
  infrastructure and user-space mechanisms), a slot for the original
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   818
  source position, and some additional flags.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   819
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   820
  \<^medskip>
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   821
  A \<open>naming\<close> provides some additional details for
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   822
  producing a long name from a binding.  Normally, the naming is
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   823
  implicit in the theory or proof context.  The \<open>full\<close>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   824
  operation (and its variants for different context types) produces a
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   825
  fully qualified internal name to be entered into a name space.  The
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   826
  main equation of this ``chemical reaction'' when binding new
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   827
  entities in a context is as follows:
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   828
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   829
  \<^medskip>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   830
  \begin{tabular}{l}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   831
  \<open>binding + naming \<longrightarrow> long name + name space accesses\<close>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   832
  \end{tabular}
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   833
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   834
  \<^bigskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   835
  As a general principle, there is a separate name space for
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   836
  each kind of formal entity, e.g.\ fact, logical constant, type
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   837
  constructor, type class.  It is usually clear from the occurrence in
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   838
  concrete syntax (or from the scope) which kind of entity a name
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   839
  refers to.  For example, the very same name \<open>c\<close> may be used
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   840
  uniformly for a constant, type constructor, and type class.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   841
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   842
  There are common schemes to name derived entities systematically
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   843
  according to the name of the main logical entity involved, e.g.\
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   844
  fact \<open>c.intro\<close> for a canonical introduction rule related to
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   845
  constant \<open>c\<close>.  This technique of mapping names from one
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   846
  space into another requires some care in order to avoid conflicts.
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   847
  In particular, theorem names derived from a type constructor or type
39839
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   848
  class should get an additional suffix in addition to the usual
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   849
  qualification.  This leads to the following conventions for derived
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   850
  names:
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   851
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   852
  \<^medskip>
39839
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   853
  \begin{tabular}{ll}
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   854
  logical entity & fact name \\\hline
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   855
  constant \<open>c\<close> & \<open>c.intro\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   856
  type \<open>c\<close> & \<open>c_type.intro\<close> \\
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   857
  class \<open>c\<close> & \<open>c_class.intro\<close> \\
39839
08f59175e541 misc tuning;
wenzelm
parents: 39837
diff changeset
   858
  \end{tabular}
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   859
\<close>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   860
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   861
text %mlref \<open>
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   862
  \begin{mldecls}
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   863
  @{index_ML_type binding} \\
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   864
  @{index_ML Binding.empty: binding} \\
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   865
  @{index_ML Binding.name: "string -> binding"} \\
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   866
  @{index_ML Binding.qualify: "bool -> string -> binding -> binding"} \\
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   867
  @{index_ML Binding.prefix: "bool -> string -> binding -> binding"} \\
59859
f9d1442c70f3 tuned signature;
wenzelm
parents: 58742
diff changeset
   868
  @{index_ML Binding.concealed: "binding -> binding"} \\
43547
f3a8476285c6 clarified Binding.pretty/print: no quotes, only markup -- Binding.str_of is rendered obsolete;
wenzelm
parents: 43329
diff changeset
   869
  @{index_ML Binding.print: "binding -> string"} \\
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   870
  \end{mldecls}
20547
wenzelm
parents: 20530
diff changeset
   871
  \begin{mldecls}
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   872
  @{index_ML_type Name_Space.naming} \\
58668
1891f17c6124 tuned signature;
wenzelm
parents: 58618
diff changeset
   873
  @{index_ML Name_Space.global_naming: Name_Space.naming} \\
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   874
  @{index_ML Name_Space.add_path: "string -> Name_Space.naming -> Name_Space.naming"} \\
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   875
  @{index_ML Name_Space.full_name: "Name_Space.naming -> binding -> string"} \\
20547
wenzelm
parents: 20530
diff changeset
   876
  \end{mldecls}
wenzelm
parents: 20530
diff changeset
   877
  \begin{mldecls}
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   878
  @{index_ML_type Name_Space.T} \\
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   879
  @{index_ML Name_Space.empty: "string -> Name_Space.T"} \\
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   880
  @{index_ML Name_Space.merge: "Name_Space.T * Name_Space.T -> Name_Space.T"} \\
47174
b9b2e183e94d updated Sign.add_type, Name_Space.declare;
wenzelm
parents: 46484
diff changeset
   881
  @{index_ML Name_Space.declare: "Context.generic -> bool ->
b9b2e183e94d updated Sign.add_type, Name_Space.declare;
wenzelm
parents: 46484
diff changeset
   882
  binding -> Name_Space.T -> string * Name_Space.T"} \\
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   883
  @{index_ML Name_Space.intern: "Name_Space.T -> string -> string"} \\
42358
b47d41d9f4b5 Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
wenzelm
parents: 40964
diff changeset
   884
  @{index_ML Name_Space.extern: "Proof.context -> Name_Space.T -> string -> string"} \\
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   885
  @{index_ML Name_Space.is_concealed: "Name_Space.T -> string -> bool"}
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   886
  \end{mldecls}
20437
wenzelm
parents: 20430
diff changeset
   887
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   888
  \<^descr> Type @{ML_type binding} represents the abstract concept of
39864
wenzelm
parents: 39863
diff changeset
   889
  name bindings.
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   890
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   891
  \<^descr> @{ML Binding.empty} is the empty binding.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   892
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   893
  \<^descr> @{ML Binding.name}~\<open>name\<close> produces a binding with base
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   894
  name \<open>name\<close>.  Note that this lacks proper source position
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   895
  information; see also the ML antiquotation @{ML_antiquotation
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   896
  binding}.
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   897
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   898
  \<^descr> @{ML Binding.qualify}~\<open>mandatory name binding\<close>
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   899
  prefixes qualifier \<open>name\<close> to \<open>binding\<close>.  The \<open>mandatory\<close> flag tells if this name component always needs to be
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   900
  given in name space accesses --- this is mostly \<open>false\<close> in
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   901
  practice.  Note that this part of qualification is typically used in
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   902
  derived specification mechanisms.
20437
wenzelm
parents: 20430
diff changeset
   903
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   904
  \<^descr> @{ML Binding.prefix} is similar to @{ML Binding.qualify}, but
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   905
  affects the system prefix.  This part of extra qualification is
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   906
  typically used in the infrastructure for modular specifications,
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   907
  notably ``local theory targets'' (see also \chref{ch:local-theory}).
20437
wenzelm
parents: 20430
diff changeset
   908
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   909
  \<^descr> @{ML Binding.concealed}~\<open>binding\<close> indicates that the
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   910
  binding shall refer to an entity that serves foundational purposes
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   911
  only.  This flag helps to mark implementation details of
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   912
  specification mechanism etc.  Other tools should not depend on the
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   913
  particulars of concealed entities (cf.\ @{ML
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   914
  Name_Space.is_concealed}).
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   915
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   916
  \<^descr> @{ML Binding.print}~\<open>binding\<close> produces a string
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   917
  representation for human-readable output, together with some formal
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   918
  markup that might get used in GUI front-ends, for example.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   919
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   920
  \<^descr> Type @{ML_type Name_Space.naming} represents the abstract
39864
wenzelm
parents: 39863
diff changeset
   921
  concept of a naming policy.
20437
wenzelm
parents: 20430
diff changeset
   922
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   923
  \<^descr> @{ML Name_Space.global_naming} is the default naming policy: it is
58668
1891f17c6124 tuned signature;
wenzelm
parents: 58618
diff changeset
   924
  global and lacks any path prefix.  In a regular theory context this is
1891f17c6124 tuned signature;
wenzelm
parents: 58618
diff changeset
   925
  augmented by a path prefix consisting of the theory name.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   926
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   927
  \<^descr> @{ML Name_Space.add_path}~\<open>path naming\<close> augments the
20488
wenzelm
parents: 20479
diff changeset
   928
  naming policy by extending its path component.
20437
wenzelm
parents: 20430
diff changeset
   929
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   930
  \<^descr> @{ML Name_Space.full_name}~\<open>naming binding\<close> turns a
30281
9ad15d8ed311 renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
wenzelm
parents: 30272
diff changeset
   931
  name binding (usually a basic name) into the fully qualified
29008
d863c103ded0 adapted to changes in binding module
haftmann
parents: 26872
diff changeset
   932
  internal name, according to the given naming policy.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   933
61439
2bf52eec4e8a more symbols;
wenzelm
parents: 61416
diff changeset
   934
  \<^descr> Type @{ML_type Name_Space.T} represents name spaces.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   935
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   936
  \<^descr> @{ML Name_Space.empty}~\<open>kind\<close> and @{ML Name_Space.merge}~\<open>(space\<^sub>1, space\<^sub>2)\<close> are the canonical operations for
20488
wenzelm
parents: 20479
diff changeset
   937
  maintaining name spaces according to theory data management
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   938
  (\secref{sec:context-data}); \<open>kind\<close> is a formal comment
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   939
  to characterize the purpose of a name space.
20437
wenzelm
parents: 20430
diff changeset
   940
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   941
  \<^descr> @{ML Name_Space.declare}~\<open>context strict binding
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   942
  space\<close> enters a name binding as fully qualified internal name into
47174
b9b2e183e94d updated Sign.add_type, Name_Space.declare;
wenzelm
parents: 46484
diff changeset
   943
  the name space, using the naming of the context.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   944
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   945
  \<^descr> @{ML Name_Space.intern}~\<open>space name\<close> internalizes a
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   946
  (partially qualified) external name.
20437
wenzelm
parents: 20430
diff changeset
   947
20488
wenzelm
parents: 20479
diff changeset
   948
  This operation is mostly for parsing!  Note that fully qualified
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   949
  names stemming from declarations are produced via @{ML
33174
1f2051f41335 adjusted to changes in corresponding ML code
haftmann
parents: 30365
diff changeset
   950
  "Name_Space.full_name"} and @{ML "Name_Space.declare"}
29008
d863c103ded0 adapted to changes in binding module
haftmann
parents: 26872
diff changeset
   951
  (or their derivatives for @{ML_type theory} and
20488
wenzelm
parents: 20479
diff changeset
   952
  @{ML_type Proof.context}).
20437
wenzelm
parents: 20430
diff changeset
   953
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   954
  \<^descr> @{ML Name_Space.extern}~\<open>ctxt space name\<close> externalizes a
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   955
  (fully qualified) internal name.
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   956
30281
9ad15d8ed311 renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
wenzelm
parents: 30272
diff changeset
   957
  This operation is mostly for printing!  User code should not rely on
9ad15d8ed311 renamed NameSpace.base to NameSpace.base_name (in accordance with "full_name");
wenzelm
parents: 30272
diff changeset
   958
  the precise result too much.
20476
6d3f144cc1bd more on names;
wenzelm
parents: 20475
diff changeset
   959
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   960
  \<^descr> @{ML Name_Space.is_concealed}~\<open>space name\<close> indicates
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   961
  whether \<open>name\<close> refers to a strictly private entity that
34927
c4c02ac736a6 more details on long names, binding/naming, name space;
wenzelm
parents: 34926
diff changeset
   962
  other tools are supposed to ignore!
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   963
\<close>
30272
2d612824e642 regenerated document;
wenzelm
parents: 30270
diff changeset
   964
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   965
text %mlantiq \<open>
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   966
  \begin{matharray}{rcl}
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   967
  @{ML_antiquotation_def "binding"} & : & \<open>ML_antiquotation\<close> \\
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   968
  \end{matharray}
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   969
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 53015
diff changeset
   970
  @{rail \<open>
42510
b9c106763325 use @{rail} antiquotation (with some nested markup);
wenzelm
parents: 42401
diff changeset
   971
  @@{ML_antiquotation binding} name
55112
b1a5d603fd12 prefer rail cartouche -- avoid back-slashed quotes;
wenzelm
parents: 53015
diff changeset
   972
  \<close>}
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   973
61493
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   974
  \<^descr> \<open>@{binding name}\<close> produces a binding with base name
0debd22f0c0e isabelle update_cartouches -t;
wenzelm
parents: 61477
diff changeset
   975
  \<open>name\<close> and the source position taken from the concrete
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   976
  syntax of this antiquotation.  In many situations this is more
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   977
  appropriate than the more basic @{ML Binding.name} function.
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   978
\<close>
39832
1080dee73a53 various concrete ML antiquotations;
wenzelm
parents: 39825
diff changeset
   979
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   980
text %mlex \<open>The following example yields the source position of some
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 39876
diff changeset
   981
  concrete binding inlined into the text:
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   982
\<close>
39833
6d54a48c859d more examples;
wenzelm
parents: 39832
diff changeset
   983
59902
6afbe5a99139 misc tuning -- keep name space more clean;
wenzelm
parents: 59859
diff changeset
   984
ML_val \<open>Binding.pos_of @{binding here}\<close>
39833
6d54a48c859d more examples;
wenzelm
parents: 39832
diff changeset
   985
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   986
text \<open>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   987
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   988
  That position can be also printed in a message as follows:\<close>
39833
6d54a48c859d more examples;
wenzelm
parents: 39832
diff changeset
   989
58742
bb55a3530709 tuned spacing;
wenzelm
parents: 58723
diff changeset
   990
ML_command
bb55a3530709 tuned spacing;
wenzelm
parents: 58723
diff changeset
   991
  \<open>writeln
bb55a3530709 tuned spacing;
wenzelm
parents: 58723
diff changeset
   992
    ("Look here" ^ Position.here (Binding.pos_of @{binding here}))\<close>
39833
6d54a48c859d more examples;
wenzelm
parents: 39832
diff changeset
   993
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
   994
text \<open>This illustrates a key virtue of formalized bindings as
39861
b8d89db3e238 use continental paragraph style, which works better with mixture of (in)formal text;
wenzelm
parents: 39857
diff changeset
   995
  opposed to raw specifications of base names: the system can use this
40126
916cb4a28ffd misc tuning;
wenzelm
parents: 39876
diff changeset
   996
  additional information for feedback given to the user (error
56071
2ffdedb0c044 added ML antiquotation @{here};
wenzelm
parents: 55837
diff changeset
   997
  messages etc.).
2ffdedb0c044 added ML antiquotation @{here};
wenzelm
parents: 55837
diff changeset
   998
61416
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
   999
  \<^medskip>
b9a3324e4e62 more symbols;
wenzelm
parents: 61046
diff changeset
  1000
  The following example refers to its source position
56071
2ffdedb0c044 added ML antiquotation @{here};
wenzelm
parents: 55837
diff changeset
  1001
  directly, which is occasionally useful for experimentation and
58618
782f0b662cae more cartouches;
wenzelm
parents: 58555
diff changeset
  1002
  diagnostic purposes:\<close>
56071
2ffdedb0c044 added ML antiquotation @{here};
wenzelm
parents: 55837
diff changeset
  1003
58742
bb55a3530709 tuned spacing;
wenzelm
parents: 58723
diff changeset
  1004
ML_command \<open>warning ("Look here" ^ Position.here @{here})\<close>
39833
6d54a48c859d more examples;
wenzelm
parents: 39832
diff changeset
  1005
18537
2681f9e34390 "The Isabelle/Isar Implementation" manual;
wenzelm
parents:
diff changeset
  1006
end