ProgTutorial/Advanced.thy
author Christian Urban <urbanc@in.tum.de>
Mon, 07 Nov 2011 16:13:26 +0000
changeset 489 1343540ed715
parent 488 780100cd4060
child 490 b8a654eabdf0
permissions -rw-r--r--
tuned
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
(*<*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
     6
setup{*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
     7
open_file_with_prelude 
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
     8
  "Advanced_Code.thy"
441
Christian Urban <urbanc@in.tum.de>
parents: 414
diff changeset
     9
  ["theory Advanced", "imports Base First_Steps", "begin"]
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    10
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    11
(*>*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    12
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 345
diff changeset
    13
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
    14
chapter {* Advanced Isabelle\label{chp:advanced} *}
381
97518188ef0e added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents: 380
diff changeset
    15
97518188ef0e added more to the unification section
Christian Urban <urbanc@in.tum.de>
parents: 380
diff changeset
    16
text {*
410
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    17
   \begin{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    18
  {\em All things are difficult before they are easy.} \\[1ex]
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    19
  proverb
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    20
  \end{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    21
Christian Urban <urbanc@in.tum.de>
parents: 408
diff changeset
    22
  \medskip
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    23
  While terms, types and theorems are the most basic data structures in
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    24
  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
    25
  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
    26
  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
    27
  The most basic layer are theories. They contain global data and
408
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    28
  can be seen as the ``long-term memory'' of Isabelle. There is usually only
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    29
  one theory active at each moment. Proof contexts and local theories, on the
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
    30
  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
    31
  ``short-term memory'' and there can be many of them that are active in
Christian Urban <urbanc@in.tum.de>
parents: 401
diff changeset
    32
  parallel.
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    35
section {* Theories and Setups\label{sec:theories} *}
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    36
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    37
text {*
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    38
  Theories, as said above, are a basic layer of abstraction in
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    39
  Isabelle. They contain definitions, syntax declarations, axioms,
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    40
  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
    41
  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
    42
  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
    43
  be stored in the theorem database of the theory in order to be
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    44
  usable. All relevant data of a theory can be queried as follows.
358
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    45
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    46
  \begin{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    47
  \isacommand{print\_theory}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    48
  @{text "> names: Pure Code_Generator HOL \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    49
  @{text "> classes: Inf < type \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    50
  @{text "> default sort: type"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    51
  @{text "> syntactic types: #prop \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    52
  @{text "> logical types: 'a \<times> 'b \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    53
  @{text "> type arities: * :: (random, random) random \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    54
  @{text "> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    55
  @{text "> abbreviations: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    56
  @{text "> axioms: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    57
  @{text "> oracles: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    58
  @{text "> definitions: \<dots>"}\\
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    59
  @{text "> theorems: \<dots>"}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    60
  \end{isabelle}
9cf3bc448210 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 356
diff changeset
    61
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    62
  Functions acting on theories often end with the suffix @{text "_global"},
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    63
  for example the function @{ML read_term_global in Syntax} in the structure
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    64
  @{ML_struct Syntax}. The reason is to set them syntactically apart from
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    65
  functions acting on contexts or local theories (which will be discussed in
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    66
  the next sections). There is a tendency in the Isabelle development to prefer
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    67
  ``non-global'' operations, because they have some advantages, as we will also
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    68
  discuss later. However, some basic management of theories will very likely
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    69
  never go away.
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
    70
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    71
  In the context of ML-programming, the most important command with theories
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    72
  is \isacommand{setup}. In the previous chapters we used it, for
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    73
  example, to make a theorem attribute known to Isabelle or to register a theorem
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    74
  under a name. What happens behind the scenes is that \isacommand{setup}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    75
  expects a function of type @{ML_type "theory -> theory"}: the input theory
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    76
  is the current theory and the output the theory where the attribute has been
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    77
  registered or the theorem has been stored.  This is a fundamental principle
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    78
  in Isabelle. A similar situation arises with declaring constants. The
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    79
  function that declares a constant on the ML-level is @{ML_ind declare_const
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    80
  in Sign} in the structure @{ML_struct Sign}. To see how \isacommand{setup}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    81
  works, consider the following code: 
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    82
*}  
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    83
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    84
ML{*let
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    85
  val thy = @{theory}
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    86
  val bar_const = ((@{binding "BAR"}, @{typ "nat"}), NoSyn)
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    87
in 
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    88
  Sign.declare_const @{context} bar_const thy  
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    89
end*}
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    90
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    91
text {*
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    92
  If you simply run this code\footnote{Recall that ML-code needs to be enclosed in
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
    93
  \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} with the
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    94
  intention of declaring a constant @{text "BAR"} with type @{typ nat}, then 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
    95
  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
    96
  constant on the Isabelle level using the command \isacommand{term}
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    97
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    98
  \begin{isabelle}
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    99
  \isacommand{term}~@{text BAR}\\
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   100
  @{text "> \"BAR\" :: \"'a\""}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   101
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   102
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   103
  you can see that you do \emph{not} obtain a constant of type @{typ nat}, but a free 
484
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 481
diff changeset
   104
  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
   105
  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
   106
  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
   107
  declared with
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   108
*}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   109
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   110
setup %gray {* fn thy => 
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   111
let
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   112
  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
   113
  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
   114
in 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   115
  thy'
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   116
end *}
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   117
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   118
text {* 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   119
  where the declaration is actually applied to the current theory and
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   120
  
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   121
  \begin{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   122
  \isacommand{term}~@{text [quotes] "BAR"}\\
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   123
  @{text "> \"BAR\" :: \"nat\""}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   124
  \end{isabelle}
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   125
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   126
  returns now a (black) constant with the type @{typ nat}, as expected.
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   127
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   128
  In a sense, \isacommand{setup} can be seen as a transaction that takes the
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   129
  current theory, applies an operation, and produces a new current theory. This
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   130
  means that we have to be careful to apply operations always to the most current
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   131
  theory, not to a \emph{stale} one. Consider again the function inside the
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   132
  \isacommand{setup}-command:
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   133
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   134
  \begin{isabelle}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   135
  \begin{graybox}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   136
  \isacommand{setup}~@{text "\<verbopen>"} @{ML
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   137
"fn thy => 
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   138
let
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   139
  val bar_const = ((@{binding \"BAR\"}, @{typ \"nat\"}), NoSyn)
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   140
  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
   141
in
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   142
  thy
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   143
end"}~@{text "\<verbclose>"}\isanewline
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   144
  @{text "> ERROR \"Stale theory encountered\""}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   145
  \end{graybox}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   146
  \end{isabelle}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   147
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   148
  This time we erroneously return the original theory @{text thy}, instead of
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   149
  the modified one @{text thy'}. Such buggy code will always result into 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   150
  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
   151
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   152
  However, sometimes it does make sense to work with two theories at the same
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   153
  time, especially in the context of parsing and typing. In the code below we
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   154
  use in Line 3 the function @{ML_ind copy in Theory} from the structure
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   155
  @{ML_struct Theory} for obtaining a new theory that contains the same
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   156
  data, but is unrelated to the existing theory.
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   157
*}
485
f3536f0b47a9 added section about testboard
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
   158
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   159
setup %graylinenos {* fn thy => 
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   160
let
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   161
  val tmp_thy = Theory.copy thy
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   162
  val foo_const = ((@{binding "FOO"}, @{typ "nat => nat"}), NoSyn)
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   163
  val (_, tmp_thy') = Sign.declare_const @{context} foo_const tmp_thy
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   164
  val trm1 = Syntax.read_term_global tmp_thy' "FOO baz"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   165
  val trm2 = Syntax.read_term_global thy "FOO baz"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   166
  val _ = writeln (@{make_string} trm1 ^ "\n" ^ @{make_string} trm2)
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   167
in
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   168
  thy
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   169
end *}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   170
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   171
text {*
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   172
  That means we can make changes to the theory @{text tmp_thy} without
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   173
  affecting the current theory @{text thy}. In this case we declare in @{text
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   174
  "tmp_thy"} the constant @{text FOO} (Lines 4 and 5). The point of this code
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   175
  is that we next, in Lines 6 and 7, parse a string to become a term (both
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   176
  times the string is @{text [quotes] "FOO baz"}). But since we parse the string
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   177
  once in the context of the theory @{text tmp_thy'} in which @{text FOO} is
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   178
  declared to be a constant of type @{typ "nat \<Rightarrow>nat"} and once in the context 
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   179
  of @{text thy} where it is not, we obtain two different terms, namely 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   180
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   181
  \begin{isabelle}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   182
  \begin{graybox}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   183
  @{text "> Const (\"Advanced.FOO\", \"nat \<Rightarrow> nat\") $ Free (\"baz\", \"nat\")"}\isanewline
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   184
  @{text "> Free (\"FOO\", \"'a \<Rightarrow> 'b\") $ Free (\"baz\", \"'a\")"}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   185
  \end{graybox}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   186
  \end{isabelle}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   187
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   188
  There are two reasons for parsing a term in a temporary theory. One is to
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   189
  obtain fully qualified names for constants and the other is appropriate type 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   190
  inference. This is relevant in situations where definitions are made later, 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   191
  but parsing and type inference has to take already proceed as if the definitions 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   192
  were already made.
348
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   193
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
341
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   195
section {* Contexts (TBD) *}
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   196
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   197
text {*
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   198
  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
   199
  contain ``short-term memory data''. The reason is that a vast number of
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   200
  functions in Isabelle depend one way or the other on contexts. Even such
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   201
  mundane operation like printing out a term make essential use of contexts.
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   202
  For this consider the following contrived proof which only purpose is to 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   203
  fix two variables 
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   204
*}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   205
488
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   206
lemma "True"
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   207
proof -
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   208
  txt_raw {*\mbox{}\\[-6mm]*} 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   209
  ML_prf {* Variable.dest_fixes @{context} *} 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   210
  txt_raw {*\mbox{}\\[-6mm]*}
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   211
  fix x y  
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   212
  txt_raw {*\mbox{}\\[-6mm]*}
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   213
  ML_prf {* Variable.dest_fixes @{context} *} 
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   214
  txt_raw {*\mbox{}\\[-6mm]*}
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   215
oops
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   216
780100cd4060 more on contexts
Christian Urban <urbanc@in.tum.de>
parents: 487
diff changeset
   217
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   218
(*
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   219
ML{*Proof_Context.debug := true*}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   220
ML{*Proof_Context.verbose := true*}
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   221
*)
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   222
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
   223
(*
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   224
lemma "True"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   225
proof -
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   226
  { -- "\<And>x. _"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   227
    fix x
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   228
    have "B x" sorry
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   229
    thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   230
  }
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   231
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   232
  thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   233
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   234
  { -- "A \<Longrightarrow> _"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   235
    assume A
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   236
    have B sorry
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   237
    thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   238
  }
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   239
   
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   240
  thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   241
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   242
  { -- "\<And>x. x = _ \<Longrightarrow> _"
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   243
    def x \<equiv> a
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   244
    have "B x" sorry
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   245
  }
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   246
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   247
  thm this
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   248
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   249
oops
487
1c4250bc6258 more on theories
Christian Urban <urbanc@in.tum.de>
parents: 486
diff changeset
   250
*)
413
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 410
diff changeset
   251
341
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   252
section {* Local Theories (TBD) *}
62dea749d5ed more work on theorem section
Christian Urban <urbanc@in.tum.de>
parents: 340
diff changeset
   253
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   254
text {*
400
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   255
  In contrast to an ordinary theory, which simply consists of a type
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   256
  signature, as well as tables for constants, axioms and theorems, a local
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   257
  theory contains additional context information, such as locally fixed
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   258
  variables and local assumptions that may be used by the package. The type
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   259
  @{ML_type local_theory} is identical to the type of \emph{proof contexts}
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   260
  @{ML_type "Proof.context"}, although not every proof context constitutes a
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   261
  valid local theory.
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   262
Christian Urban <urbanc@in.tum.de>
parents: 396
diff changeset
   263
  @{ML "Context.>> o Context.map_theory"}
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   264
  @{ML_ind "Local_Theory.declaration"}
486
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   265
45cfd2ece7bd a section about theories and setups
Christian Urban <urbanc@in.tum.de>
parents: 485
diff changeset
   266
   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
   267
  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
   268
  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
   269
  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
   270
  the current simpset.
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   271
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
336
a12bb28fe2bd polished on the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 335
diff changeset
   273
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   274
section {* Morphisms (TBD) *}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   275
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   276
text {*
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   277
  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
   278
  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
   279
  which expects a record with functions of type
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   280
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   281
  \begin{isabelle}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   282
  \begin{tabular}{rl}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   283
  @{text "binding:"} & @{text "binding -> binding"}\\
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   284
  @{text "typ:"}     & @{text "typ -> typ"}\\
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   285
  @{text "term:"}    & @{text "term -> term"}\\
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   286
  @{text "fact:"}    & @{text "thm list -> thm list"}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   287
  \end{tabular}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   288
  \end{isabelle}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   289
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   290
  The simplest morphism is the  @{ML_ind identity in Morphism}-morphism defined as
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   291
*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   292
481
32323727af96 updated
Christian Urban <urbanc@in.tum.de>
parents: 475
diff changeset
   293
ML{*val identity = Morphism.morphism {binding = [], typ = [], term = [], fact = []}*}
394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   294
  
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   295
text {*
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   296
  Morphisms can be composed with the function @{ML_ind "$>" in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   297
*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   298
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   299
ML{*fun trm_phi (Free (x, T)) = Var ((x, 0), T) 
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   300
  | 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
   301
  | trm_phi (t $ s) = (trm_phi t) $ (trm_phi s)
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   302
  | trm_phi t = t*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   303
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   304
ML{*val phi = Morphism.term_morphism trm_phi*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   305
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   306
ML{*Morphism.term phi @{term "P x y"}*}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   307
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   308
text {*
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   309
  @{ML_ind term_morphism in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   310
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   311
  @{ML_ind term in Morphism},
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   312
  @{ML_ind thm in Morphism}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   313
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   314
  \begin{readmore}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   315
  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
   316
  \end{readmore}
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   317
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
section {* Misc (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
ML {*Datatype.get_info @{theory} "List.list"*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   323
text {* 
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   324
FIXME: association lists:
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   325
@{ML_file "Pure/General/alist.ML"}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   326
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   327
FIXME: calling the ML-compiler
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   328
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   329
*}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   330
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   331
section {* What Is In an Isabelle Name? (TBD) *}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   332
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   333
text {*
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   334
  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
   335
  These are strings with some additional information, such as positional
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   336
  information and qualifiers. Such qualified names can be generated with the
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   337
  antiquotation @{text "@{binding \<dots>}"}. For example
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   338
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   339
  @{ML_response [display,gray]
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   340
  "@{binding \"name\"}"
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   341
  "name"}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   342
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   343
  An example where a qualified name is needed is the function 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   344
  @{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
   345
  the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   346
*}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   347
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   348
local_setup %gray {* 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   349
  Local_Theory.define ((@{binding "TrueConj"}, NoSyn), 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   350
      (Attrib.empty_binding, @{term "True \<and> True"})) #> snd *}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   351
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   352
text {* 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   353
  Now querying the definition you obtain:
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   354
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   355
  \begin{isabelle}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   356
  \isacommand{thm}~@{text "TrueConj_def"}\\
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   357
  @{text "> "}~@{thm TrueConj_def}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   358
  \end{isabelle}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   359
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   360
  \begin{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   361
  The basic operations on bindings are implemented in 
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   362
  @{ML_file "Pure/General/binding.ML"}.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   363
  \end{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   364
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   365
  \footnote{\bf FIXME give a better example why bindings are important}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   366
  \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
   367
  why @{ML snd} is needed.}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   368
  \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
   369
  and sign.}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   370
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   371
*}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   372
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
360
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   374
ML {* Sign.intern_type @{theory} "list" *}
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   375
ML {* Sign.intern_const @{theory} "prod_fun" *}
Christian Urban <urbanc@in.tum.de>
parents: 359
diff changeset
   376
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   377
text {*
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   378
  \footnote{\bf FIXME: Explain the following better; maybe put in a separate
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   379
  section and link with the comment in the antiquotation section.}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   380
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   381
  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
   382
  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
   383
462
1d1e795bc3ad updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   384
  @{ML_response [display,gray] "Long_Name.base_name \"List.list.Nil\"" "\"Nil\""}
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   385
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   386
  \begin{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   387
  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
   388
  functions about signatures in @{ML_file "Pure/sign.ML"}.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   389
  \end{readmore}
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   390
*}
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   391
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   392
text {* 
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   393
  @{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
   394
0019ebf76e10 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 393
diff changeset
   395
  @{ML_ind "Binding.conceal"} 
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   396
*}
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   397
388
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   398
section {* Concurrency (TBD) *}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   399
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   400
text {*
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   401
  @{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
   402
  @{ML_ind future_result in Goal}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   403
  @{ML_ind fork_pri in Future}
0b337dedc306 added Skip_Proof.mk_thm and some pointers about concurrency
Christian Urban <urbanc@in.tum.de>
parents: 387
diff changeset
   404
*}
387
5dcee4d751ad completed the unification section
Christian Urban <urbanc@in.tum.de>
parents: 386
diff changeset
   405
396
Christian Urban <urbanc@in.tum.de>
parents: 395
diff changeset
   406
section {* Parse and Print Translations (TBD) *}
Christian Urban <urbanc@in.tum.de>
parents: 395
diff changeset
   407
349
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   408
section {* Summary *}
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   409
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   410
text {*
395
2c392f61f400 spilt the Essential's chapter
Christian Urban <urbanc@in.tum.de>
parents: 394
diff changeset
   411
  TBD
349
9e374cd891e1 updated to Isabelle changes
Christian Urban <urbanc@in.tum.de>
parents: 348
diff changeset
   412
*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
end