ProgTutorial/Advanced.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 23 May 2019 00:58:11 +0100
changeset 577 d1523393dd5a
parent 572 438703674711
permissions -rw-r--r--
updated pdf-document
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
     1
theory Advanced
441
Christian Urban <urbanc@in.tum.de>
parents: 414
diff changeset
     2
imports Base First_Steps
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
     5
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
     6
chapter \<open>Advanced Isabelle\label{chp:advanced}\<close>
381
97518188ef0e added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents: 380
diff changeset
     7
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
     8
text \<open>
410
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
     9
   \begin{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    10
  {\em All things are difficult before they are easy.} \\[1ex]
539
12861a362099 updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 518
diff changeset
    11
  proverb\\[2ex]
12861a362099 updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 518
diff changeset
    12
  {\em Programming is not just an act of telling a computer what 
12861a362099 updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 518
diff changeset
    13
  to do: it is also an act of telling other programmers what you 
12861a362099 updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 518
diff changeset
    14
  wished the computer to do. Both are important, and the latter 
12861a362099 updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 518
diff changeset
    15
  deserves care.}\\[1ex]
12861a362099 updated to new isabelle
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 518
diff changeset
    16
  ---Andrew Morton, Linux Kernel mailinglist, 13 March 2012
410
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    17
  \end{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    18
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    19
  \medskip
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    20
  While terms, types and theorems are the most basic data structures in
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    21
  Isabelle, there are a number of layers built on top of them. Most of these
408
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    22
  layers are concerned with storing and manipulating data. Handling them
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    23
  properly is an essential skill for programming on the ML-level of Isabelle. 
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    24
  The most basic layer are theories. They contain global data and
408
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    25
  can be seen as the ``long-term memory'' of Isabelle. There is usually only
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    26
  one theory active at each moment. Proof contexts and local theories, on the
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    27
  other hand, store local data for a task at hand. They act like the
408
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    28
  ``short-term memory'' and there can be many of them that are active in
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    29
  parallel.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    30
\<close>
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    32
section \<open>Theories\label{sec:theories}\<close>
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    33
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    34
text \<open>
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
    35
  Theories, as said above, are the most basic layer of abstraction in
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
    36
  Isabelle. They record information about definitions, syntax declarations, axioms,
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    37
  theorems and much more.  For example, if a definition is made, it
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    38
  must be stored in a theory in order to be usable later on. Similar
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    39
  with proofs: once a proof is finished, the proved theorem needs to
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    40
  be stored in the theorem database of the theory in order to be
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
    41
  usable. All relevant data of a theory can be queried with the
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
    42
  Isabelle command \isacommand{print\_theory}.
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    43
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    44
  \begin{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    45
  \isacommand{print\_theory}\\
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    46
  \<open>> names: Pure Code_Generator HOL \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    47
  \<open>> classes: Inf < type \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    48
  \<open>> default sort: type\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    49
  \<open>> syntactic types: #prop \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    50
  \<open>> logical types: 'a \<times> 'b \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    51
  \<open>> type arities: * :: (random, random) random \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    52
  \<open>> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    53
  \<open>> abbreviations: \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    54
  \<open>> axioms: \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    55
  \<open>> oracles: \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    56
  \<open>> definitions: \<dots>\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    57
  \<open>> theorems: \<dots>\<close>
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    58
  \end{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    59
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    60
  Functions acting on theories often end with the suffix \<open>_global\<close>,
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    61
  for example the function @{ML read_term_global in Syntax} in the structure
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    62
  @{ML_structure Syntax}. The reason is to set them syntactically apart from
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
    63
  functions acting on contexts or local theories, which will be discussed in
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
    64
  the next sections. There is a tendency amongst Isabelle developers to prefer
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    65
  ``non-global'' operations, because they have some advantages, as we will also
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    66
  discuss later. However, some basic understanding of theories is still necessary
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    67
  for effective Isabelle programming.
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    68
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
    69
  An important Isabelle command with theories is \isacommand{setup}. In the
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
    70
  previous chapters we used it already to make a theorem attribute known
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
    71
  to Isabelle and to register a theorem under a name. What happens behind the
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    72
  scenes is that \isacommand{setup} expects a function of type @{ML_type
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    73
  "theory -> theory"}: the input theory is the current theory and the output
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    74
  the theory where the attribute has been registered or the theorem has been
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    75
  stored.  This is a fundamental principle in Isabelle. A similar situation
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
    76
  arises with declaring a constant, which can be done on the ML-level with 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
    77
  function @{ML_ind declare_const in Sign} from the structure @{ML_structure
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    78
  Sign}. To see how \isacommand{setup} works, consider the following code:
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
    79
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    80
\<close>  
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    81
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    82
ML %grayML\<open>let
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    83
  val thy = @{theory}
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    84
  val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    85
in 
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    86
  Sign.declare_const @{context} bar_const thy  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    87
end\<close>
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    88
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    89
text \<open>
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    90
  If you simply run this code\footnote{Recall that ML-code needs to be enclosed in
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    91
  \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close>.} with the
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    92
  intention of declaring a constant \<open>BAR\<close> having type @{typ nat}, then 
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    93
  indeed you obtain a theory as result. But if you query the
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    94
  constant on the Isabelle level using the command \isacommand{term}
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    95
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    96
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    97
  \isacommand{term}~\<open>BAR\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
    98
  \<open>> "BAR" :: "'a"\<close>
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    99
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   100
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   101
  you can see that you do \emph{not} obtain the expected constant of type @{typ nat}, but a free 
484
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 481
diff changeset
   102
  variable (printed in blue) of polymorphic type. The problem is that the 
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 481
diff changeset
   103
  ML-expression above did not ``register'' the declaration with the current theory. 
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 481
diff changeset
   104
  This is what the command \isacommand{setup} is for. The constant is properly 
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 481
diff changeset
   105
  declared with
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   106
\<close>
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   107
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   108
setup %gray \<open>fn thy => 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   109
let
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   110
  val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   111
  val (_, thy') = Sign.declare_const @{context} bar_const thy
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   112
in 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   113
  thy'
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   114
end\<close>
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   115
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   116
text \<open>
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   117
  where the declaration is actually applied to the current theory and
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   118
  
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   119
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   120
  \isacommand{term}~@{text [quotes] "BAR"}\\
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   121
  \<open>> "BAR" :: "nat"\<close>
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   122
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   123
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
   124
  now returns a (black) constant with the type @{typ nat}, as expected.
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   125
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
   126
  In a sense, \isacommand{setup} can be seen as a transaction that
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   127
  takes the current theory \<open>thy\<close>, applies an operation, and
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   128
  produces a new current theory \<open>thy'\<close>. This means that we have
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
   129
  to be careful to apply operations always to the most current theory,
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
   130
  not to a \emph{stale} one. Consider again the function inside the
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   131
  \isacommand{setup}-command:
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   132
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   133
  \begin{isabelle}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   134
  \begin{graybox}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   135
  \isacommand{setup}~\<open>\<verbopen>\<close> @{ML \<open>fn thy => 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   136
let
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   137
  val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   138
  val (_, thy') = Sign.declare_const @{context} bar_const thy
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   139
in
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   140
  thy
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   141
end\<close>}~\<open>\<verbclose>\<close>\isanewline
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   142
  \<open>> ERROR: "Stale theory encountered"\<close>
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   143
  \end{graybox}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   144
  \end{isabelle}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   145
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   146
  This time we erroneously return the original theory \<open>thy\<close>, instead of
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   147
  the modified one \<open>thy'\<close>. Such buggy code will always result into 
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   148
  a runtime error message about stale theories.
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   149
502
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   150
  \begin{readmore}
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   151
  Most of the functions about theories are implemented in
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   152
  @{ML_file "Pure/theory.ML"} and @{ML_file "Pure/global_theory.ML"}.
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   153
  \end{readmore}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   154
\<close>
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   156
section \<open>Contexts\<close>
341
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   157
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   158
text \<open>
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   159
  Contexts are arguably more important than theories, even though they only 
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   160
  contain ``short-term memory data''. The reason is that a vast number of
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   161
  functions in Isabelle depend in one way or another on contexts. Even such
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   162
  mundane operations like printing out a term make essential use of contexts.
494
Christian Urban <urbanc@in.tum.de>
parents: 493
diff changeset
   163
  For this consider the following contrived proof-snippet whose purpose is to 
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   164
  fix two variables:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   165
\<close>
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   166
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   167
lemma "True"
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   168
proof -
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   169
  ML_prf \<open>Variable.dest_fixes @{context}\<close> 
518
7ff1a681f758 updated
Christian Urban <urbanc@in.tum.de>
parents: 517
diff changeset
   170
  fix x y  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   171
  ML_prf \<open>Variable.dest_fixes @{context}\<close>(*<*)oops(*>*)
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   172
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   173
text \<open>
496
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   174
  The interesting point is that we injected ML-code before and after
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   175
  the variables are fixed. For this remember that ML-code inside a proof
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   176
  needs to be enclosed inside \isacommand{ML\_prf}~\<open>\<verbopen> \<dots> \<verbclose>\<close>,
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   177
  not \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close>. The function 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   178
  @{ML_ind dest_fixes in Variable} from the structure @{ML_structure Variable} takes 
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   179
  a context and returns all its currently fixed variable (names). That 
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   180
  means a context has a dataslot containing information about fixed variables.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   181
  The ML-antiquotation \<open>@{context}\<close> points to the context that is
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   182
  active at that point of the theory. Consequently, in the first call to 
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   183
  @{ML dest_fixes in Variable} this dataslot is  empty; in the second it is 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   184
  filled with \<open>x\<close> and \<open>y\<close>. What is interesting is that contexts
494
Christian Urban <urbanc@in.tum.de>
parents: 493
diff changeset
   185
  can be stacked. For this consider the following proof fragment:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   186
\<close>
490
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   187
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   188
lemma "True"
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   189
proof -
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   190
  fix x y
Christian Urban <urbanc@in.tum.de>
parents: 488
diff changeset
   191
  { fix z w
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   192
    ML_prf \<open>Variable.dest_fixes @{context}\<close> 
518
7ff1a681f758 updated
Christian Urban <urbanc@in.tum.de>
parents: 517
diff changeset
   193
  }
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   194
  ML_prf \<open>Variable.dest_fixes @{context}\<close>(*<*)oops(*>*)
491
Christian Urban <urbanc@in.tum.de>
parents: 490
diff changeset
   195
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   196
text \<open>
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   197
  Here the first time we call @{ML dest_fixes in Variable} we have four fixes variables;
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   198
  the second time we get only the fixes variables \<open>x\<close> and \<open>y\<close> as answer, since
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   199
  \<open>z\<close> and \<open>w\<close> are not anymore in the scope of the context. 
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   200
  This means the curly-braces act on the Isabelle level as opening and closing statements 
496
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   201
  for a context. The above proof fragment corresponds roughly to the following ML-code
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   202
\<close>
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   203
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   204
ML %grayML\<open>val ctxt0 = @{context};
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
   205
val ([x, y], ctxt1) = Variable.add_fixes ["x", "y"] ctxt0;
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   206
val ([z, w], ctxt2) = Variable.add_fixes ["z", "w"] ctxt1\<close>
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
   207
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   208
text \<open>
494
Christian Urban <urbanc@in.tum.de>
parents: 493
diff changeset
   209
  where the function @{ML_ind add_fixes in Variable} fixes a list of variables
Christian Urban <urbanc@in.tum.de>
parents: 493
diff changeset
   210
  specified by strings. Let us come back to the point about printing terms. Consider
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   211
  printing out the term \mbox{\<open>(x, y, z, w)\<close>} using our function @{ML_ind pretty_term}.
496
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   212
  This function takes a term and a context as argument. Notice how the printing
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   213
  of the term changes according to which context is used.
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
   214
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   215
  \begin{isabelle}
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   216
  \begin{graybox}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   217
  @{ML \<open>let
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   218
  val trm = @{term "(x, y, z, w)"}
493
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   219
in
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   220
  pwriteln (Pretty.chunks 
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   221
    [ pretty_term ctxt0 trm,
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   222
      pretty_term ctxt1 trm,
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   223
      pretty_term ctxt2 trm ])
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   224
end\<close>}\\
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   225
  \setlength{\fboxsep}{0mm}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   226
  \<open>>\<close>~\<open>(\<close>\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   227
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   228
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   229
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   230
  \<open>>\<close>~\<open>(\<close>\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   231
  \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   232
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   233
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   234
  \<open>>\<close>~\<open>(\<close>\colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   235
  \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   236
  \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}\<open>,\<close>~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   237
  \colorbox{gray!20}{\raisebox{0mm}[3mm][1mm]{\<open>w\<close>}}\<open>)\<close>
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   238
  \end{graybox}
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   239
  \end{isabelle}
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   240
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   241
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   242
  The term we are printing is in all three cases the same---a tuple containing
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   243
  four free variables. As you can see, however, in case of @{ML \<open>ctxt0\<close>} all
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   244
  variables are highlighted indicating a problem, while in case of @{ML \<open>ctxt1\<close>} \<open>x\<close> and \<open>y\<close> are printed as normal (blue) free
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   245
  variables, but not \<open>z\<close> and \<open>w\<close>. In the last case all variables
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   246
  are printed as expected. The point of this example is that the context
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   247
  contains the information which variables are fixed, and designates all other
497
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   248
  free variables as being alien or faulty.  Therefore the highlighting. 
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   249
  While this seems like a minor detail, the concept of making the context aware 
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   250
  of fixed variables is actually quite useful. For example it prevents us from 
501
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   251
  fixing a variable twice
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   252
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   253
  @{ML_response [gray, display]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   254
  \<open>@{context}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   255
|> Variable.add_fixes ["x", "x"]\<close> 
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   256
  \<open>Duplicate fixed variable(s): "x"\<close>}
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   257
501
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   258
  More importantly it also allows us to easily create fresh names for
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   259
  fixed variables.  For this you have to use the function @{ML_ind
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   260
  variant_fixes in Variable} from the structure @{ML_structure Variable}.
501
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   261
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   262
  @{ML_response [gray, display]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   263
  \<open>@{context}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   264
|> Variable.variant_fixes ["y", "y", "z"]\<close> 
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   265
  \<open>(["y", "ya", "z"],\<dots>\<close>}
501
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   266
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   267
  Now a fresh variant for the second occurence of \<open>y\<close> is created
502
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   268
  avoiding any clash. In this way we can also create fresh free variables
501
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   269
  that avoid any clashes with fixed variables. In Line~3 below we fix
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   270
  the variable \<open>x\<close> in the context \<open>ctxt1\<close>. Next we want to
501
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   271
  create two fresh variables of type @{typ nat} as variants of the
f56fc3305a08 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 500
diff changeset
   272
  string @{text [quotes] "x"} (Lines 6 and 7).
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   273
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   274
  @{ML_matchresult [display, gray, linenos]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   275
  \<open>let
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   276
  val ctxt0 = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   277
  val (_, ctxt1) = Variable.add_fixes ["x"] ctxt0
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   278
  val frees = replicate 2 ("x", @{typ nat})
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   279
in
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   280
  (Variable.variant_frees ctxt0 [] frees,
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   281
   Variable.variant_frees ctxt1 [] frees)
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   282
end\<close>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   283
  \<open>([("x", _), ("xa", _)], 
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   284
 [("xa", _), ("xb", _)])\<close>}
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   285
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   286
  As you can see, if we create the fresh variables with the context \<open>ctxt0\<close>,
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   287
  then we obtain \<open>x\<close> and \<open>xa\<close>; but in \<open>ctxt1\<close> we obtain \<open>xa\<close>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   288
  and \<open>xb\<close> avoiding \<open>x\<close>, which was fixed in this context.
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   289
496
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   290
  Often one has the problem that given some terms, create fresh variables
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   291
  avoiding any variable occurring in those terms. For this you can use the
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   292
  function @{ML_ind declare_term in Variable} from the structure @{ML_structure Variable}.
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   293
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   294
  @{ML_matchresult [gray, display]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   295
  \<open>let
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   296
  val ctxt1 = Variable.declare_term @{term "(x, xa)"} @{context}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   297
  val frees = replicate 2 ("x", @{typ nat})
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   298
in
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   299
  Variable.variant_frees ctxt1 [] frees
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   300
end\<close>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   301
  \<open>[("xb", _), ("xc", _)]\<close>}
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   302
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   303
  The result is \<open>xb\<close> and \<open>xc\<close> for the names of the fresh
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   304
  variables, since \<open>x\<close> and \<open>xa\<close> occur in the term we declared. 
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   305
  Note that @{ML_ind declare_term in Variable} does not fix the
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   306
  variables; it just makes them ``known'' to the context. You can see
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   307
  that if you print out a declared term.
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   308
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   309
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   310
  \begin{graybox}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   311
  @{ML \<open>let
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   312
  val trm = @{term "P x y z"}
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   313
  val ctxt1 = Variable.declare_term trm @{context}
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   314
in
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   315
  pwriteln (pretty_term ctxt1 trm)
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   316
end\<close>}\\
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   317
  \setlength{\fboxsep}{0mm}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   318
  \<open>>\<close>~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>P\<close>}}~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   319
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>x\<close>}}~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   320
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>y\<close>}}~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   321
  \colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>z\<close>}}
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   322
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   323
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   324
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   325
  All variables are highligted, indicating that they are not
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   326
  fixed. However, declaring a term is helpful when parsing terms using
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   327
  the function @{ML_ind read_term in Syntax} from the structure
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   328
  @{ML_structure Syntax}. Consider the following code:
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   329
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   330
  @{ML_response [gray, display]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   331
  \<open>let
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   332
  val ctxt0 = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   333
  val ctxt1 = Variable.declare_term @{term "x::nat"} ctxt0
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   334
in
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   335
  (Syntax.read_term ctxt0 "x", 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   336
   Syntax.read_term ctxt1 "x") 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   337
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   338
  \<open>(Free ("x", "'a"), Free ("x", "nat"))\<close>}
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   339
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   340
  Parsing the string in the context \<open>ctxt0\<close> results in a free variable 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   341
  with a default polymorphic type, but in case of \<open>ctxt1\<close> we obtain a
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   342
  free variable of type @{typ nat} as previously declared. Which
496
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   343
  type the parsed term receives depends upon the last declaration that
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   344
  is made, as the next example illustrates.
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   345
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   346
  @{ML_response [gray, display]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   347
  \<open>let
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   348
  val ctxt1 = Variable.declare_term @{term "x::nat"} @{context}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   349
  val ctxt2 = Variable.declare_term @{term "x::int"} ctxt1
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   350
in
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   351
  (Syntax.read_term ctxt1 "x", 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   352
   Syntax.read_term ctxt2 "x") 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   353
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   354
  \<open>(Free ("x", "nat"), Free ("x", "int"))\<close>}
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   355
499
42bac8b16951 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   356
  The most useful feature of contexts is that one can export, or transfer, 
42bac8b16951 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   357
  terms and theorems between them. We show this first for terms. 
497
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   358
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   359
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   360
  \begin{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   361
  \begin{linenos}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   362
  @{ML \<open>let
497
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   363
  val ctxt0 = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   364
  val (_, ctxt1) = Variable.add_fixes ["x", "y", "z"] ctxt0 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   365
  val foo_trm = @{term "P x y z"}
497
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   366
in
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   367
  singleton (Variable.export_terms ctxt1 ctxt0) foo_trm
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   368
  |> pretty_term ctxt0
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   369
  |> pwriteln
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   370
end\<close>}
497
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   371
  \end{linenos}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   372
  \setlength{\fboxsep}{0mm}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   373
  \<open>>\<close>~\colorbox{gray!5}{\raisebox{0mm}[3mm][1mm]{\<open>P\<close>}}~%
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   374
  \<open>?x ?y ?z\<close>
497
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   375
  \end{graybox}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   376
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   377
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   378
  In Line 3 we fix the variables @{term x}, @{term y} and @{term z} in
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   379
  context \<open>ctxt1\<close>.  The function @{ML_ind export_terms in
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   380
  Variable} from the structure @{ML_structure Variable} can be used to transfer
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   381
  terms between contexts. Transferring means to turn all (free)
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   382
  variables that are fixed in one context, but not in the other, into
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   383
  schematic variables. In our example, we are transferring the term
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   384
  \<open>P x y z\<close> from context \<open>ctxt1\<close> to \<open>ctxt0\<close>,
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   385
  which means @{term x}, @{term y} and @{term z} become schematic
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   386
  variables (as can be seen by the leading question marks in the result). 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   387
  Note that the variable \<open>P\<close> stays a free variable, since it not fixed in
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   388
  \<open>ctxt1\<close>; it is even highlighed, because \<open>ctxt0\<close> does
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   389
  not know about it. Note also that in Line 6 we had to use the
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   390
  function @{ML_ind singleton}, because the function @{ML_ind
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   391
  export_terms in Variable} normally works over lists of terms.
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   392
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   393
  The case of transferring theorems is even more useful. The reason is
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   394
  that the generalisation of fixed variables to schematic variables is
499
42bac8b16951 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   395
  not trivial if done manually. For illustration purposes we use in the 
42bac8b16951 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   396
  following code the function @{ML_ind make_thm in Skip_Proof} from the 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   397
  structure @{ML_structure Skip_Proof}. This function will turn an arbitray 
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   398
  term, in our case @{term "P x y z x y z"}, into a theorem (disregarding 
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   399
  whether it is actually provable).
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   400
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   401
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   402
  \<open>let
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   403
  val thy = @{theory}
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   404
  val ctxt0 = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   405
  val (_, ctxt1) = Variable.add_fixes ["P", "x", "y", "z"] ctxt0 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   406
  val foo_thm = Skip_Proof.make_thm thy @{prop "P x y z x y z"}
498
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   407
in
Christian Urban <urbanc@in.tum.de>
parents: 497
diff changeset
   408
  singleton (Proof_Context.export ctxt1 ctxt0) foo_thm
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   409
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   410
  \<open>?P ?x ?y ?z ?x ?y ?z\<close>}
499
42bac8b16951 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 498
diff changeset
   411
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   412
  Since we fixed all variables in \<open>ctxt1\<close>, in the exported
502
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   413
  result all of them are schematic. The great point of contexts is
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   414
  that exporting from one to another is not just restricted to
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   415
  variables, but also works with assumptions. For this we can use the
Christian Urban <urbanc@in.tum.de>
parents: 501
diff changeset
   416
  function @{ML_ind export in Assumption} from the structure
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   417
  @{ML_structure Assumption}. Consider the following code.
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   418
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   419
  @{ML_response [display, gray, linenos]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   420
  \<open>let
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   421
  val ctxt0 = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   422
  val ([eq], ctxt1) = Assumption.add_assumes [@{cprop "x \<equiv> y"}] ctxt0
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   423
  val eq' = Thm.symmetric eq
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   424
in
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   425
  Assumption.export false ctxt1 ctxt0 eq'
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   426
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   427
  \<open>x \<equiv> y \<Longrightarrow> y \<equiv> x\<close>}
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   428
  
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   429
  The function @{ML_ind add_assumes in Assumption} from the structure
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   430
  @{ML_structure Assumption} adds the assumption \mbox{\<open>x \<equiv> y\<close>}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   431
  to the context \<open>ctxt1\<close> (Line 3). This function expects a list
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   432
  of @{ML_type cterm}s and returns them as theorems, together with the
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   433
  new context in which they are ``assumed''. In Line 4 we use the
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   434
  function @{ML_ind symmetric in Thm} from the structure @{ML_structure
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   435
  Thm} in order to obtain the symmetric version of the assumed
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   436
  meta-equality. Now exporting the theorem \<open>eq'\<close> from \<open>ctxt1\<close> to \<open>ctxt0\<close> means @{term "y \<equiv> x"} will be prefixed with
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   437
  the assumed theorem. The boolean flag in @{ML_ind export in
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   438
  Assumption} indicates whether the assumptions should be marked with
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   439
  the goal marker (see Section~\ref{sec:basictactics}). In normal
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   440
  circumstances this is not necessary and so should be set to @{ML
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   441
  false}.  The result of the export is then the theorem \mbox{@{term
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   442
  "x \<equiv> y \<Longrightarrow> y \<equiv> x"}}.  As can be seen this is an easy way for obtaing
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   443
  simple theorems. We will explain this in more detail in
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   444
  Section~\ref{sec:structured}.
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   445
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   446
  The function @{ML_ind export in Proof_Context} from the structure 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   447
  @{ML_structure Proof_Context} combines both export functions from 
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   448
  @{ML_structure Variable} and @{ML_structure Assumption}. This can be seen
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   449
  in the following example.
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   450
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   451
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   452
  \<open>let
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   453
  val ctxt0 = @{context}
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   454
  val ((fvs, [eq]), ctxt1) = ctxt0
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   455
    |> Variable.add_fixes ["x", "y"]
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   456
    ||>> Assumption.add_assumes [@{cprop "x \<equiv> y"}]
500
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   457
  val eq' = Thm.symmetric eq
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   458
in
6cfde4ff13e3 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 499
diff changeset
   459
  Proof_Context.export ctxt1 ctxt0 [eq']
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   460
end\<close>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   461
  \<open>["?x \<equiv> ?y \<Longrightarrow> ?y \<equiv> ?x"]\<close>}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   462
\<close>
495
f3f24874e8be more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 494
diff changeset
   463
496
Christian Urban <urbanc@in.tum.de>
parents: 495
diff changeset
   464
493
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   465
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   466
text \<open>
493
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   467
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   468
\<close>
493
Christian Urban <urbanc@in.tum.de>
parents: 492
diff changeset
   469
492
Christian Urban <urbanc@in.tum.de>
parents: 491
diff changeset
   470
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   471
(*
517
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 514
diff changeset
   472
ML %grayML{*Proof_Context.debug := true*}
d8c376662bb4 removed special ML-setup and replaced it by explicit markups (i.e., %grayML)
Christian Urban <urbanc@in.tum.de>
parents: 514
diff changeset
   473
ML %grayML{*Proof_Context.verbose := true*}
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   474
*)
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   475
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
   476
(*
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   477
lemma "True"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   478
proof -
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   479
  { -- "\<And>x. _"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   480
    fix x
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   481
    have "B x" sorry
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   482
    thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   483
  }
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   484
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   485
  thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   487
  { -- "A \<Longrightarrow> _"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   488
    assume A
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   489
    have B sorry
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   490
    thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   491
  }
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   492
   
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   493
  thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   494
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   495
  { -- "\<And>x. x = _ \<Longrightarrow> _"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   496
    def x \<equiv> a
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   497
    have "B x" sorry
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   498
  }
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   499
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   500
  thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   501
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   502
oops
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
   503
*)
413
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   504
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   505
section \<open>Local Theories and Local Setups\label{sec:local} (TBD)\<close>
341
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   506
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   507
text \<open>
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   508
  In contrast to an ordinary theory, which simply consists of a type
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   509
  signature, as well as tables for constants, axioms and theorems, a local
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   510
  theory contains additional context information, such as locally fixed
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   511
  variables and local assumptions that may be used by the package. The type
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   512
  @{ML_type local_theory} is identical to the type of \emph{proof contexts}
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   513
  @{ML_type "Proof.context"}, although not every proof context constitutes a
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   514
  valid local theory.
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   515
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   516
  @{ML \<open>Context.>> o Context.map_theory\<close>}
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   517
  @{ML_ind "Local_Theory.declaration"}
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   518
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   519
   A similar command is \isacommand{local\_setup}, which expects a function
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   520
  of type @{ML_type "local_theory -> local_theory"}. Later on we will also
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   521
  use the commands \isacommand{method\_setup} for installing methods in the
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   522
  current theory and \isacommand{simproc\_setup} for adding new simprocs to
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   523
  the current simpset.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   524
\<close>
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
336
a12bb28fe2bd polished on the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 335
diff changeset
   526
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   527
section \<open>Morphisms (TBD)\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   528
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   529
text \<open>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   530
  Morphisms are arbitrary transformations over terms, types, theorems and bindings.
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   531
  They can be constructed using the function @{ML_ind morphism in Morphism},
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   532
  which expects a record with functions of type
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   533
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   534
  \begin{isabelle}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   535
  \begin{tabular}{rl}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   536
  \<open>binding:\<close> & \<open>binding -> binding\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   537
  \<open>typ:\<close>     & \<open>typ -> typ\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   538
  \<open>term:\<close>    & \<open>term -> term\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   539
  \<open>fact:\<close>    & \<open>thm list -> thm list\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   540
  \end{tabular}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   541
  \end{isabelle}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   542
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   543
  The simplest morphism is the  @{ML_ind identity in Morphism}-morphism defined as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   544
\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   545
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   546
ML %grayML\<open>val identity = Morphism.morphism "" {binding = [], typ = [], term = [], fact = []}\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   547
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   548
text \<open>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   549
  Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   550
\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   551
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   552
ML %grayML\<open>fun trm_phi (Free (x, T)) = Var ((x, 0), T) 
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   553
  | trm_phi (Abs (x, T, t)) = Abs (x, T, trm_phi t)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   554
  | trm_phi (t $ s) = (trm_phi t) $ (trm_phi s)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   555
  | trm_phi t = t\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   556
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   557
ML %grayML\<open>val phi = Morphism.term_morphism "foo" trm_phi\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   558
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   559
ML %grayML\<open>Morphism.term phi @{term "P x y"}\<close>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   560
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   561
text \<open>
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   562
  @{ML_ind term_morphism in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   563
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   564
  @{ML_ind term in Morphism},
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   565
  @{ML_ind thm in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   566
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   567
  \begin{readmore}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   568
  Morphisms are implemented in the file @{ML_file "Pure/morphism.ML"}.
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   569
  \end{readmore}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   570
\<close>
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   572
section \<open>Misc (TBD)\<close>
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   574
text \<open>
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   575
FIXME: association lists:
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   576
@{ML_file "Pure/General/alist.ML"}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   577
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   578
FIXME: calling the ML-compiler
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   579
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   580
\<close>
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   581
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   582
section \<open>What Is In an Isabelle Name? (TBD)\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   583
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   584
text \<open>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   585
  On the ML-level of Isabelle, you often have to work with qualified names.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   586
  These are strings with some additional information, such as positional
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   587
  information and qualifiers. Such qualified names can be generated with the
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   588
  antiquotation \<open>@{binding \<dots>}\<close>. For example
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   589
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 565
diff changeset
   590
  @{ML_matchresult [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   591
  \<open>@{binding "name"}\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   592
  \<open>name\<close>}
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   593
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   594
  An example where a qualified name is needed is the function 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   595
  @{ML_ind define in Local_Theory}.  This function is used below to define 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   596
  the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   597
\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   598
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   599
local_setup %gray \<open>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   600
  Local_Theory.define ((@{binding "TrueConj"}, NoSyn), 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   601
      ((@{binding "TrueConj_def"}, []), @{term "True \<and> True"})) #> snd\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   602
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   603
text \<open>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   604
  Now querying the definition you obtain:
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   605
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   606
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   607
  \isacommand{thm}~\<open>TrueConj_def\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   608
  \<open>> \<close>~@{thm TrueConj_def}
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   609
  \end{isabelle}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   610
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   611
  \begin{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   612
  The basic operations on bindings are implemented in 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   613
  @{ML_file "Pure/General/binding.ML"}.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   614
  \end{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   615
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   616
  \footnote{\bf FIXME give a better example why bindings are important}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   617
  \footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   618
  why @{ML snd} is needed.}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   619
  \footnote{\bf FIXME: There should probably a separate section on binding, long-names
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   620
  and sign.}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   621
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   622
\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   623
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   625
ML \<open>Sign.intern_type @{theory} "list"\<close>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   626
ML \<open>Sign.intern_const @{theory} "prod_fun"\<close>
360
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   627
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   628
text \<open>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   629
  \footnote{\bf FIXME: Explain the following better; maybe put in a separate
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   630
  section and link with the comment in the antiquotation section.}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   631
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   632
  Occasionally you have to calculate what the ``base'' name of a given
462
1d1e795bc3ad updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   633
  constant is. For this you can use the function @{ML_ind  Long_Name.base_name}. For example:
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   634
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   635
  @{ML_matchresult [display,gray] \<open>Long_Name.base_name "List.list.Nil"\<close> \<open>"Nil"\<close>}
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   636
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   637
  \begin{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   638
  Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   639
  functions about signatures in @{ML_file "Pure/sign.ML"}.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   640
  \end{readmore}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   641
\<close>
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   642
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   643
text \<open>
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   644
  @{ML_ind "Binding.name_of"} returns the string without markup
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   645
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 560
diff changeset
   646
  @{ML_ind "Binding.concealed"} 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   647
\<close>
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   648
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   649
section \<open>Concurrency (TBD)\<close>
388
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   650
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   651
text \<open>
388
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   652
  @{ML_ind prove_future in Goal}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   653
  @{ML_ind future_result in Goal}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   654
\<close>
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   655
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   656
section \<open>Parse and Print Translations (TBD)\<close>
396
Christian Urban <urbanc@in.tum.de>
parents: 395
diff changeset
   657
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   658
section \<open>Summary\<close>
349
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   659
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   660
text \<open>
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   661
  TBD
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   662
\<close>
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
end