ProgTutorial/General.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 03 Oct 2009 13:01:39 +0200
changeset 328 c0cae24b9d46
parent 323 92f6a772e013
child 329 5dffcab68680
permissions -rw-r--r--
updated to new Isabelle; more work on the data section
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory General
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Base FirstSteps
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
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 319
diff changeset
     5
chapter {* Isabelle -- The Good, the Bad and the Ugly *}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
     7
text {*
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
     8
  Isabelle is build around a few central ideas. One is the LCF-approach to
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
     9
  theorem proving where there is a small trusted core and everything else is
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
    10
  build on top of this trusted core.
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
    11
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
    12
*}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
    13
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
    14
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
section {* Terms and Types *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
  One way to construct Isabelle terms, is by using the antiquotation 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
  \mbox{@{text "@{term \<dots>}"}}. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  @{ML_response [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
"@{term \"(a::nat) + b = c\"}" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
"Const (\"op =\", \<dots>) $ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
  (Const (\"HOL.plus_class.plus\", \<dots>) $ \<dots> $ \<dots>) $ \<dots>"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
  will show the term @{term "(a::nat) + b = c"}, but printed using the internal
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
  representation corresponding to the datatype @{ML_type  "term"} defined as follows: 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
*}  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
ML_val %linenosgray{*datatype term =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
  Const of string * typ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
| Free of string * typ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
| Var of indexname * typ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
| Bound of int 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
| Abs of string * typ * term 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
| $ of term * term *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
  This datatype implements lambda-terms typed in Church-style.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
  As can be seen in Line 5, terms use the usual de Bruijn index mechanism
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  for representing bound variables.  For
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  example in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  @{ML_response_fake [display, gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
  "@{term \"\<lambda>x y. x y\"}"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  "Abs (\"x\", \"'a \<Rightarrow> 'b\", Abs (\"y\", \"'a\", Bound 1 $ Bound 0))"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
  the indices refer to the number of Abstractions (@{ML Abs}) that we need to
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
  skip until we hit the @{ML Abs} that binds the corresponding
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  variable. Constructing a term with dangling de Bruijn indices is possible,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
  but will be flagged as ill-formed when you try to typecheck or certify it
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
  (see Section~\ref{sec:typechecking}). Note that the names of bound variables
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
  are kept at abstractions for printing purposes, and so should be treated
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  only as ``comments''.  Application in Isabelle is realised with the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
  term-constructor @{ML $}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
  Isabelle makes a distinction between \emph{free} variables (term-constructor
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  @{ML Free} and written on the user level in blue colour) and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
  \emph{schematic} variables (term-constructor @{ML Var} and written with a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  leading question mark). Consider the following two examples
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
  @{ML_response_fake [display, gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
  val v1 = Var ((\"x\", 3), @{typ bool})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
  val v2 = Var ((\"x1\", 3), @{typ bool})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
  val v3 = Free (\"x\", @{typ bool})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  string_of_terms @{context} [v1, v2, v3]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  |> tracing
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
"?x3, ?x1.3, x"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  When constructing terms, you are usually concerned with free variables (as
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  mentioned earlier, you cannot construct schematic variables using the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  antiquotation @{text "@{term \<dots>}"}). If you deal with theorems, you have to,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  however, observe the distinction. The reason is that only schematic
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
  varaibles can be instantiated with terms when a theorem is applied. A
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
  similar distinction between free and schematic variables holds for types
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  (see below).
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
  Terms and types are described in detail in \isccite{sec:terms}. Their
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  definition and many useful operations are implemented in @{ML_file "Pure/term.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  For constructing terms involving HOL constants, many helper functions are defined
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
  in @{ML_file "HOL/Tools/hologic.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
  Constructing terms via antiquotations has the advantage that only typable
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
  terms can be constructed. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
  @{ML_response_fake_both [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
  "@{term \"x x\"}"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  "Type unification failed: Occurs check!"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  raises a typing error, while it perfectly ok to construct the term
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
  @{ML_response_fake [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
  val omega = Free (\"x\", @{typ nat}) $ Free (\"x\", @{typ nat})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
in 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
  tracing (string_of_term @{context} omega)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
  "x x"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
  with the raw ML-constructors.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
  Sometimes the internal representation of terms can be surprisingly different
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  from what you see at the user-level, because the layers of
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
  parsing/type-checking/pretty printing can be quite elaborate. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
  \begin{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
  Look at the internal term representation of the following terms, and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
  find out why they are represented like this:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  \begin{itemize}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  \item @{term "case x of 0 \<Rightarrow> 0 | Suc y \<Rightarrow> y"}  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
  \item @{term "\<lambda>(x,y). P y x"}  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
  \item @{term "{ [x::int] | x. x \<le> -2 }"}  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
  \end{itemize}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  Hint: The third term is already quite big, and the pretty printer
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
  may omit parts of it by default. If you want to see all of it, you
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
  can use the following ML-function to set the printing depth to a higher 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  value:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
  @{ML [display,gray] "print_depth 50"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
  \end{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  The antiquotation @{text "@{prop \<dots>}"} constructs terms by inserting the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  usually invisible @{text "Trueprop"}-coercions whenever necessary. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  Consider for example the pairs
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
@{ML_response [display,gray] "(@{term \"P x\"}, @{prop \"P x\"})" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
"(Free (\"P\", \<dots>) $ Free (\"x\", \<dots>),
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
 Const (\"Trueprop\", \<dots>) $ (Free (\"P\", \<dots>) $ Free (\"x\", \<dots>)))"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
  where a coercion is inserted in the second component and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  @{ML_response [display,gray] "(@{term \"P x \<Longrightarrow> Q x\"}, @{prop \"P x \<Longrightarrow> Q x\"})" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  "(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
 Const (\"==>\", \<dots>) $ \<dots> $ \<dots>)"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
  where it is not (since it is already constructed by a meta-implication). 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  The purpose of the @{text "Trueprop"}-coercion is to embed formulae of
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
  an object logic, for example HOL, into the meta-logic of Isabelle. It
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  is needed whenever a term is constructed that will be proved as a theorem. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
  As already seen above, types can be constructed using the antiquotation 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  @{text "@{typ \<dots>}"}. For example:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
  @{ML_response_fake [display,gray] "@{typ \"bool \<Rightarrow> nat\"}" "bool \<Rightarrow> nat"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  The corresponding datatype is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
ML_val{*datatype typ =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
  Type  of string * typ list 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
| TFree of string * sort 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
| TVar  of indexname * sort *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
  Like with terms, there is the distinction between free type
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
  variables (term-constructor @{ML "TFree"} and schematic
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
  type variables (term-constructor @{ML "TVar"}). A type constant,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
  like @{typ "int"} or @{typ bool}, are types with an empty list
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  of argument types. However, it is a bit difficult to show an
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
  example, because Isabelle always pretty-prints types (unlike terms).
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  Here is a contrived example:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
  @{ML_response [display, gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
  "if Type (\"bool\", []) = @{typ \"bool\"} then true else false"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
  "true"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
  Types are described in detail in \isccite{sec:types}. Their
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
  definition and many useful operations are implemented 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  in @{ML_file "Pure/type.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
section {* Constructing Terms and Types Manually\label{sec:terms_types_manually} *} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
  While antiquotations are very convenient for constructing terms, they can
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  only construct fixed terms (remember they are ``linked'' at compile-time).
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
  However, you often need to construct terms dynamically. For example, a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
  function that returns the implication @{text "\<And>(x::nat). P x \<Longrightarrow> Q x"} taking
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
  @{term P} and @{term Q} as arguments can only be written as:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
ML{*fun make_imp P Q =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
  val x = Free ("x", @{typ nat})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
in 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
  Logic.all x (Logic.mk_implies (P $ x, Q $ x))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
end *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
  The reason is that you cannot pass the arguments @{term P} and @{term Q} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
  into an antiquotation.\footnote{At least not at the moment.} For example 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  the following does \emph{not} work.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
ML{*fun make_wrong_imp P Q = @{prop "\<And>(x::nat). P x \<Longrightarrow> Q x"} *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
  To see this, apply @{text "@{term S}"} and @{text "@{term T}"} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
  to both functions. With @{ML make_imp} you obtain the intended term involving 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
  the given arguments
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
  @{ML_response [display,gray] "make_imp @{term S} @{term T}" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
"Const \<dots> $ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
  Abs (\"x\", Type (\"nat\",[]),
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
    Const \<dots> $ (Free (\"S\",\<dots>) $ \<dots>) $ (Free (\"T\",\<dots>) $ \<dots>))"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
  whereas with @{ML make_wrong_imp} you obtain a term involving the @{term "P"} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
  and @{text "Q"} from the antiquotation.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
  @{ML_response [display,gray] "make_wrong_imp @{term S} @{term T}" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
"Const \<dots> $ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
  Abs (\"x\", \<dots>,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
    Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
                (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
  There are a number of handy functions that are frequently used for 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
  constructing terms. One is the function @{ML_ind  list_comb}, which takes a term
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
  and a list of terms as arguments, and produces as output the term
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  list applied to the term. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
  val trm = @{term \"P::nat\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
  val args = [@{term \"True\"}, @{term \"False\"}]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  list_comb (trm, args)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
"Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
  Another handy function is @{ML_ind  lambda}, which abstracts a variable 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
  in a term. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
  val x_nat = @{term \"x::nat\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
  val trm = @{term \"(P::nat \<Rightarrow> bool) x\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  lambda x_nat trm
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
  "Abs (\"x\", \"nat\", Free (\"P\", \"bool \<Rightarrow> bool\") $ Bound 0)"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  In this example, @{ML lambda} produces a de Bruijn index (i.e.~@{ML "Bound 0"}), 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
  and an abstraction. It also records the type of the abstracted
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
  variable and for printing purposes also its name.  Note that because of the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
  typing annotation on @{text "P"}, the variable @{text "x"} in @{text "P x"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
  is of the same type as the abstracted variable. If it is of different type,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
  as in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
  val x_int = @{term \"x::int\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
  val trm = @{term \"(P::nat \<Rightarrow> bool) x\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
  lambda x_int trm
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
  "Abs (\"x\", \"int\", Free (\"P\", \"nat \<Rightarrow> bool\") $ Free (\"x\", \"nat\"))"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
  then the variable @{text "Free (\"x\", \"int\")"} is \emph{not} abstracted. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
  This is a fundamental principle
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  of Church-style typing, where variables with the same name still differ, if they 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
  have different type.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
  There is also the function @{ML_ind  subst_free} with which terms can be
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  replaced by other terms. For example below, we will replace in @{term
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
  @{term y}, and @{term x} by @{term True}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
  val sub1 = (@{term \"(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0\"}, @{term \"y::nat \<Rightarrow> nat\"})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
  val sub2 = (@{term \"x::nat\"}, @{term \"True\"})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
  val trm = @{term \"((f::nat \<Rightarrow> nat \<Rightarrow> nat) 0) x\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
  subst_free [sub1, sub2] trm
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
  "Free (\"y\", \"nat \<Rightarrow> nat\") $ Const (\"True\", \"bool\")"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
  As can be seen, @{ML subst_free} does not take typability into account.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
  However it takes alpha-equivalence into account:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  @{ML_response_fake [display, gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
  val sub = (@{term \"(\<lambda>y::nat. y)\"}, @{term \"x::nat\"})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
  val trm = @{term \"(\<lambda>x::nat. x)\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
  subst_free [sub] trm
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
  "Free (\"x\", \"nat\")"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
  Similarly the function @{ML_ind  subst_bounds}, replaces lose bound 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  variables with terms. To see how this function works, let us implement a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
  function that strips off the outermost quantifiers in a term.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
ML{*fun strip_alls (Const ("All", _) $ Abs (n, T, t)) =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
         strip_alls t |>> cons (Free (n, T))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
  | strip_alls t = ([], t) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
  The function returns a pair consisting of the stripped off variables and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
  the body of the universal quantifications. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
  @{ML_response_fake [display, gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
  "strip_alls @{term \"\<forall>x y. x = (y::bool)\"}"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
"([Free (\"x\", \"bool\"), Free (\"y\", \"bool\")],
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
  Const (\"op =\", \<dots>) $ Bound 1 $ Bound 0)"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
  After calling @{ML strip_alls}, you obtain a term with lose bound variables. With
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
  the function @{ML subst_bounds}, you can replace these lose @{ML_ind 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
  Bound}s with the stripped off variables.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
  @{ML_response_fake [display, gray, linenos]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
  "let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  val (vrs, trm) = strip_alls @{term \"\<forall>x y. x = (y::bool)\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
in 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
  subst_bounds (rev vrs, trm) 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
  |> string_of_term @{context}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
  |> tracing
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
  "x = y"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  Note that in Line 4 we had to reverse the list of variables that @{ML strip_alls}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  returned. The reason is that the head of the list the function @{ML subst_bounds}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
  takes is the replacement for @{ML "Bound 0"}, the next element for @{ML "Bound 1"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
  and so on. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
  There are many convenient functions that construct specific HOL-terms. For
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
  example @{ML_ind  mk_eq in HOLogic} constructs an equality out of two terms.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
  The types needed in this equality are calculated from the type of the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
  arguments. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
@{ML_response_fake [gray,display]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
  val eq = HOLogic.mk_eq (@{term \"True\"}, @{term \"False\"})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
  string_of_term @{context} eq
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  |> tracing
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
  "True = False"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
  There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
  "Pure/logic.ML"} and @{ML_file "HOL/Tools/hologic.ML"} that make such manual
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
  constructions of terms and types easier.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
  When constructing terms manually, there are a few subtle issues with
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
  constants. They usually crop up when pattern matching terms or types, or
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
  when constructing them. While it is perfectly ok to write the function
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
  @{text is_true} as follows
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
ML{*fun is_true @{term True} = true
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
  | is_true _ = false*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
  this does not work for picking out @{text "\<forall>"}-quantified terms. Because
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
  the function 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
ML{*fun is_all (@{term All} $ _) = true
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
  | is_all _ = false*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
  will not correctly match the formula @{prop[source] "\<forall>x::nat. P x"}: 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
  @{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "false"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
  The problem is that the @{text "@term"}-antiquotation in the pattern 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
  fixes the type of the constant @{term "All"} to be @{typ "('a \<Rightarrow> bool) \<Rightarrow> bool"} for 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
  an arbitrary, but fixed type @{typ "'a"}. A properly working alternative 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
  for this function is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387
ML{*fun is_all (Const ("All", _) $ _) = true
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   388
  | is_all _ = false*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   389
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   390
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
  because now
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   393
@{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "true"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   394
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   395
  matches correctly (the first wildcard in the pattern matches any type and the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   396
  second any term).
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   397
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   398
  However there is still a problem: consider the similar function that
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   399
  attempts to pick out @{text "Nil"}-terms:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   400
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   401
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   402
ML{*fun is_nil (Const ("Nil", _)) = true
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   403
  | is_nil _ = false *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   405
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   406
  Unfortunately, also this function does \emph{not} work as expected, since
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   407
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   408
  @{ML_response [display,gray] "is_nil @{term \"Nil\"}" "false"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   409
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
  The problem is that on the ML-level the name of a constant is more
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
  subtle than you might expect. The function @{ML is_all} worked correctly,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   412
  because @{term "All"} is such a fundamental constant, which can be referenced
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   413
  by @{ML "Const (\"All\", some_type)" for some_type}. However, if you look at
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   414
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
  @{ML_response [display,gray] "@{term \"Nil\"}" "Const (\"List.list.Nil\", \<dots>)"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   416
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   417
  the name of the constant @{text "Nil"} depends on the theory in which the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   418
  term constructor is defined (@{text "List"}) and also in which datatype
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   419
  (@{text "list"}). Even worse, some constants have a name involving
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   420
  type-classes. Consider for example the constants for @{term "zero"} and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   421
  \mbox{@{text "(op *)"}}:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   422
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   423
  @{ML_response [display,gray] "(@{term \"0::nat\"}, @{term \"(op *)\"})" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   424
 "(Const (\"HOL.zero_class.zero\", \<dots>), 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   425
 Const (\"HOL.times_class.times\", \<dots>))"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
  While you could use the complete name, for example 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   428
  @{ML "Const (\"List.list.Nil\", some_type)" for some_type}, for referring to or
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   429
  matching against @{text "Nil"}, this would make the code rather brittle. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   430
  The reason is that the theory and the name of the datatype can easily change. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
  To make the code more robust, it is better to use the antiquotation 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   432
  @{text "@{const_name \<dots>}"}. With this antiquotation you can harness the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   433
  variable parts of the constant's name. Therefore a function for 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   434
  matching against constants that have a polymorphic type should 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   435
  be written as follows.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   436
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   437
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   438
ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   439
  | is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   440
  | is_nil_or_all _ = false *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   441
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   442
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   443
  The antiquotation for properly referencing type constants is @{text "@{type_name \<dots>}"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   444
  For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   445
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   446
  @{ML_response [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   447
  "@{type_name \"list\"}" "\"List.list\""}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   448
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   449
  (FIXME: Explain the following better.)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   450
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   451
  Occasionally you have to calculate what the ``base'' name of a given
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   452
  constant is. For this you can use the function @{ML_ind  "Sign.extern_const"} or
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
  @{ML_ind  Long_Name.base_name}. For example:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   454
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   455
  @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   456
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   457
  The difference between both functions is that @{ML extern_const in Sign} returns
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   458
  the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   459
  strips off all qualifiers.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   460
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   461
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
  Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   463
  functions about signatures in @{ML_file "Pure/sign.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   464
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   465
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   466
  Although types of terms can often be inferred, there are many
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   467
  situations where you need to construct types manually, especially  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   468
  when defining constants. For example the function returning a function 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   469
  type is as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   470
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   471
*} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   472
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   473
ML{*fun make_fun_type ty1 ty2 = Type ("fun", [ty1, ty2]) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   474
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   475
text {* This can be equally written with the combinator @{ML_ind  "-->"} as: *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   476
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   477
ML{*fun make_fun_type ty1 ty2 = ty1 --> ty2 *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   478
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   479
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   480
  If you want to construct a function type with more than one argument
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   481
  type, then you can use @{ML_ind  "--->"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   482
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   483
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   484
ML{*fun make_fun_types tys ty = tys ---> ty *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   485
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   486
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
  A handy function for manipulating terms is @{ML_ind  map_types}: it takes a 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   488
  function and applies it to every type in a term. You can, for example,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   489
  change every @{typ nat} in a term into an @{typ int} using the function:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
ML{*fun nat_to_int ty =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   493
  (case ty of
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   494
     @{typ nat} => @{typ int}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   495
   | Type (s, tys) => Type (s, map nat_to_int tys)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   496
   | _ => ty)*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   497
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   498
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   499
  Here is an example:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   500
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   501
@{ML_response_fake [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   502
"map_types nat_to_int @{term \"a = (1::nat)\"}" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   503
"Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\")
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   504
           $ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   505
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   506
  If you want to obtain the list of free type-variables of a term, you
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   507
  can use the function @{ML_ind  add_tfrees in Term} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
  (similarly @{ML_ind  add_tvars in Term} for the schematic type-variables). 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
  One would expect that such functions
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   510
  take a term as input and return a list of types. But their type is actually 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
  @{text[display] "Term.term -> (string * Term.sort) list -> (string * Term.sort) list"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   513
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   514
  that is they take, besides a term, also a list of type-variables as input. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   515
  So in order to obtain the list of type-variables of a term you have to 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   516
  call them as follows
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   517
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   518
  @{ML_response [gray,display]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   519
  "Term.add_tfrees @{term \"(a, b)\"} []"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   520
  "[(\"'b\", [\"HOL.type\"]), (\"'a\", [\"HOL.type\"])]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   521
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   522
  The reason for this definition is that @{ML add_tfrees in Term} can
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   523
  be easily folded over a list of terms. Similarly for all functions
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   524
  named @{text "add_*"} in @{ML_file "Pure/term.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   525
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   526
  \begin{exercise}\label{fun:revsum}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   527
  Write a function @{text "rev_sum : term -> term"} that takes a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   528
  term of the form @{text "t\<^isub>1 + t\<^isub>2 + \<dots> + t\<^isub>n"} (whereby @{text "n"} might be one)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   529
  and returns the reversed sum @{text "t\<^isub>n + \<dots> + t\<^isub>2 + t\<^isub>1"}. Assume
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   530
  the @{text "t\<^isub>i"} can be arbitrary expressions and also note that @{text "+"} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   531
  associates to the left. Try your function on some examples. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   532
  \end{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   533
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   534
  \begin{exercise}\label{fun:makesum}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   535
  Write a function which takes two terms representing natural numbers
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   536
  in unary notation (like @{term "Suc (Suc (Suc 0))"}), and produces the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   537
  number representing their sum.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   538
  \end{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   539
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   540
  \begin{exercise}\footnote{Personal communication of
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   541
  de Bruijn to Dyckhoff.}\label{ex:debruijn}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   542
  Implement the function, which we below name deBruijn, that depends on a natural
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   543
  number n$>$0 and constructs terms of the form:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   544
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   545
  \begin{center}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   546
  \begin{tabular}{r@ {\hspace{2mm}}c@ {\hspace{2mm}}l}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   547
  {\it rhs n} & $\dn$ & {\large$\bigwedge$}{\it i=1\ldots n.  P\,i}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   548
  {\it lhs n} & $\dn$ & {\large$\bigwedge$}{\it i=1\ldots n. P\,i = P (i + 1 mod n)}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   549
                        $\longrightarrow$ {\it rhs n}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   550
  {\it deBruijn n} & $\dn$ & {\it lhs n} $\longrightarrow$ {\it rhs n}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   551
  \end{tabular}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   552
  \end{center}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   553
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   554
  For n=3 this function returns the term
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   555
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   556
  \begin{center}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   557
  \begin{tabular}{l}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   558
  (P 1 = P 2 $\longrightarrow$ P 1 $\wedge$ P 2 $\wedge$ P 3) $\wedge$\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   559
  (P 2 = P 3 $\longrightarrow$ P 1 $\wedge$ P 2 $\wedge$ P 3) $\wedge$\\ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   560
  (P 3 = P 1 $\longrightarrow$ P 1 $\wedge$ P 2 $\wedge$ P 3) $\longrightarrow$ P 1 $\wedge$ P 2 $\wedge$ P 3
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   561
  \end{tabular}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   562
  \end{center}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   563
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   564
  Make sure you use the functions defined in @{ML_file "HOL/Tools/hologic.ML"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   565
  for constructing the terms for the logical connectives. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   566
  \end{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   567
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   568
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   569
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   570
section {* Type-Checking\label{sec:typechecking} *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   571
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   572
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   573
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   574
  You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   575
  typ}es, since they are just arbitrary unchecked trees. However, you
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   576
  eventually want to see if a term is well-formed, or type-checks, relative to
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   577
  a theory.  Type-checking is done via the function @{ML_ind  cterm_of}, which
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   578
  converts a @{ML_type  term} into a @{ML_type  cterm}, a \emph{certified}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   579
  term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   580
  abstract objects that are guaranteed to be type-correct, and they can only
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   581
  be constructed via ``official interfaces''.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   582
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   583
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   584
  Type-checking is always relative to a theory context. For now we use
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
  the @{ML "@{theory}"} antiquotation to get hold of the current theory.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
  For example you can write:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
  @{ML_response_fake [display,gray] "cterm_of @{theory} @{term \"(a::nat) + b = c\"}" "a + b = c"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
  This can also be written with an antiquotation:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
  @{ML_response_fake [display,gray] "@{cterm \"(a::nat) + b = c\"}" "a + b = c"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
  Attempting to obtain the certified term for
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
  @{ML_response_fake_both [display,gray] "@{cterm \"1 + True\"}" "Type unification failed \<dots>"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
  yields an error (since the term is not typable). A slightly more elaborate
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
  example that type-checks is:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
@{ML_response_fake [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   602
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   603
  val natT = @{typ \"nat\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   604
  val zero = @{term \"0::nat\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   605
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   606
  cterm_of @{theory} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   607
      (Const (@{const_name plus}, natT --> natT --> natT) $ zero $ zero)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   608
end" "0 + 0"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   609
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   610
  In Isabelle not just terms need to be certified, but also types. For example, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   611
  you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   612
  the ML-level as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   613
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   614
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   615
  "ctyp_of @{theory} (@{typ nat} --> @{typ bool})"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   616
  "nat \<Rightarrow> bool"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   617
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   618
  or with the antiquotation:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   619
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   620
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   621
  "@{ctyp \"nat \<Rightarrow> bool\"}"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   622
  "nat \<Rightarrow> bool"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   623
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   624
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   625
  For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   626
  the file @{ML_file "Pure/thm.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   627
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   628
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   629
  Remember Isabelle follows the Church-style typing for terms, i.e., a term contains 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   630
  enough typing information (constants, free variables and abstractions all have typing 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   631
  information) so that it is always clear what the type of a term is. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   632
  Given a well-typed term, the function @{ML_ind  type_of} returns the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   633
  type of a term. Consider for example:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   634
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   635
  @{ML_response [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   636
  "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   637
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   638
  To calculate the type, this function traverses the whole term and will
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   639
  detect any typing inconsistency. For example changing the type of the variable 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   640
  @{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message: 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   641
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   642
  @{ML_response_fake [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   643
  "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   644
  "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   645
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   646
  Since the complete traversal might sometimes be too costly and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   647
  not necessary, there is the function @{ML_ind  fastype_of}, which 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   648
  also returns the type of a term.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   649
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   650
  @{ML_response [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   651
  "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   652
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   653
  However, efficiency is gained on the expense of skipping some tests. You 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   654
  can see this in the following example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   655
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   656
   @{ML_response [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   657
  "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" "bool"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   658
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   659
  where no error is detected.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   660
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   661
  Sometimes it is a bit inconvenient to construct a term with 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   662
  complete typing annotations, especially in cases where the typing 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   663
  information is redundant. A short-cut is to use the ``place-holder'' 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   664
  type @{ML_ind  dummyT} and then let type-inference figure out the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   665
  complete type. An example is as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   666
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   667
  @{ML_response_fake [display,gray] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   668
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   669
  val c = Const (@{const_name \"plus\"}, dummyT) 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   670
  val o = @{term \"1::nat\"} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   671
  val v = Free (\"x\", dummyT)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   672
in   
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   673
  Syntax.check_term @{context} (c $ o $ v)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
"Const (\"HOL.plus_class.plus\", \"nat \<Rightarrow> nat \<Rightarrow> nat\") $
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   676
  Const (\"HOL.one_class.one\", \"nat\") $ Free (\"x\", \"nat\")"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   677
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   678
  Instead of giving explicitly the type for the constant @{text "plus"} and the free 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   679
  variable @{text "x"}, type-inference fills in the missing information.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   680
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   681
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   682
  See @{ML_file "Pure/Syntax/syntax.ML"} where more functions about reading,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   683
  checking and pretty-printing of terms are defined. Functions related to
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   684
  type-inference are implemented in @{ML_file "Pure/type.ML"} and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   685
  @{ML_file "Pure/type_infer.ML"}. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   686
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   687
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   688
  (FIXME: say something about sorts)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   689
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   690
  \begin{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
  Check that the function defined in Exercise~\ref{fun:revsum} returns a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
  result that type-checks. See what happens to the  solutions of this 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
  exercise given in \ref{ch:solutions} when they receive an ill-typed term
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
  as input.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
  \end{exercise}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   698
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
section {* Theorems *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
  Just like @{ML_type cterm}s, theorems are abstract objects of type @{ML_type thm} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
  that can only be built by going through interfaces. As a consequence, every proof 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
  in Isabelle is correct by construction. This follows the tradition of the LCF approach
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
  \cite{GordonMilnerWadsworth79}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
  To see theorems in ``action'', let us give a proof on the ML-level for the following 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
  statement:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
  lemma 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
   assumes assm\<^isub>1: "\<And>(x::nat). P x \<Longrightarrow> Q x" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
   and     assm\<^isub>2: "P t"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
   shows "Q t" (*<*)oops(*>*) 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
  The corresponding ML-code is as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
  val assm1 = @{cprop \"\<And>(x::nat). P x \<Longrightarrow> Q x\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
  val assm2 = @{cprop \"(P::nat\<Rightarrow>bool) t\"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
  val Pt_implies_Qt = 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
        assume assm1
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
        |> forall_elim @{cterm \"t::nat\"};
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   729
  val Qt = implies_elim Pt_implies_Qt (assume assm2);
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
  Qt 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
  |> implies_intr assm2
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
  |> implies_intr assm1
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
end" "\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
  This code-snippet constructs the following proof:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
  \[
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
  \infer[(@{text "\<Longrightarrow>"}$-$intro)]{\vdash @{prop "(\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> P t \<Longrightarrow> Q t"}}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   740
    {\infer[(@{text "\<Longrightarrow>"}$-$intro)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   741
       {\infer[(@{text "\<Longrightarrow>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"}, @{prop "P t"} \vdash @{prop "Q t"}}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
          {\infer[(@{text "\<And>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
                 {\infer[(assume)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "\<And>x. P x \<Longrightarrow> Q x"}}{}}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
                 &
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
           \infer[(assume)]{@{prop "P t"} \vdash @{prop "P t"}}{}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
          }
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
       }
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
    }
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
  \]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
  However, while we obtained a theorem as result, this theorem is not
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
  yet stored in Isabelle's theorem database. So it cannot be referenced later
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
  on. How to store theorems will be explained in Section~\ref{sec:storing}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
  For the functions @{text "assume"}, @{text "forall_elim"} etc 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
  see \isccite{sec:thms}. The basic functions for theorems are defined in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
  @{ML_file "Pure/thm.ML"}. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   759
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
  (FIXME: handy functions working on theorems, like @{ML_ind  rulify in ObjectLogic} and so on) 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
  (FIXME: @{ML_ind prove in Goal})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   764
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   765
  (FIXME: how to add case-names to goal states - maybe in the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
  next section)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   768
  (FIXME: example for how to add theorem styles)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
ML {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
fun strip_assums_all (params, Const("all",_) $ Abs(a, T, t)) =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
      strip_assums_all ((a, T)::params, t)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
  | strip_assums_all (params, B) = (params, B)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   775
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
fun style_parm_premise i ctxt t =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
  let val prems = Logic.strip_imp_prems t in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
    if i <= length prems
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
    then let val (params,t) = strip_assums_all([], nth prems (i - 1))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
         in subst_bounds(map Free params, t) end
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
    else error ("Not enough premises for prem" ^ string_of_int i ^
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
      " in propositon: " ^ string_of_term ctxt t)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
  end;
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   786
ML {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
strip_assums_all ([], @{term "\<And>x y. A x y"})
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
setup %gray {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
  TermStyle.add_style "no_all_prem1" (style_parm_premise 1) #>
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
  TermStyle.add_style "no_all_prem2" (style_parm_premise 2)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   794
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
lemma 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
  shows "A \<Longrightarrow> B"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
  and   "C \<Longrightarrow> D"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   798
oops
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   799
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   800
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   801
section {* Setups (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   802
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   803
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
  In the previous section we used \isacommand{setup} in order to make
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805
  a theorem attribute known to Isabelle. What happens behind the scenes
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   806
  is that \isacommand{setup} expects a function of type 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   807
  @{ML_type "theory -> theory"}: the input theory is the current theory and the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   808
  output the theory where the theory attribute has been stored.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   809
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   810
  This is a fundamental principle in Isabelle. A similar situation occurs 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   811
  for example with declaring constants. The function that declares a 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
  constant on the ML-level is @{ML_ind  add_consts_i in Sign}. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
  If you write\footnote{Recall that ML-code  needs to be 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
  enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
*}  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
  for declaring the constant @{text "BAR"} with type @{typ nat} and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
  run the code, then you indeed obtain a theory as result. But if you 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
  query the constant on the Isabelle level using the command \isacommand{term}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
  \isacommand{term}~@{text [quotes] "BAR"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
  @{text "> \"BAR\" :: \"'a\""}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
  you do not obtain a constant of type @{typ nat}, but a free variable (printed in 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
  blue) of polymorphic type. The problem is that the ML-expression above did 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
  not register the declaration with the current theory. This is what the command
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
  \isacommand{setup} is for. The constant is properly declared with
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
setup %gray {* Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   836
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   837
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   838
  Now 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   839
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   841
  \isacommand{term}~@{text [quotes] "BAR"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   842
  @{text "> \"BAR\" :: \"nat\""}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   843
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   844
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   845
  returns a (black) constant with the type @{typ nat}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   846
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   847
  A similar command is \isacommand{local\_setup}, which expects a function
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   848
  of type @{ML_type "local_theory -> local_theory"}. Later on we will also
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   849
  use the commands \isacommand{method\_setup} for installing methods in the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   850
  current theory and \isacommand{simproc\_setup} for adding new simprocs to
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   851
  the current simpset.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   852
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   853
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   854
section {* Theorem Attributes\label{sec:attributes} *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   855
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   856
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   857
  Theorem attributes are @{text "[symmetric]"}, @{text "[THEN \<dots>]"}, @{text
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   858
  "[simp]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   859
  annotated to theorems, but functions that do further processing once a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   860
  theorem is proved. In particular, it is not possible to find out
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   861
  what are all theorems that have a given attribute in common, unless of course
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   862
  the function behind the attribute stores the theorems in a retrievable 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   863
  data structure. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   864
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   865
  If you want to print out all currently known attributes a theorem can have, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   866
  you can use the Isabelle command
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   867
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   868
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   869
  \isacommand{print\_attributes}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   870
  @{text "> COMP:  direct composition with rules (no lifting)"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   871
  @{text "> HOL.dest:  declaration of Classical destruction rule"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   872
  @{text "> HOL.elim:  declaration of Classical elimination rule"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   873
  @{text "> \<dots>"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   874
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   875
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   876
  The theorem attributes fall roughly into two categories: the first category manipulates
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   877
  the proved theorem (for example @{text "[symmetric]"} and @{text "[THEN \<dots>]"}), and the second
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   878
  stores the proved theorem somewhere as data (for example @{text "[simp]"}, which adds
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   879
  the theorem to the current simpset).
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   880
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   881
  To explain how to write your own attribute, let us start with an extremely simple 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   882
  version of the attribute @{text "[symmetric]"}. The purpose of this attribute is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   883
  to produce the ``symmetric'' version of an equation. The main function behind 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   884
  this attribute is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   885
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   886
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   887
ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   888
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   889
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   890
  where the function @{ML_ind  rule_attribute in Thm} expects a function taking a 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   891
  context (which we ignore in the code above) and a theorem (@{text thm}), and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   892
  returns another theorem (namely @{text thm} resolved with the theorem 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   893
  @{thm [source] sym}: @{thm sym[no_vars]}; the function @{ML_ind "RS"} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   894
  is explained in Section~\ref{sec:simpletacs}). The function 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   895
  @{ML rule_attribute in Thm} then returns  an attribute.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   896
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   897
  Before we can use the attribute, we need to set it up. This can be done
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   898
  using the Isabelle command \isacommand{attribute\_setup} as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   899
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   900
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   901
attribute_setup %gray my_sym = {* Scan.succeed my_symmetric *} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   902
  "applying the sym rule"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   903
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   904
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   905
  Inside the @{text "\<verbopen> \<dots> \<verbclose>"}, we have to specify a parser
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   906
  for the theorem attribute. Since the attribute does not expect any further 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   907
  arguments (unlike @{text "[THEN \<dots>]"}, for example), we use the parser @{ML
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   908
  Scan.succeed}. Later on we will also consider attributes taking further
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   909
  arguments. An example for the attribute @{text "[my_sym]"} is the proof
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   910
*} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   911
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   912
lemma test[my_sym]: "2 = Suc (Suc 0)" by simp
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   913
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   914
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   915
  which stores the theorem @{thm test} under the name @{thm [source] test}. You
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   916
  can see this, if you query the lemma: 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   917
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   918
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   919
  \isacommand{thm}~@{text "test"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   920
  @{text "> "}~@{thm test}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   921
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   922
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   923
  We can also use the attribute when referring to this theorem:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   924
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   925
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   926
  \isacommand{thm}~@{text "test[my_sym]"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   927
  @{text "> "}~@{thm test[my_sym]}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   928
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   929
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   930
  An alternative for setting up an attribute is the function @{ML_ind  setup in Attrib}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   931
  So instead of using \isacommand{attribute\_setup}, you can also set up the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   932
  attribute as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   933
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   934
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   935
ML{*Attrib.setup @{binding "my_sym"} (Scan.succeed my_symmetric)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   936
  "applying the sym rule" *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   937
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   938
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   939
  This gives a function from @{ML_type "Context.theory -> Context.theory"}, which
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   940
  can be used for example with \isacommand{setup}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   941
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   942
  As an example of a slightly more complicated theorem attribute, we implement 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   943
  our own version of @{text "[THEN \<dots>]"}. This attribute will take a list of theorems
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   944
  as argument and resolve the proved theorem with this list (one theorem 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   945
  after another). The code for this attribute is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   946
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   947
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   948
ML{*fun MY_THEN thms = 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   949
  Thm.rule_attribute (fn _ => fn thm => foldl ((op RS) o swap) thm thms)*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   950
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   951
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   952
  where @{ML swap} swaps the components of a pair. The setup of this theorem
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   953
  attribute uses the parser @{ML thms in Attrib}, which parses a list of
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   954
  theorems. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   955
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   956
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   957
attribute_setup %gray MY_THEN = {* Attrib.thms >> MY_THEN *} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   958
  "resolving the list of theorems with the proved theorem"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   959
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   960
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   961
  You can, for example, use this theorem attribute to turn an equation into a 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   962
  meta-equation:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   963
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   964
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   965
  \isacommand{thm}~@{text "test[MY_THEN eq_reflection]"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   966
  @{text "> "}~@{thm test[MY_THEN eq_reflection]}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   967
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   968
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   969
  If you need the symmetric version as a meta-equation, you can write
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   970
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   971
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   972
  \isacommand{thm}~@{text "test[MY_THEN sym eq_reflection]"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   973
  @{text "> "}~@{thm test[MY_THEN sym eq_reflection]}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   974
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   975
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   976
  It is also possible to combine different theorem attributes, as in:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   977
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   978
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   979
  \isacommand{thm}~@{text "test[my_sym, MY_THEN eq_reflection]"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   980
  @{text "> "}~@{thm test[my_sym, MY_THEN eq_reflection]}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   981
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   983
  However, here also a weakness of the concept
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   984
  of theorem attributes shows through: since theorem attributes can be
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   985
  arbitrary functions, they do not in general commute. If you try
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   986
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   987
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   988
  \isacommand{thm}~@{text "test[MY_THEN eq_reflection, my_sym]"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   989
  @{text "> "}~@{text "exception THM 1 raised: RSN: no unifiers"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   990
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   991
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   992
  you get an exception indicating that the theorem @{thm [source] sym}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   993
  does not resolve with meta-equations. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   994
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   995
  The purpose of @{ML_ind  rule_attribute in Thm} is to directly manipulate theorems.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   996
  Another usage of theorem attributes is to add and delete theorems from stored data.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   997
  For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   998
  current simpset. For these applications, you can use @{ML_ind  declaration_attribute in Thm}. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   999
  To illustrate this function, let us introduce a reference containing a list
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1000
  of theorems.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1001
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1002
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
  1003
ML{*val my_thms = Unsynchronized.ref ([] : thm list)*}
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1004
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1005
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1006
  The purpose of this reference is to store a list of theorems.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1007
  We are going to modify it by adding and deleting theorems.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1008
  However, a word of warning: such references must not 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1009
  be used in any code that is meant to be more than just for testing purposes! 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1010
  Here it is only used to illustrate matters. We will show later how to store 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1011
  data properly without using references.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1012
 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1013
  We need to provide two functions that add and delete theorems from this list. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1014
  For this we use the two functions:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1015
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1016
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1017
ML{*fun my_thm_add thm ctxt =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1018
  (my_thms := Thm.add_thm thm (!my_thms); ctxt)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1019
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1020
fun my_thm_del thm ctxt =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1021
  (my_thms := Thm.del_thm thm (!my_thms); ctxt)*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1022
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1023
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1024
  These functions take a theorem and a context and, for what we are explaining
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1025
  here it is sufficient that they just return the context unchanged. They change
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1026
  however the reference @{ML my_thms}, whereby the function 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1027
  @{ML_ind  add_thm in Thm} adds a theorem if it is not already included in 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1028
  the list, and @{ML_ind  del_thm in Thm} deletes one (both functions use the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1029
  predicate @{ML_ind  eq_thm_prop in Thm}, which compares theorems according to 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1030
  their proved propositions modulo alpha-equivalence).
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1031
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1032
  You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1033
  attributes with the code
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1034
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
ML{*val my_add = Thm.declaration_attribute my_thm_add
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1037
val my_del = Thm.declaration_attribute my_thm_del *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1038
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1039
text {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1040
  and set up the attributes as follows
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1041
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1042
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1043
attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
  "maintaining a list of my_thms - rough test only!" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1046
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
  The parser @{ML_ind  add_del in Attrib} is a predefined parser for 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1048
  adding and deleting lemmas. Now if you prove the next lemma 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1049
  and attach to it the attribute @{text "[my_thms]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1050
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
lemma trueI_2[my_thms]: "True" by simp
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
  then you can see it is added to the initially empty list.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1058
  "!my_thms" "[\"True\"]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
  You can also add theorems using the command \isacommand{declare}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1061
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1062
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
declare test[my_thms] trueI_2[my_thms add] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1066
  With this attribute, the @{text "add"} operation is the default and does 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
  not need to be explicitly given. These three declarations will cause the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1068
  theorem list to be updated as:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1070
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1071
  "!my_thms"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
  "[\"True\", \"Suc (Suc 0) = 2\"]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1073
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1074
  The theorem @{thm [source] trueI_2} only appears once, since the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
  function @{ML_ind  add_thm in Thm} tests for duplicates, before extending
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1076
  the list. Deletion from the list works as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1078
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1079
declare test[my_thms del]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1081
text {* After this, the theorem list is again: 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1082
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1084
  "!my_thms"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1085
  "[\"True\"]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1087
  We used in this example two functions declared as @{ML_ind  declaration_attribute in Thm}, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
  but there can be any number of them. We just have to change the parser for reading
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
  the arguments accordingly. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1090
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
  However, as said at the beginning of this example, using references for storing theorems is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
  \emph{not} the received way of doing such things. The received way is to 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
  start a ``data slot'', below called @{text MyThmsData}, generated by the functor 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1094
  @{text GenericDataFun}:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1097
ML{*structure MyThmsData = GenericDataFun
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
 (type T = thm list
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1099
  val empty = []
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1100
  val extend = I
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
  fun merge _ = Thm.merge_thms) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
  The type @{text "T"} of this data slot is @{ML_type "thm list"}.\footnote{FIXME: give a pointer
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
  to where data slots are explained properly.}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
  To use this data slot, you only have to change @{ML my_thm_add} and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1107
  @{ML my_thm_del} to:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1109
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1110
ML{*val my_thm_add = MyThmsData.map o Thm.add_thm
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1111
val my_thm_del = MyThmsData.map o Thm.del_thm*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1112
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1113
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1114
  where @{ML MyThmsData.map} updates the data appropriately. The
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
  corresponding theorem attributes are
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1116
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1117
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1118
ML{*val my_add = Thm.declaration_attribute my_thm_add
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1119
val my_del = Thm.declaration_attribute my_thm_del *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1120
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1121
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1122
  and the setup is as follows
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1123
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1124
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1125
attribute_setup %gray my_thms2 = {* Attrib.add_del my_add my_del *} 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1126
  "properly maintaining a list of my_thms"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1127
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1128
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1129
  Initially, the data slot is empty 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1130
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1131
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1132
  "MyThmsData.get (Context.Proof @{context})"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1133
  "[]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1134
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1135
  but if you prove
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1136
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1137
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1138
lemma three[my_thms2]: "3 = Suc (Suc (Suc 0))" by simp
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1139
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1140
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1141
  then the lemma is recorded. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1142
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1143
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1144
  "MyThmsData.get (Context.Proof @{context})"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1145
  "[\"3 = Suc (Suc (Suc 0))\"]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1146
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1147
  With theorem attribute @{text my_thms2} you can also nicely see why it 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1148
  is important to 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1149
  store data in a ``data slot'' and \emph{not} in a reference. Backtrack
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1150
  to the point just before the lemma @{thm [source] three} was proved and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1151
  check the the content of @{ML_struct MyThmsData}: it should be empty. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1152
  The addition has been properly retracted. Now consider the proof:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1153
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1154
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1155
lemma four[my_thms]: "4 = Suc (Suc (Suc (Suc 0)))" by simp
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1156
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1157
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1158
  Checking the content of @{ML my_thms} gives
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1159
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1160
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1161
  "!my_thms"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1162
  "[\"4 = Suc (Suc (Suc (Suc 0)))\", \"True\"]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1163
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1164
  as expected, but if you backtrack before the lemma @{thm [source] four}, the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1165
  content of @{ML my_thms} is unchanged. The backtracking mechanism
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1166
  of Isabelle is completely oblivious about what to do with references, but
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1167
  properly treats ``data slots''!
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1168
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1169
  Since storing theorems in a list is such a common task, there is the special
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1170
  functor @{ML_functor Named_Thms}, which does most of the work for you. To obtain
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1171
  a named theorem list, you just declare
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1172
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1173
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1174
ML{*structure FooRules = Named_Thms
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1175
 (val name = "foo" 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1176
  val description = "Rules for foo") *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1177
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1178
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1179
  and set up the @{ML_struct FooRules} with the command
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1180
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1181
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1182
setup %gray {* FooRules.setup *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1183
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1184
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1185
  This code declares a data slot where the theorems are stored,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1186
  an attribute @{text foo} (with the @{text add} and @{text del} options
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1187
  for adding and deleting theorems) and an internal ML interface to retrieve and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1188
  modify the theorems.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1189
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1190
  Furthermore, the facts are made available on the user-level under the dynamic 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1191
  fact name @{text foo}. For example you can declare three lemmas to be of the kind
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1192
  @{text foo} by:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1193
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1194
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1195
lemma rule1[foo]: "A" sorry
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1196
lemma rule2[foo]: "B" sorry
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1197
lemma rule3[foo]: "C" sorry
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1198
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1199
text {* and undeclare the first one by: *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1200
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1201
declare rule1[foo del]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1202
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1203
text {* and query the remaining ones with:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1204
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1205
  \begin{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1206
  \isacommand{thm}~@{text "foo"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1207
  @{text "> ?C"}\\
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1208
  @{text "> ?B"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1209
  \end{isabelle}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1210
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1211
  On the ML-level the rules marked with @{text "foo"} can be retrieved
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1212
  using the function @{ML FooRules.get}:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1213
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1214
  @{ML_response_fake [display,gray] "FooRules.get @{context}" "[\"?C\",\"?B\"]"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1215
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1216
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1217
  For more information see @{ML_file "Pure/Tools/named_thms.ML"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1218
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1219
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1220
  (FIXME What are: @{text "theory_attributes"}, @{text "proof_attributes"}?)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1221
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1222
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1223
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1224
  FIXME: @{ML_file "Pure/more_thm.ML"}; parsers for attributes is in 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1225
  @{ML_file "Pure/Isar/attrib.ML"}...also explained in the chapter about
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1226
  parsing.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1227
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1228
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1229
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1230
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1231
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1232
section {* Theories, Contexts and Local Theories (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1233
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1234
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1235
  There are theories, proof contexts and local theories (in this order, if you
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1236
  want to order them). 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1237
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1238
  In contrast to an ordinary theory, which simply consists of a type
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1239
  signature, as well as tables for constants, axioms and theorems, a local
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1240
  theory contains additional context information, such as locally fixed
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1241
  variables and local assumptions that may be used by the package. The type
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1242
  @{ML_type local_theory} is identical to the type of \emph{proof contexts}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1243
  @{ML_type "Proof.context"}, although not every proof context constitutes a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1244
  valid local theory.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1245
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1246
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1247
section {* Storing Theorems\label{sec:storing} (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1248
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1249
text {* @{ML_ind  add_thms_dynamic in PureThy} *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1250
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1251
local_setup %gray {* 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1252
  LocalTheory.note Thm.theoremK 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1253
    ((@{binding "allI_alt"}, []), [@{thm allI}]) #> snd *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1254
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1255
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1256
(* FIXME: some code below *)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1257
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1258
(*<*)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1259
(*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1260
setup {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1261
 Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1262
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1263
*)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1264
lemma "bar = (1::nat)"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1265
  oops
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1266
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1267
(*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1268
setup {*   
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1269
  Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1270
 #> PureThy.add_defs false [((@{binding "foo_def"}, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1271
       Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])] 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1272
 #> snd
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1273
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1274
*)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1275
(*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1276
lemma "foo = (1::nat)"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1277
  apply(simp add: foo_def)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1278
  done
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1279
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1280
thm foo_def
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1281
*)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1282
(*>*)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1283
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1284
section {* Pretty-Printing\label{sec:pretty} *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1285
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1286
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1287
  So far we printed out only plain strings without any formatting except for
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1288
  occasional explicit line breaks using @{text [quotes] "\\n"}. This is
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1289
  sufficient for ``quick-and-dirty'' printouts. For something more
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1290
  sophisticated, Isabelle includes an infrastructure for properly formatting text.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1291
  This infrastructure is loosely based on a paper by Oppen~\cite{Oppen80}. Most of
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1292
  its functions do not operate on @{ML_type string}s, but on instances of the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1293
  type:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1294
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1295
  @{ML_type_ind [display, gray] "Pretty.T"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1296
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1297
  The function @{ML str in Pretty} transforms a (plain) string into such a pretty 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1298
  type. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1299
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1300
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1301
  "Pretty.str \"test\"" "String (\"test\", 4)"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1302
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1303
  where the result indicates that we transformed a string with length 4. Once
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1304
  you have a pretty type, you can, for example, control where linebreaks may
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1305
  occur in case the text wraps over a line, or with how much indentation a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1306
  text should be printed.  However, if you want to actually output the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1307
  formatted text, you have to transform the pretty type back into a @{ML_type
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1308
  string}. This can be done with the function @{ML_ind  string_of in Pretty}. In what
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1309
  follows we will use the following wrapper function for printing a pretty
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1310
  type:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1311
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1312
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1313
ML{*fun pprint prt = tracing (Pretty.string_of prt)*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1314
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1315
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1316
  The point of the pretty-printing infrastructure is to give hints about how to
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1317
  layout text and let Isabelle do the actual layout. Let us first explain
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1318
  how you can insert places where a line break can occur. For this assume the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1319
  following function that replicates a string n times:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1320
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1321
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1322
ML{*fun rep n str = implode (replicate n str) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1323
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1324
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1325
  and suppose we want to print out the string:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1326
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1327
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1328
ML{*val test_str = rep 8 "fooooooooooooooobaaaaaaaaaaaar "*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1329
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1330
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1331
  We deliberately chose a large string so that it spans over more than one line. 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1332
  If we print out the string using the usual ``quick-and-dirty'' method, then
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1333
  we obtain the ugly output:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1334
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1335
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1336
"tracing test_str"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1337
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1338
ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1339
aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1340
oooooooooooooobaaaaaaaaaaaar"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1341
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1342
  We obtain the same if we just use
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1343
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1344
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1345
"pprint (Pretty.str test_str)"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1346
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1347
ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1348
aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1349
oooooooooooooobaaaaaaaaaaaar"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1350
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1351
  However by using pretty types you have the ability to indicate a possible
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1352
  line break for example at each space. You can achieve this with the function
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1353
  @{ML_ind  breaks in Pretty}, which expects a list of pretty types and inserts a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1354
  possible line break in between every two elements in this list. To print
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1355
  this list of pretty types as a single string, we concatenate them 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1356
  with the function @{ML_ind  blk in Pretty} as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1357
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1358
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1359
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1360
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1361
  val ptrs = map Pretty.str (space_explode \" \" test_str)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1362
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1363
  pprint (Pretty.blk (0, Pretty.breaks ptrs))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1364
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1365
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1366
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1367
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1368
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1369
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1370
  Here the layout of @{ML test_str} is much more pleasing to the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1371
  eye. The @{ML "0"} in @{ML_ind  blk in Pretty} stands for no indentation
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1372
  of the printed string. You can increase the indentation and obtain
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1373
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1374
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1375
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1376
  val ptrs = map Pretty.str (space_explode \" \" test_str)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1377
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1378
  pprint (Pretty.blk (3, Pretty.breaks ptrs))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1379
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1380
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1381
   fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1382
   fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1383
   fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1384
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1385
  where starting from the second line the indent is 3. If you want
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1386
  that every line starts with the same indent, you can use the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1387
  function @{ML_ind  indent in Pretty} as follows:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1388
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1389
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1390
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1391
  val ptrs = map Pretty.str (space_explode \" \" test_str)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1392
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1393
  pprint (Pretty.indent 10 (Pretty.blk (0, Pretty.breaks ptrs)))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1394
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1395
"          fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1396
          fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1397
          fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1398
          fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1399
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1400
  If you want to print out a list of items separated by commas and 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1401
  have the linebreaks handled properly, you can use the function 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1402
  @{ML_ind  commas in Pretty}. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1403
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1404
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1405
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1406
  val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1407
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1408
  pprint (Pretty.blk (0, Pretty.commas ptrs))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1409
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1410
"99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1411
100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1412
100016, 100017, 100018, 100019, 100020"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1413
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1414
  where @{ML upto} generates a list of integers. You can print out this
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1415
  list as a ``set'', that means enclosed inside @{text [quotes] "{"} and
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1416
  @{text [quotes] "}"}, and separated by commas using the function
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1417
  @{ML_ind  enum in Pretty}. For example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1418
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1419
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1420
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1421
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1422
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1423
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1424
  val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1425
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1426
  pprint (Pretty.enum \",\" \"{\" \"}\" ptrs)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1427
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1428
"{99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1429
  100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1430
  100016, 100017, 100018, 100019, 100020}"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1431
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1432
  As can be seen, this function prints out the ``set'' so that starting 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1433
  from the second, each new line as an indentation of 2.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1434
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1435
  If you print out something that goes beyond the capabilities of the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1436
  standard functions, you can do relatively easily the formatting
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1437
  yourself. Assume you want to print out a list of items where like in ``English''
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1438
  the last two items are separated by @{text [quotes] "and"}. For this you can
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1439
  write the function
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1440
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1441
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1442
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1443
ML %linenosgray{*fun and_list [] = []
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1444
  | and_list [x] = [x]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1445
  | and_list xs = 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1446
      let 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1447
        val (front, last) = split_last xs
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1448
      in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1449
        (Pretty.commas front) @ 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1450
        [Pretty.brk 1, Pretty.str "and", Pretty.brk 1, last]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1451
      end *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1452
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1453
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1454
  where Line 7 prints the beginning of the list and Line
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1455
  8 the last item. We have to use @{ML "Pretty.brk 1"} in order
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1456
  to insert a space (of length 1) before the 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1457
  @{text [quotes] "and"}. This space is also a place where a line break 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1458
  can occur. We do the same after the @{text [quotes] "and"}. This gives you
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1459
  for example
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1460
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1461
@{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1462
"let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1463
  val ptrs = map (Pretty.str o string_of_int) (1 upto 22)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1464
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1465
  pprint (Pretty.blk (0, and_list ptrs))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1466
end"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1467
"1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1468
and 22"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1469
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1470
  Next we like to pretty-print a term and its type. For this we use the
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1471
  function @{text "tell_type"}.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1472
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1473
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1474
ML %linenosgray{*fun tell_type ctxt t = 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1475
let
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1476
  fun pstr s = Pretty.breaks (map Pretty.str (space_explode " " s))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1477
  val ptrm = Pretty.quote (Syntax.pretty_term ctxt t)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1478
  val pty  = Pretty.quote (Syntax.pretty_typ ctxt (fastype_of t))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1479
in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1480
  pprint (Pretty.blk (0, 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1481
    (pstr "The term " @ [ptrm] @ pstr " has type " 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1482
      @ [pty, Pretty.str "."])))
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1483
end*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1484
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1485
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1486
  In Line 3 we define a function that inserts possible linebreaks in places
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1487
  where a space is. In Lines 4 and 5 we pretty-print the term and its type
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1488
  using the functions @{ML_ind  pretty_term in Syntax} and @{ML_ind 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1489
  pretty_typ in Syntax}. We also use the function @{ML_ind quote in
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1490
  Pretty} in order to enclose the term and type inside quotation marks. In
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1491
  Line 9 we add a period right after the type without the possibility of a
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1492
  line break, because we do not want that a line break occurs there.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1493
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1494
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1495
  Now you can write
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1496
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1497
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1498
  "tell_type @{context} @{term \"min (Suc 0)\"}"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1499
  "The term \"min (Suc 0)\" has type \"nat \<Rightarrow> nat\"."}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1500
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1501
  To see the proper line breaking, you can try out the function on a bigger term 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1502
  and type. For example:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1503
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1504
  @{ML_response_fake [display,gray]
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1505
  "tell_type @{context} @{term \"op = (op = (op = (op = (op = op =))))\"}"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1506
  "The term \"op = (op = (op = (op = (op = op =))))\" has type 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1507
\"((((('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool) \<Rightarrow> bool\"."}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1508
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1509
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1510
  FIXME: TBD below
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1511
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1512
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1513
ML {* pprint (Pretty.big_list "header"  (map (Pretty.str o string_of_int) (4 upto 10))) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1514
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1515
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1516
chunks inserts forced breaks (unlike blk where you have to do this yourself)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1517
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1518
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1519
ML {* (Pretty.chunks [Pretty.str "a", Pretty.str "b"], 
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1520
       Pretty.blk (0, [Pretty.str "a", Pretty.str "b"])) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1521
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1522
ML {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1523
fun setmp_show_all_types f =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1524
   setmp show_all_types
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1525
         (! show_types orelse ! show_sorts orelse ! show_all_types) f;
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1526
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1527
val ctxt = @{context};
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1528
val t1 = @{term "undefined::nat"};
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1529
val t2 = @{term "Suc y"};
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1530
val pty =        Pretty.block (Pretty.breaks
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1531
             [(setmp show_question_marks false o setmp_show_all_types)
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1532
                  (Syntax.pretty_term ctxt) t1,
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1533
              Pretty.str "=", Syntax.pretty_term ctxt t2]);
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1534
pty |> Pretty.string_of |> priority
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1535
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1536
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1537
text {* the infrastructure of Syntax-pretty-term makes sure it is printed nicely  *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1538
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1539
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1540
ML {* Pretty.mark Markup.hilite (Pretty.str "foo") |> Pretty.string_of  |> tracing *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1541
ML {* (Pretty.str "bar") |> Pretty.string_of |> tracing *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1542
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1543
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1544
ML {* Pretty.mark Markup.subgoal (Pretty.str "foo") |> Pretty.string_of  |> tracing *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1545
ML {* (Pretty.str "bar") |> Pretty.string_of |> tracing *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1546
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1547
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1548
  Still to be done:
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1549
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1550
  What happens with big formulae?
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1551
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1552
  \begin{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1553
  The general infrastructure for pretty-printing is implemented in the file
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1554
  @{ML_file "Pure/General/pretty.ML"}. The file @{ML_file "Pure/Syntax/syntax.ML"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1555
  contains pretty-printing functions for terms, types, theorems and so on.
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1556
  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1557
  @{ML_file "Pure/General/markup.ML"}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1558
  \end{readmore}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1559
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1560
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1561
text {*
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1562
  printing into the goal buffer as part of the proof state
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1563
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1564
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1565
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1566
ML {* Pretty.mark Markup.hilite (Pretty.str "foo") |> Pretty.string_of  |> tracing *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1567
ML {* (Pretty.str "bar") |> Pretty.string_of |> tracing *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1568
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1569
text {* writing into the goal buffer *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1570
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1571
ML {* fun my_hook interactive state =
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1572
         (interactive ? Proof.goal_message (fn () => Pretty.str  
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1573
"foo")) state
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1574
*}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1575
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1576
setup %gray {* Context.theory_map (Specification.add_theorem_hook my_hook) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1577
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1578
lemma "False"
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1579
oops
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1580
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1581
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1582
section {* Misc (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1583
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1584
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
  1585
319
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1586
text {* 
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1587
FIXME: association lists:
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1588
@{ML_file "Pure/General/alist.ML"}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1589
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1590
FIXME: calling the ML-compiler
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1591
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1592
*}
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1593
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1594
6bce4acf7f2a added file for producing a keyword file
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1595
318
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1596
section {* Name Space Issues (TBD) *}
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1597
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1598
efb5fff99c96 split up the first-steps section into two chapters
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1599
end