| author | Christian Urban <urbanc@in.tum.de> | 
| Tue, 10 Nov 2009 12:17:08 +0100 | |
| changeset 383 | 72a53b07d264 | 
| parent 382 | 3f153aa4f231 | 
| child 384 | 0b24a016f6dd | 
| permissions | -rw-r--r-- | 
| 
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  | 
|
| 
346
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
5  | 
(*<*)  | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
6  | 
setup{*
 | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
7  | 
open_file_with_prelude  | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
8  | 
"General_Code.thy"  | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
9  | 
["theory General", "imports Base FirstSteps", "begin"]  | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
10  | 
*}  | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
11  | 
(*>*)  | 
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
12  | 
|
| 
 
0fea8b7a14a1
tuned the ML-output mechanism; tuned slightly the text
 
Christian Urban <urbanc@in.tum.de> 
parents: 
345 
diff
changeset
 | 
13  | 
|
| 358 | 14  | 
chapter {* Isabelle Essentials *}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
15  | 
|
| 
319
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
16  | 
text {*
 | 
| 345 | 17  | 
Isabelle is build around a few central ideas. One central idea is the  | 
18  | 
LCF-approach to theorem proving where there is a small trusted core and  | 
|
| 351 | 19  | 
everything else is build on top of this trusted core  | 
20  | 
  \cite{GordonMilnerWadsworth79}. The fundamental data
 | 
|
| 350 | 21  | 
structures involved in this core are certified terms and certified types,  | 
22  | 
as well as theorems.  | 
|
| 
319
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
23  | 
*}  | 
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
24  | 
|
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
25  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
26  | 
section {* Terms and Types *}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
28  | 
text {*
 | 
| 350 | 29  | 
In Isabelle, there are certified terms and uncertified terms (respectively types).  | 
30  | 
Uncertified terms are often just called terms. One way to construct them is by  | 
|
| 329 | 31  | 
  using the antiquotation \mbox{@{text "@{term \<dots>}"}}. For example
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
33  | 
  @{ML_response [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
34  | 
"@{term \"(a::nat) + b = c\"}" 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
35  | 
"Const (\"op =\", \<dots>) $  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
36  | 
(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
 | 
37  | 
|
| 350 | 38  | 
  constructs the term @{term "(a::nat) + b = c"}. The resulting term is printed using 
 | 
39  | 
  the internal representation corresponding to the datatype @{ML_type_ind "term"}, 
 | 
|
40  | 
which is defined as follows:  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
41  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
42  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
43  | 
ML_val %linenosgray{*datatype term =
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
44  | 
Const of string * typ  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
45  | 
| Free of string * typ  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
46  | 
| Var of indexname * typ  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
47  | 
| Bound of int  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
48  | 
| Abs of string * typ * term  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
49  | 
| $ of term * term *}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
50  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
51  | 
text {*
 | 
| 345 | 52  | 
This datatype implements Church-style lambda-terms, where types are  | 
| 350 | 53  | 
explicitly recorded in variables, constants and abstractions. As can be  | 
| 345 | 54  | 
seen in Line 5, terms use the usual de Bruijn index mechanism for  | 
55  | 
representing bound variables. For example in  | 
|
| 
318
 
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  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
58  | 
  "@{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
 | 
59  | 
"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
 | 
60  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
61  | 
  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
 | 
62  | 
  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
 | 
63  | 
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
 | 
64  | 
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
 | 
65  | 
  (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
 | 
66  | 
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
 | 
67  | 
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
 | 
68  | 
  term-constructor @{ML $}.
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
69  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
70  | 
  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
 | 
71  | 
  @{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
 | 
72  | 
  \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
 | 
73  | 
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
 | 
74  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
75  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
76  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
77  | 
  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
 | 
78  | 
  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
 | 
79  | 
  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
 | 
80  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
81  | 
  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
 | 
82  | 
|> tracing  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
83  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
84  | 
"?x3, ?x1.3, x"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
85  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
86  | 
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
 | 
87  | 
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
 | 
88  | 
  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
 | 
89  | 
however, observe the distinction. The reason is that only schematic  | 
| 345 | 90  | 
variables can be instantiated with terms when a theorem is applied. A  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
91  | 
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
 | 
92  | 
(see below).  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
93  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
94  | 
  \begin{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
95  | 
  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
 | 
96  | 
  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
 | 
97  | 
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
 | 
98  | 
  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
 | 
99  | 
  \end{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
101  | 
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
 | 
102  | 
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
 | 
103  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
104  | 
  @{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
 | 
105  | 
  "@{term \"x x\"}"
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
106  | 
"Type unification failed: Occurs check!"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
107  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
108  | 
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
 | 
109  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
110  | 
  @{ML_response_fake [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
111  | 
"let  | 
| 345 | 112  | 
  val omega = Free (\"x\", @{typ \"nat \<Rightarrow> nat\"}) $ Free (\"x\", @{typ nat})
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
113  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
114  | 
  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
 | 
115  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
116  | 
"x x"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
117  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
118  | 
with the raw ML-constructors.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
119  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
120  | 
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
 | 
121  | 
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
 | 
122  | 
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
 | 
123  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
124  | 
  \begin{exercise}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
125  | 
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
 | 
126  | 
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
 | 
127  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
128  | 
  \begin{itemize}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
129  | 
  \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
 | 
130  | 
  \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
 | 
131  | 
  \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
 | 
132  | 
  \end{itemize}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
133  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
134  | 
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
 | 
135  | 
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
 | 
136  | 
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
 | 
137  | 
value:  | 
| 
 
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 [display,gray] "print_depth 50"}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
140  | 
  \end{exercise}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
141  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
142  | 
  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
 | 
143  | 
  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
 | 
144  | 
Consider for example the pairs  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
145  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
146  | 
@{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
 | 
147  | 
"(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
 | 
148  | 
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
 | 
149  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
150  | 
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
 | 
151  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
152  | 
  @{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
 | 
153  | 
"(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>,  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
154  | 
Const (\"==>\", \<dots>) $ \<dots> $ \<dots>)"}  | 
| 
 
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  | 
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
 | 
157  | 
  The purpose of the @{text "Trueprop"}-coercion is to embed formulae of
 | 
| 350 | 158  | 
an object logic, for example HOL, into the meta-logic of Isabelle. The coercion  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
159  | 
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
 | 
160  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
161  | 
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
 | 
162  | 
  @{text "@{typ \<dots>}"}. For example:
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
163  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
164  | 
  @{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
 | 
165  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
166  | 
The corresponding datatype is  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
167  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
168  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
169  | 
ML_val{*datatype typ =
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
170  | 
Type of string * typ list  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
171  | 
| TFree of string * sort  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
172  | 
| TVar of indexname * sort *}  | 
| 
 
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  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
175  | 
Like with terms, there is the distinction between free type  | 
| 350 | 176  | 
  variables (term-constructor @{ML "TFree"}) and schematic
 | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
177  | 
  type variables (term-constructor @{ML "TVar"} and printed with
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
178  | 
a leading question mark). A type constant,  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
179  | 
  like @{typ "int"} or @{typ bool}, are types with an empty list
 | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
180  | 
of argument types. However, it needs a bit of effort to show an  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
181  | 
example, because Isabelle always pretty prints types (unlike terms).  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
182  | 
  Using just the antiquotation @{text "@{typ \"bool\"}"} we only see
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
183  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
184  | 
  @{ML_response [display, gray]
 | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
185  | 
  "@{typ \"bool\"}"
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
186  | 
"bool"}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
187  | 
|
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
188  | 
  that is the pretty printed version of @{text "bool"}. However, in PolyML it is 
 | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
189  | 
easy to install your own pretty printer. With the function below we  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
190  | 
mimic the behaviour of the usual pretty printer for  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
191  | 
  datatypes.\footnote{Thanks to David Matthews for providing this
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
192  | 
code.}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
193  | 
*}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
194  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
195  | 
ML{*fun typ_raw_pretty_printer depth _ ty =
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
196  | 
let  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
197  | 
fun cond str a =  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
198  | 
if depth <= 0  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
199  | 
then PolyML.PrettyString "..."  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
200  | 
else PolyML.PrettyBlock(1, false, [],  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
201  | 
[PolyML.PrettyString str, PolyML.PrettyBreak(1, 0), a])  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
202  | 
in  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
203  | 
case ty of  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
204  | 
Type a => cond "Type" (PolyML.prettyRepresentation(a, depth - 1))  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
205  | 
| TFree a => cond "TFree" (PolyML.prettyRepresentation(a, depth - 1))  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
206  | 
| TVar a => cond "TVar" (PolyML.prettyRepresentation(a, depth - 1))  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
207  | 
end*}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
208  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
209  | 
text {*
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
210  | 
We can install this pretty printer with the function  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
211  | 
  @{ML_ind addPrettyPrinter in PolyML} as follows.
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
212  | 
*}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
213  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
214  | 
ML{*PolyML.addPrettyPrinter typ_raw_pretty_printer*}
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
215  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
216  | 
text {*
 | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
217  | 
Now the type bool is printed out in full detail.  | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
218  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
219  | 
  @{ML_response [display,gray]
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
220  | 
  "@{typ \"bool\"}"
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
221  | 
"Type (\"bool\", [])"}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
222  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
223  | 
When printing out a list-type  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
224  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
225  | 
  @{ML_response [display,gray]
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
226  | 
  "@{typ \"'a list\"}"
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
227  | 
"Type (\"List.list\", [TFree (\"'a\", [\"HOL.type\"])])"}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
228  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
229  | 
  we can see the full name of the type is actually @{text "List.list"}, indicating
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
230  | 
  that it is defined in the theory @{text "List"}. However, one has to be 
 | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
231  | 
careful with names of types, because even if  | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
232  | 
  @{text "fun"}, @{text "bool"} and @{text "nat"} are defined in the 
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
233  | 
  theories @{text "HOL"} and @{text "Nat"}, respectively, they are 
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
234  | 
still represented by their simple name.  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
235  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
236  | 
   @{ML_response [display,gray]
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
237  | 
  "@{typ \"bool \<Rightarrow> nat\"}"
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
238  | 
"Type (\"fun\", [Type (\"bool\", []), Type (\"nat\", [])])"}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
239  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
240  | 
We can restore the usual behaviour of Isabelle's pretty printer  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
241  | 
with the code  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
242  | 
*}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
243  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
244  | 
ML{*fun stnd_pretty_printer _ _ =   
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
245  | 
ml_pretty o Pretty.to_ML o Proof_Display.pp_typ Pure.thy;  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
246  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
247  | 
PolyML.addPrettyPrinter stnd_pretty_printer*}  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
248  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
249  | 
text {*
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
250  | 
After that the types for booleans, lists and so on are printed out again  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
251  | 
the standard Isabelle way.  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
252  | 
|
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
253  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
254  | 
  "@{typ \"bool\"};
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
255  | 
@{typ \"'a list\"}"
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
256  | 
"\"bool\"  | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
257  | 
\"'a List.list\""}  | 
| 
318
 
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  | 
  \begin{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
260  | 
  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
 | 
261  | 
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
 | 
262  | 
  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
 | 
263  | 
  \end{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
264  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
265  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
266  | 
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
 | 
267  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
268  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
269  | 
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
 | 
270  | 
only construct fixed terms (remember they are ``linked'' at compile-time).  | 
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
271  | 
However, you often need to construct terms manually. For example, a  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
272  | 
  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
 | 
273  | 
  @{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
 | 
274  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
275  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
276  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
277  | 
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
 | 
278  | 
let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
279  | 
  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
 | 
280  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
281  | 
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
 | 
282  | 
end *}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
283  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
284  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
285  | 
  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
 | 
286  | 
  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
 | 
287  | 
  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
 | 
288  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
289  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
290  | 
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
 | 
291  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
292  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
293  | 
  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
 | 
294  | 
  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
 | 
295  | 
the given arguments  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
296  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
297  | 
  @{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
 | 
298  | 
"Const \<dots> $  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
299  | 
Abs (\"x\", Type (\"nat\",[]),  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
300  | 
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
 | 
301  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
302  | 
  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
 | 
303  | 
  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
 | 
304  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
305  | 
  @{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
 | 
306  | 
"Const \<dots> $  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
307  | 
Abs (\"x\", \<dots>,  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
308  | 
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
 | 
309  | 
(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
 | 
310  | 
|
| 345 | 311  | 
There are a number of handy functions that are frequently used for  | 
312  | 
  constructing terms. One is the function @{ML_ind list_comb in Term}, which
 | 
|
| 350 | 313  | 
takes as argument a term and a list of terms, and produces as output the  | 
| 345 | 314  | 
term list applied to the term. For example  | 
315  | 
||
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
316  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
317  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
318  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
319  | 
  val trm = @{term \"P::nat\"}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
320  | 
  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
 | 
321  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
322  | 
list_comb (trm, args)  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
323  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
324  | 
"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
 | 
325  | 
|
| 345 | 326  | 
  Another handy function is @{ML_ind lambda in Term}, which abstracts a variable 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
327  | 
in a term. For example  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
328  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
329  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
330  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
331  | 
  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
 | 
332  | 
  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
 | 
333  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
334  | 
lambda x_nat trm  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
335  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
336  | 
"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
 | 
337  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
338  | 
  In this example, @{ML lambda} produces a de Bruijn index (i.e.~@{ML "Bound 0"}), 
 | 
| 350 | 339  | 
and an abstraction, where it also records the type of the abstracted  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
340  | 
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
 | 
341  | 
  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
 | 
342  | 
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
 | 
343  | 
as in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
344  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
345  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
346  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
347  | 
  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
 | 
348  | 
  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
 | 
349  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
350  | 
lambda x_int trm  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
351  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
352  | 
"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
 | 
353  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
354  | 
  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
 | 
355  | 
This is a fundamental principle  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
356  | 
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
 | 
357  | 
have different type.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
358  | 
|
| 345 | 359  | 
  There is also the function @{ML_ind subst_free in Term} with which terms can be
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
360  | 
  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
 | 
361  | 
  "(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
 | 
362  | 
  @{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
 | 
363  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
364  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
365  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
366  | 
  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
 | 
367  | 
  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
 | 
368  | 
  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
 | 
369  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
370  | 
subst_free [sub1, sub2] trm  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
371  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
372  | 
"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
 | 
373  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
374  | 
  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
 | 
375  | 
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
 | 
376  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
377  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
378  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
379  | 
  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
 | 
380  | 
  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
 | 
381  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
382  | 
subst_free [sub] trm  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
383  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
384  | 
"Free (\"x\", \"nat\")"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
385  | 
|
| 345 | 386  | 
  Similarly the function @{ML_ind subst_bounds in Term}, replaces lose bound 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
387  | 
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
 | 
388  | 
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
 | 
389  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
390  | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
391  | 
ML{*fun strip_alls t =
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
392  | 
let  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
393  | 
fun aux (x, T, t) = strip_alls t |>> cons (Free (x, T))  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
394  | 
in  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
395  | 
case t of  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
396  | 
    Const ("All", _) $ Abs body => aux body
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
397  | 
| _ => ([], t)  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
398  | 
end*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
399  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
400  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
401  | 
The function returns a pair consisting of the stripped off variables and  | 
| 350 | 402  | 
the body of the universal quantification. For example  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
403  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
404  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
405  | 
  "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
 | 
406  | 
"([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
 | 
407  | 
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
 | 
408  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
409  | 
  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
 | 
410  | 
  the function @{ML subst_bounds}, you can replace these lose @{ML_ind 
 | 
| 345 | 411  | 
Bound in Term}s with the stripped off variables.  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
412  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
413  | 
  @{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
 | 
414  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
415  | 
  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
 | 
416  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
417  | 
subst_bounds (rev vrs, trm)  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
418  | 
  |> string_of_term @{context}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
419  | 
|> tracing  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
420  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
421  | 
"x = y"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
422  | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
423  | 
  Note that in Line 4 we had to reverse the list of variables that @{ML
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
424  | 
strip_alls} returned. The reason is that the head of the list the function  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
425  | 
  @{ML subst_bounds} takes is the replacement for @{ML "Bound 0"}, the next
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
426  | 
  element for @{ML "Bound 1"} and so on. 
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
427  | 
|
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
428  | 
Notice also that this function might introduce name clashes, since we  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
429  | 
substitute just a variable with the name recorded in an abstraction. This  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
430  | 
name is by no means unique. If clashes need to be avoided, then we should  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
431  | 
  use the function @{ML_ind dest_abs in Term}, which returns the body where
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
432  | 
the lose de Bruijn index is replaced by a unique free variable. For example  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
433  | 
|
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
434  | 
|
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
435  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
436  | 
"let  | 
| 374 | 437  | 
  val body = Bound 0 $ Free (\"x\", @{typ nat})
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
438  | 
in  | 
| 374 | 439  | 
  Term.dest_abs (\"x\", @{typ \"nat \<Rightarrow> bool\"}, body)
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
440  | 
end"  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
441  | 
"(\"xa\", Free (\"xa\", \"nat \<Rightarrow> bool\") $ Free (\"x\", \"nat\"))"}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
442  | 
|
| 350 | 443  | 
There are also many convenient functions that construct specific HOL-terms  | 
444  | 
  in the structure @{ML_struct HOLogic}. For
 | 
|
| 345 | 445  | 
  example @{ML_ind mk_eq in HOLogic} constructs an equality out of two terms.
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
446  | 
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
 | 
447  | 
arguments. For example  | 
| 
 
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  | 
@{ML_response_fake [gray,display]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
450  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
451  | 
  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
 | 
452  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
453  | 
  string_of_term @{context} eq
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
454  | 
|> tracing  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
455  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
456  | 
"True = False"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
457  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
458  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
459  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
460  | 
  \begin{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
461  | 
  There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file
 | 
| 374 | 462  | 
  "Pure/logic.ML"} and @{ML_file "HOL/Tools/hologic.ML"} that make manual
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
463  | 
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
 | 
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  | 
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
 | 
467  | 
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
 | 
468  | 
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
 | 
469  | 
  @{text is_true} 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  | 
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
 | 
473  | 
| is_true _ = false*}  | 
| 
 
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 {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
476  | 
  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
 | 
477  | 
the function  | 
| 
 
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  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
480  | 
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
 | 
481  | 
| is_all _ = false*}  | 
| 
 
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  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
484  | 
  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
 | 
485  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
486  | 
  @{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
 | 
487  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
488  | 
  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
 | 
489  | 
  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
 | 
490  | 
  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
 | 
491  | 
for this function is  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
492  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
493  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
494  | 
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
 | 
495  | 
| is_all _ = false*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
496  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
497  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
498  | 
because now  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
499  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
500  | 
@{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
 | 
501  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
502  | 
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
 | 
503  | 
second any term).  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
504  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
505  | 
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
 | 
506  | 
  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
 | 
507  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
508  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
509  | 
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
 | 
510  | 
| is_nil _ = false *}  | 
| 
 
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 {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
513  | 
  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
 | 
514  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
515  | 
  @{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
 | 
516  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
517  | 
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
 | 
518  | 
  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
 | 
519  | 
  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
 | 
520  | 
  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
 | 
521  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
522  | 
  @{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
 | 
523  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
524  | 
  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
 | 
525  | 
  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
 | 
526  | 
  (@{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
 | 
527  | 
  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
 | 
528  | 
  \mbox{@{text "(op *)"}}:
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
529  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
530  | 
  @{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
 | 
531  | 
"(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
 | 
532  | 
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
 | 
533  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
534  | 
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
 | 
535  | 
  @{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
 | 
536  | 
  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
 | 
537  | 
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
 | 
538  | 
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
 | 
539  | 
  @{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
 | 
540  | 
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
 | 
541  | 
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
 | 
542  | 
be written as follows.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
543  | 
*}  | 
| 
 
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  | 
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
 | 
546  | 
  | 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
 | 
547  | 
| is_nil_or_all _ = false *}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
549  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
550  | 
  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
 | 
551  | 
For example  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
552  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
553  | 
  @{ML_response [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
554  | 
  "@{type_name \"list\"}" "\"List.list\""}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
555  | 
|
| 329 | 556  | 
  \footnote{\bf FIXME: Explain the following better; maybe put in a separate
 | 
557  | 
section and link with the comment in the antiquotation section.}  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
558  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
559  | 
Occasionally you have to calculate what the ``base'' name of a given  | 
| 345 | 560  | 
  constant is. For this you can use the function @{ML_ind "Sign.extern_const"} or
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
561  | 
  @{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
 | 
562  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
563  | 
  @{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
 | 
564  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
565  | 
  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
 | 
566  | 
  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
 | 
567  | 
strips off all qualifiers.  | 
| 
 
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  | 
  \begin{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
570  | 
  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
 | 
571  | 
  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
 | 
572  | 
  \end{readmore}
 | 
| 
 
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  | 
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
 | 
575  | 
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
 | 
576  | 
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
 | 
577  | 
type is as follows:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
578  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
579  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
580  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
581  | 
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
 | 
582  | 
|
| 345 | 583  | 
text {* This can be equally written with the combinator @{ML_ind "-->" in Term} as: *}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
584  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
585  | 
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
 | 
586  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
587  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
588  | 
If you want to construct a function type with more than one argument  | 
| 345 | 589  | 
  type, then you can use @{ML_ind "--->" in Term}.
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
590  | 
*}  | 
| 
 
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{*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
 | 
593  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
594  | 
text {*
 | 
| 369 | 595  | 
  A handy function for manipulating terms is @{ML_ind map_types in Term}: it takes a 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
596  | 
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
 | 
597  | 
  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
 | 
598  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
599  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
600  | 
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
 | 
601  | 
(case ty of  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
602  | 
     @{typ nat} => @{typ int}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
603  | 
| 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
 | 
604  | 
| _ => ty)*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
605  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
606  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
607  | 
Here is an example:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
608  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
609  | 
@{ML_response_fake [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
610  | 
"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
 | 
611  | 
"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
 | 
612  | 
$ 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
 | 
613  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
614  | 
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
 | 
615  | 
  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
 | 
616  | 
  (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
 | 
617  | 
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
 | 
618  | 
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
 | 
619  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
620  | 
  @{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
 | 
621  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
622  | 
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
 | 
623  | 
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
 | 
624  | 
call them as follows  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
625  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
626  | 
  @{ML_response [gray,display]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
627  | 
  "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
 | 
628  | 
"[(\"'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
 | 
629  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
630  | 
  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
 | 
631  | 
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
 | 
632  | 
  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
 | 
633  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
634  | 
  \begin{exercise}\label{fun:revsum}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
635  | 
  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
 | 
636  | 
  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
 | 
637  | 
  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
 | 
638  | 
  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
 | 
639  | 
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
 | 
640  | 
  \end{exercise}
 | 
| 
 
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  | 
  \begin{exercise}\label{fun:makesum}
 | 
| 350 | 643  | 
Write a function that takes two terms representing natural numbers  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
644  | 
  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
 | 
645  | 
number representing their sum.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
646  | 
  \end{exercise}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
647  | 
|
| 329 | 648  | 
  \begin{exercise}\label{ex:debruijn}
 | 
| 350 | 649  | 
Implement the function, which we below name deBruijn, that depends on a natural  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
650  | 
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
 | 
651  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
652  | 
  \begin{center}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
653  | 
  \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
 | 
654  | 
  {\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
 | 
655  | 
  {\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
 | 
656  | 
                        $\longrightarrow$ {\it rhs n}\\
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
657  | 
  {\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
 | 
658  | 
  \end{tabular}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
659  | 
  \end{center}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
660  | 
|
| 329 | 661  | 
This function returns for n=3 the term  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
662  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
663  | 
  \begin{center}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
664  | 
  \begin{tabular}{l}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
665  | 
(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
 | 
666  | 
(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
 | 
667  | 
(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
 | 
668  | 
  \end{tabular}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
669  | 
  \end{center}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
670  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
671  | 
  Make sure you use the functions defined in @{ML_file "HOL/Tools/hologic.ML"}
 | 
| 350 | 672  | 
  for constructing the terms for the logical connectives.\footnote{Thanks to Roy
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
673  | 
Dyckhoff for suggesting this exercise and working out the details.}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
674  | 
  \end{exercise}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
675  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
676  | 
|
| 383 | 677  | 
section {* Unification and Matching *}
 | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
678  | 
|
| 
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
679  | 
text {* 
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
680  | 
Isabelle's terms and types may contain schematic term variables  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
681  | 
  (term-constructor @{ML Var}) and schematic type variables (term-constructor
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
682  | 
  @{ML TVar}). These variables stand for unknown entities, which can be made
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
683  | 
more concrete by instantiations. Such instantiations might be a result of  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
684  | 
unification or matching. While in case of types, unification and matching is  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
685  | 
relatively straightforward, in case of terms the algorithms are  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
686  | 
substantially more complicated, because terms need higher-order versions of  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
687  | 
the unification and matching algorithms. Below we shall use the  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
688  | 
  antiquotations @{text "@{typ_pat \<dots>}"} and @{text "@{term_pat \<dots>}"} from
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
689  | 
  Section~\ref{sec:antiquote} in order to construct examples involving
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
690  | 
schematic variables.  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
691  | 
|
| 382 | 692  | 
Let us begin with describing the unification and matching function for  | 
| 383 | 693  | 
  types.  Both return type environments (ML-type @{ML_type "Type.tyenv"})
 | 
694  | 
which map schematic type variables to types and sorts. Below we use the  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
695  | 
  function @{ML_ind typ_unify in Sign} from the structure @{ML_struct Sign}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
696  | 
  for unifying the types @{text "?'a * ?'b"} and @{text "?'b list *
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
697  | 
  nat"}. This will produce the mapping, or type environment, @{text "[?'a :=
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
698  | 
?'b list, ?'b := nat]"}.  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
699  | 
*}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
700  | 
|
| 382 | 701  | 
ML %linenosgray{*val (tyenv_unif, _) = let
 | 
| 379 | 702  | 
  val ty1 = @{typ_pat "?'a * ?'b"}
 | 
703  | 
  val ty2 = @{typ_pat "?'b list * nat"}
 | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
704  | 
in  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
705  | 
  Sign.typ_unify @{theory} (ty1, ty2) (Vartab.empty, 0) 
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
706  | 
end*}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
707  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
708  | 
text {* 
 | 
| 383 | 709  | 
  The environment @{ML_ind "Vartab.empty"} in line 5 stands for the empty type
 | 
710  | 
environment, which is needed for starting the unification without any  | 
|
711  | 
  (pre)instantiations. The @{text 0} is an integer index that we will explain
 | 
|
712  | 
  below. In case of failure @{ML typ_unify in Sign} will throw the exception
 | 
|
713  | 
  @{text TUNIFY}.  We can print out the resulting type environment @{ML
 | 
|
714  | 
  tyenv_unif} with the built-in function @{ML_ind dest in Vartab} from the
 | 
|
715  | 
  structure @{ML_struct Vartab}.
 | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
716  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
717  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
718  | 
"Vartab.dest tyenv_unif"  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
719  | 
"[((\"'a\", 0), ([\"HOL.type\"], \"?'b List.list\")),  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
720  | 
((\"'b\", 0), ([\"HOL.type\"], \"nat\"))]"}  | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
721  | 
*}  | 
| 
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
722  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
723  | 
text_raw {*
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
724  | 
  \begin{figure}[t]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
725  | 
  \begin{minipage}{\textwidth}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
726  | 
  \begin{isabelle}*}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
727  | 
ML{*fun pretty_helper f env =
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
728  | 
env |> Vartab.dest  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
729  | 
|> map f  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
730  | 
|> commas  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
731  | 
|> enclose "[" "]"  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
732  | 
|> tracing  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
733  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
734  | 
fun pretty_tyenv ctxt tyenv =  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
735  | 
let  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
736  | 
fun aux (v, (s, T)) =  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
737  | 
Syntax.string_of_typ ctxt (TVar (v, s))  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
738  | 
^ " := " ^ Syntax.string_of_typ ctxt T  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
739  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
740  | 
pretty_helper aux tyenv  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
741  | 
end*}  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
742  | 
text_raw {*
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
743  | 
  \end{isabelle}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
744  | 
  \end{minipage}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
745  | 
  \caption{A pretty printing function for type environments, which are 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
746  | 
  produced by unification and matching.\label{fig:prettyenv}}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
747  | 
  \end{figure}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
748  | 
*}  | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
749  | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
750  | 
text {*
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
751  | 
The first components in this list stand for the schematic type variables and  | 
| 383 | 752  | 
the second are the associated sorts and types. In this example the sort is  | 
753  | 
  the default sort @{text "HOL.type"}. We will use in what follows the
 | 
|
754  | 
  pretty-printing function from Figure~\ref{fig:prettyenv} for @{ML_type
 | 
|
755  | 
"Type.tyenv"}s. For the type environment in the example this function prints  | 
|
756  | 
out the more legible:  | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
757  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
758  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
759  | 
  "pretty_tyenv @{context} tyenv_unif"
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
760  | 
"[?'a := ?'b list, ?'b := nat]"}  | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
761  | 
|
| 383 | 762  | 
  The way the unification function @{ML typ_unify in Sign} is implemented 
 | 
763  | 
using an initial type environment and initial index makes it easy to  | 
|
764  | 
unify more than two terms. For example  | 
|
765  | 
*}  | 
|
766  | 
||
767  | 
ML %linenosgray{*val (tyenvs, _) = let
 | 
|
768  | 
  val tys1 = (@{typ_pat "?'a"}, @{typ_pat "?'b list"})
 | 
|
769  | 
  val tys2 = (@{typ_pat "?'b"}, @{typ_pat "nat"})
 | 
|
770  | 
in  | 
|
771  | 
  fold (Sign.typ_unify @{theory}) [tys1, tys2] (Vartab.empty, 0) 
 | 
|
772  | 
end*}  | 
|
773  | 
||
774  | 
text {*
 | 
|
775  | 
  The index @{text 0} in Line 5 is the maximal index of the schematic type
 | 
|
776  | 
  variables occurring in @{text ty1} and @{text ty2}. This index will be
 | 
|
777  | 
increased whenever a new schematic type variable is introduced during  | 
|
778  | 
unification. This is for example the case when two schematic type variables  | 
|
779  | 
have different, incomparable sorts. Then a new schematic type variable is  | 
|
780  | 
introduced with the combined sorts. To show this let us assume two sorts,  | 
|
781  | 
  say @{text "s1"} and @{text "s2"}, which we attach to the schematic type
 | 
|
782  | 
  variables @{text "?'a"} and @{text "?'b"}. Since we do not make any
 | 
|
783  | 
assumption about the sorts, they are incomparable.  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
784  | 
*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
785  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
786  | 
ML{*val (tyenv, index) = let
 | 
| 383 | 787  | 
  val ty1 = @{typ_pat "?'a::s1"}
 | 
788  | 
  val ty2 = @{typ_pat "?'b::s2"}
 | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
789  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
790  | 
  Sign.typ_unify @{theory} (ty1, ty2) (Vartab.empty, 0) 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
791  | 
end*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
792  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
793  | 
text {*
 | 
| 383 | 794  | 
To print out the result type environment we switch on the printing  | 
795  | 
  of sort information by setting @{ML_ind show_sorts in Syntax} to 
 | 
|
796  | 
true. This allows us to inspect the typing environment.  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
797  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
798  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
799  | 
  "pretty_tyenv @{context} tyenv"
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
800  | 
  "[?'a::s1 := ?'a1::{s1,s2}, ?'b::s2 := ?'a1::{s1,s2}]"}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
801  | 
|
| 383 | 802  | 
  As can be seen, the type variables @{text "?'a"} and @{text "?'b"} are instantiated
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
803  | 
  with a new type variable @{text "?'a1"} with sort @{text "{s1,s2}"}. Since a new
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
804  | 
  type variable has been introduced the @{ML index}, originally being @{text 0}, 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
805  | 
  has been increased by @{text 1}.
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
806  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
807  | 
  @{ML_response [display,gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
808  | 
"index"  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
809  | 
"1"}  | 
| 383 | 810  | 
|
811  | 
  Let us now return again to the unification problem @{text "?'a * ?'b"} and 
 | 
|
812  | 
  @{text "?'b list * nat"} from the beginning of the section, and the 
 | 
|
813  | 
  calculated type environment @{ML tyenv_unif}:
 | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
814  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
815  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
816  | 
  "pretty_tyenv @{context} tyenv_unif"
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
817  | 
"[?'a := ?'b list, ?'b := nat]"}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
818  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
819  | 
  Observe that the type environment which the function @{ML typ_unify in
 | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
820  | 
  Sign} returns is in \emph{not} an instantiation in fully solved form: while @{text
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
821  | 
  "?'b"} is instantiated to @{typ nat}, this is not propagated to the
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
822  | 
  instantiation for @{text "?'a"}.  In unification theory, this is often
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
823  | 
  called an instantiation in \emph{triangular form}. These not fully solved 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
824  | 
instantiations are used because of performance reasons.  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
825  | 
  To apply a type environment in triangular form to a type, say @{text "?'a *
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
826  | 
  ?'b"}, you should use the function @{ML_ind norm_type in Envir}:
 | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
827  | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
828  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
829  | 
  "Envir.norm_type tyenv_unif @{typ_pat \"?'a * ?'b\"}"
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
830  | 
"nat list * nat"}  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
831  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
832  | 
  With this function the type variables @{text "?'a"} and @{text "?'b"} are 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
833  | 
correctly instantiated.  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
834  | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
835  | 
  Matching of types can be done with the function @{ML_ind typ_match in Sign}
 | 
| 383 | 836  | 
  also from the structure @{ML_struct Sign}. This function returns a @{ML_type
 | 
837  | 
  Type.tyenv} as well, but might raise the exception @{text TYPE_MATCH} in case
 | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
838  | 
of failure. For example  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
839  | 
*}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
840  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
841  | 
ML{*val tyenv_match = let
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
842  | 
  val pat = @{typ_pat "?'a * ?'b"}
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
843  | 
  and ty  = @{typ_pat "bool list * nat"}
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
844  | 
in  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
845  | 
  Sign.typ_match @{theory} (pat, ty) Vartab.empty 
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
846  | 
end*}  | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
847  | 
|
| 
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
848  | 
text {*
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
849  | 
Printing out the calculated matcher gives  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
850  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
851  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
852  | 
  "pretty_tyenv @{context} tyenv_match"
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
853  | 
"[?'a := bool list, ?'b := nat]"}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
854  | 
|
| 383 | 855  | 
  Unlike unification, which uses the function @{ML norm_type in Envir}, 
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
856  | 
applying the matcher to a type needs to be done with the function  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
857  | 
  @{ML_ind subst_type in Envir}.
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
858  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
859  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
860  | 
  "Envir.subst_type tyenv_match @{typ_pat \"?'a * ?'b\"}"
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
861  | 
"bool list * nat"}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
862  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
863  | 
Be careful to observe the difference: use always  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
864  | 
  @{ML subst_type in Envir} for matchers and @{ML norm_type in Envir} 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
865  | 
for unifiers. To make the difference explicit, let us calculate the  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
866  | 
following matcher:  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
867  | 
*}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
868  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
869  | 
ML{*val tyenv_match' = let
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
870  | 
  val pat = @{typ_pat "?'a * ?'b"}
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
871  | 
  and ty  = @{typ_pat "?'b list * nat"}
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
872  | 
in  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
873  | 
  Sign.typ_match @{theory} (pat, ty) Vartab.empty 
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
874  | 
end*}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
875  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
876  | 
text {*
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
877  | 
  Now @{ML tyenv_unif} is equal to @{ML tyenv_match'}. If we apply 
 | 
| 383 | 878  | 
  @{ML norm_type in Envir} to the type @{text "?'a * ?'b"} we obtain
 | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
879  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
880  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
881  | 
  "Envir.norm_type tyenv_match' @{typ_pat \"?'a * ?'b\"}"
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
882  | 
"nat list * nat"}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
883  | 
|
| 383 | 884  | 
which does not solve the matching problem, and if  | 
885  | 
  we apply @{ML subst_type in Envir} to the same type we obtain
 | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
886  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
887  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
888  | 
  "Envir.subst_type tyenv_unif @{typ_pat \"?'a * ?'b\"}"
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
889  | 
"?'b list * nat"}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
890  | 
|
| 383 | 891  | 
which does not solve the unification problem.  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
892  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
893  | 
  \begin{readmore}
 | 
| 383 | 894  | 
Unification and matching for types is implemented  | 
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
895  | 
  in @{ML_file "Pure/type.ML"}. The ``interface'' functions for them
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
896  | 
  are in @{ML_file "Pure/sign.ML"}. Matching and unification produce type environments
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
897  | 
  as results. These are implemented in @{ML_file "Pure/envir.ML"}.
 | 
| 379 | 898  | 
This file also includes the substitution and normalisation functions,  | 
899  | 
that apply a type environment to a type. Type environments are lookup  | 
|
900  | 
  tables which are implemented in @{ML_file "Pure/term_ord.ML"}.
 | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
901  | 
  \end{readmore}
 | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
902  | 
*}  | 
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
903  | 
|
| 
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
904  | 
text {*
 | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
905  | 
Unification and matching of terms is substantially more complicated than the  | 
| 383 | 906  | 
type-case. The reason is that terms have abstractions and, in this context,  | 
907  | 
unification or matching modulo plain equality is often not meaningful.  | 
|
908  | 
  Nevertheless, Isabelle implements the function @{ML_ind
 | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
909  | 
first_order_match in Pattern} for terms. This matching function returns a  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
910  | 
type environment and a term environment. To pretty print the latter we use  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
911  | 
  the function @{text "pretty_env"}:
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
912  | 
*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
913  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
914  | 
ML{*fun pretty_env ctxt env =
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
915  | 
let  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
916  | 
fun aux (v, (T, t)) =  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
917  | 
string_of_term ctxt (Var (v, T)) ^ " := " ^ string_of_term ctxt t  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
918  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
919  | 
pretty_helper aux env  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
920  | 
end*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
921  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
922  | 
text {*
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
923  | 
  As can be seen from the @{text "aux"}-function, a term environment associates 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
924  | 
a schematic term variable with a type and a term. An example of a first-order  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
925  | 
  matching problem is the term @{term "P (\<lambda>a b. Q b a)"} and the pattern 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
926  | 
  @{text "?X ?Y"}.
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
927  | 
*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
928  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
929  | 
ML{*val (_, fo_env) = let
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
930  | 
  val fo_pat = @{term_pat "(?X::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool) ?Y"}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
931  | 
  val trm = @{term "P::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool"} 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
932  | 
              $ @{term "\<lambda>a b. (Q::nat\<Rightarrow>nat\<Rightarrow>nat) b a"}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
933  | 
val init = (Vartab.empty, Vartab.empty)  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
934  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
935  | 
  Pattern.first_order_match @{theory} (fo_pat, trm) init
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
936  | 
end *}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
937  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
938  | 
text {*
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
939  | 
In this example we annotated explicitly types because then  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
940  | 
the type environment is empty and can be ignored. The  | 
| 383 | 941  | 
resulting term environment is  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
942  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
943  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
944  | 
  "pretty_env @{context} fo_env"
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
945  | 
"[?X := P, ?Y := \<lambda>a b. Q b a]"}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
946  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
947  | 
  The matcher can be applied to a term using the function @{ML_ind subst_term
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
948  | 
  in Envir} (remember the same convention for types applies to terms: @{ML
 | 
| 383 | 949  | 
  subst_term in Envir} is for matchers and @{ML norm_term in Envir} for
 | 
950  | 
  unifiers). The function @{ML subst_term in Envir} expects a type environment,
 | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
951  | 
which is set to empty in the example below, and a term environment.  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
952  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
953  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
954  | 
"let  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
955  | 
  val trm = @{term_pat \"(?X::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool) ?Y\"}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
956  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
957  | 
Envir.subst_term (Vartab.empty, fo_env) trm  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
958  | 
  |> string_of_term @{context}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
959  | 
|> tracing  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
960  | 
end"  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
961  | 
"P (\<lambda>a b. Q b a)"}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
962  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
963  | 
First-order matching is useful for matching against applications and  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
964  | 
variables. It can deal also with abstractions and a limited form of  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
965  | 
alpha-equivalence, but this kind of matching should be used with care, since  | 
| 383 | 966  | 
it is not clear whether the result is meaningful. A meaningful example is  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
967  | 
  matching @{text "\<lambda>x. P x"} against the pattern @{text "\<lambda>y. ?X y"}. In this
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
968  | 
  case first-order matching produces @{text "[?X := P]"}.
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
969  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
970  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
971  | 
"let  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
972  | 
  val fo_pat = @{term_pat \"\<lambda>y. (?X::nat\<Rightarrow>bool) y\"}
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
973  | 
  val trm = @{term \"\<lambda>x. (P::nat\<Rightarrow>bool) x\"} 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
974  | 
val init = (Vartab.empty, Vartab.empty)  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
975  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
976  | 
  Pattern.first_order_match @{theory} (fo_pat, trm) init
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
977  | 
|> snd  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
978  | 
  |> pretty_env @{context} 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
979  | 
end"  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
980  | 
"[?X := P]"}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
981  | 
*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
982  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
983  | 
text {*
 | 
| 383 | 984  | 
Unification of abstractions is more thoroughly studied in the context  | 
985  | 
of higher-order pattern unification and higher-order pattern matching. A  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
986  | 
  \emph{\index*{pattern}} is a lambda term whose ``head symbol'' (that is the
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
987  | 
first symbol under abstractions) is either a constant, or a schematic or a free  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
988  | 
variable. If it is a schematic variable then it can be only applied by  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
989  | 
distinct bound variables. This excludes that a schematic variable is an  | 
| 383 | 990  | 
argument of another one and that a schematic variable is applied  | 
991  | 
  twice with the same bound variable. The function @{ML_ind pattern in Pattern}
 | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
992  | 
  in the structure @{ML_struct Pattern} tests whether a term satisfies these
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
993  | 
restrictions.  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
994  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
995  | 
  @{ML_response [display, gray]
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
996  | 
"let  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
997  | 
val trm_list =  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
998  | 
        [@{term_pat \"?X\"},              @{term_pat \"a\"},
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
999  | 
         @{term_pat \"\<lambda>a b. ?X a b\"},    @{term_pat \"\<lambda>a b. (op +) a b\"},
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1000  | 
         @{term_pat \"\<lambda>a. (op +) a ?Y\"}, @{term_pat \"?X ?Y\"},
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1001  | 
         @{term_pat \"\<lambda>a b. ?X a b ?Y\"}, @{term_pat \"\<lambda>a. ?X a a\"}] 
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1002  | 
in  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1003  | 
map Pattern.pattern trm_list  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1004  | 
end"  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1005  | 
"[true, true, true, true, true, false, false, false]"}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1006  | 
|
| 383 | 1007  | 
The point of the restriction to patterns is that unification and matching  | 
1008  | 
are decidable and produce most general unifiers, respectively matchers.  | 
|
1009  | 
In this way, matching and unification can be implemented as functions that  | 
|
1010  | 
produce a type and term environment (unification actually returns a  | 
|
1011  | 
  record of type @{ML_type Envir.env} containing a maxind, a type environment 
 | 
|
1012  | 
  and a term environment). The former function is @{ML_ind match in Pattern},
 | 
|
1013  | 
  the latter @{ML_ind unify in Pattern} both implemented in the structure
 | 
|
1014  | 
  @{ML_struct Pattern}. An example for higher-order pattern unification is
 | 
|
1015  | 
||
1016  | 
\ldots  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1017  | 
*}  | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1018  | 
|
| 383 | 1019  | 
ML{*let
 | 
1020  | 
  val pat = @{term_pat "?X"}
 | 
|
1021  | 
  val trm = @{term "a"}
 | 
|
1022  | 
val init = (Vartab.empty, Vartab.empty)  | 
|
1023  | 
in  | 
|
1024  | 
  Pattern.match @{theory} (pat, trm) init
 | 
|
1025  | 
end *}  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1026  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1027  | 
text {*
 | 
| 383 | 1028  | 
An assumption of this function is that the terms to be unified have  | 
1029  | 
already the same type. In case of failure, the exceptions that are raised  | 
|
1030  | 
  are either @{text Pattern}, @{text MATCH} or @{text Unif}.
 | 
|
1031  | 
||
1032  | 
As mentioned before, ``full'' higher-order unification, respectively, matching  | 
|
1033  | 
problems are in general undecidable and might in general not posses a single  | 
|
1034  | 
most general solution. Therefore Isabelle implements Huet's pre-unification  | 
|
1035  | 
algorithm which does not return a single result, rather a lazy list of potentially  | 
|
1036  | 
  infinite unifiers. The corresponding function is called @{ML_ind unifiers in Unify}.
 | 
|
1037  | 
An example is as follows  | 
|
1038  | 
||
1039  | 
\ldots  | 
|
1040  | 
||
1041  | 
In case of failure this function does not raise an exception, rather returns  | 
|
1042  | 
the empty sequence. In order to find a reasonable solution for a unification  | 
|
1043  | 
problem, Isabelle also tries first to solve the problem by higher-order  | 
|
1044  | 
pattern unification.  | 
|
1045  | 
||
1046  | 
\ldots  | 
|
1047  | 
||
1048  | 
  For higher-order matching the function is called @{ML_ind matchers in Unify}.
 | 
|
1049  | 
Also this function might potentially return sequences with more than one  | 
|
1050  | 
element.  | 
|
1051  | 
||
1052  | 
||
1053  | 
  Like @{ML unifiers in Unify}, this function does not raise an exception
 | 
|
1054  | 
in case of failure. It also tries out first whether the matching problem  | 
|
1055  | 
can be solved by first-order matching.  | 
|
1056  | 
||
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
1057  | 
  \begin{readmore}
 | 
| 383 | 1058  | 
Unification and matching of higher-order patterns is implemented in  | 
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1059  | 
  @{ML_file "Pure/pattern.ML"}. This file also contains a first-order matcher
 | 
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1060  | 
for terms. Full higher-order unification is implemented  | 
| 383 | 1061  | 
  in @{ML_file "Pure/unify.ML"}. It uses lazy sequences which are implemented
 | 
1062  | 
  in @{ML_file "Pure/General/seq.ML"}.
 | 
|
| 
378
 
8d160d79b48c
section about matching and unification of types
 
Christian Urban <urbanc@in.tum.de> 
parents: 
377 
diff
changeset
 | 
1063  | 
  \end{readmore}
 | 
| 
377
 
272ba2cceeb2
added a section about unification and matching
 
Christian Urban <urbanc@in.tum.de> 
parents: 
375 
diff
changeset
 | 
1064  | 
*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1065  | 
|
| 
381
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1066  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1067  | 
|
| 
 
97518188ef0e
added more to the unification section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
380 
diff
changeset
 | 
1068  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1069  | 
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
 | 
1070  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1071  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1072  | 
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
 | 
1073  | 
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
 | 
1074  | 
information) so that it is always clear what the type of a term is.  | 
| 369 | 1075  | 
  Given a well-typed term, the function @{ML_ind type_of in Term} returns the 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1076  | 
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
 | 
1077  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1078  | 
  @{ML_response [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1079  | 
  "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
 | 
1080  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1081  | 
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
 | 
1082  | 
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
 | 
1083  | 
  @{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
 | 
1084  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1085  | 
  @{ML_response_fake [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1086  | 
  "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
 | 
1087  | 
"*** 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
 | 
1088  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1089  | 
Since the complete traversal might sometimes be too costly and  | 
| 369 | 1090  | 
  not necessary, there is the function @{ML_ind fastype_of in Term}, which 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1091  | 
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
 | 
1092  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1093  | 
  @{ML_response [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1094  | 
  "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
 | 
1095  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1096  | 
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
 | 
1097  | 
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
 | 
1098  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1099  | 
   @{ML_response [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1100  | 
  "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
 | 
1101  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1102  | 
where no error is detected.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1103  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1104  | 
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
 | 
1105  | 
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
 | 
1106  | 
information is redundant. A short-cut is to use the ``place-holder''  | 
| 345 | 1107  | 
  type @{ML_ind dummyT in Term} and then let type-inference figure out the 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1108  | 
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
 | 
1109  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1110  | 
  @{ML_response_fake [display,gray] 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1111  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1112  | 
  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
 | 
1113  | 
  val o = @{term \"1::nat\"} 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1114  | 
val v = Free (\"x\", dummyT)  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1115  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1116  | 
  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
 | 
1117  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1118  | 
"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
 | 
1119  | 
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
 | 
1120  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1121  | 
  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
 | 
1122  | 
  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
 | 
1123  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1124  | 
  \begin{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1125  | 
  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
 | 
1126  | 
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
 | 
1127  | 
  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
 | 
1128  | 
  @{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
 | 
1129  | 
  \end{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1130  | 
|
| 329 | 1131  | 
  \footnote{\bf FIXME: say something about sorts.}
 | 
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1132  | 
  \footnote{\bf FIXME: give a ``readmore''.}
 | 
| 
318
 
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  | 
  \begin{exercise}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1135  | 
  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
 | 
1136  | 
result that type-checks. See what happens to the solutions of this  | 
| 329 | 1137  | 
  exercise given in Appendix \ref{ch:solutions} when they receive an 
 | 
1138  | 
ill-typed term as input.  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1139  | 
  \end{exercise}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1140  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1141  | 
|
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1142  | 
section {* Certified Terms and Certified Types *}
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1143  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1144  | 
text {*
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1145  | 
  You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1146  | 
typ}es, since they are just arbitrary unchecked trees. However, you  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1147  | 
eventually want to see if a term is well-formed, or type-checks, relative to  | 
| 369 | 1148  | 
  a theory.  Type-checking is done via the function @{ML_ind cterm_of in Thm}, which
 | 
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1149  | 
  converts a @{ML_type term} into a @{ML_type cterm}, a \emph{certified}
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1150  | 
  term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1151  | 
abstract objects that are guaranteed to be type-correct, and they can only  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1152  | 
be constructed via ``official interfaces''.  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1153  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1154  | 
Certification is always relative to a theory context. For example you can  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1155  | 
write:  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1156  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1157  | 
  @{ML_response_fake [display,gray] 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1158  | 
  "cterm_of @{theory} @{term \"(a::nat) + b = c\"}" 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1159  | 
"a + b = c"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1160  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1161  | 
This can also be written with an antiquotation:  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1162  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1163  | 
  @{ML_response_fake [display,gray] 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1164  | 
  "@{cterm \"(a::nat) + b = c\"}" 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1165  | 
"a + b = c"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1166  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1167  | 
Attempting to obtain the certified term for  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1168  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1169  | 
  @{ML_response_fake_both [display,gray] 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1170  | 
  "@{cterm \"1 + True\"}" 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1171  | 
"Type unification failed \<dots>"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1172  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1173  | 
yields an error (since the term is not typable). A slightly more elaborate  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1174  | 
example that type-checks is:  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1175  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1176  | 
@{ML_response_fake [display,gray] 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1177  | 
"let  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1178  | 
  val natT = @{typ \"nat\"}
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1179  | 
  val zero = @{term \"0::nat\"}
 | 
| 356 | 1180  | 
  val plus = Const (@{const_name plus}, [natT, natT] ---> natT)
 | 
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1181  | 
in  | 
| 356 | 1182  | 
  cterm_of @{theory} (plus $ zero $ zero)
 | 
1183  | 
end"  | 
|
1184  | 
"0 + 0"}  | 
|
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1185  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1186  | 
In Isabelle not just terms need to be certified, but also types. For example,  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1187  | 
  you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1188  | 
the ML-level as follows:  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1189  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1190  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1191  | 
  "ctyp_of @{theory} (@{typ nat} --> @{typ bool})"
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1192  | 
"nat \<Rightarrow> bool"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1193  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1194  | 
or with the antiquotation:  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1195  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1196  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1197  | 
  "@{ctyp \"nat \<Rightarrow> bool\"}"
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1198  | 
"nat \<Rightarrow> bool"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1199  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1200  | 
Since certified terms are, unlike terms, abstract objects, we cannot  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1201  | 
pattern-match against them. However, we can construct them. For example  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1202  | 
  the function @{ML_ind capply in Thm} produces a certified application.
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1203  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1204  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1205  | 
  "Thm.capply @{cterm \"P::nat \<Rightarrow> bool\"} @{cterm \"3::nat\"}"
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1206  | 
"P 3"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1207  | 
|
| 351 | 1208  | 
  Similarly the function @{ML_ind list_comb in Drule} from the structure @{ML_struct Drule}
 | 
1209  | 
  applies a list of @{ML_type cterm}s.
 | 
|
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1210  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1211  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1212  | 
"let  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1213  | 
  val chead = @{cterm \"P::unit \<Rightarrow> nat \<Rightarrow> bool\"}
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1214  | 
  val cargs = [@{cterm \"()\"}, @{cterm \"3::nat\"}]
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1215  | 
in  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1216  | 
Drule.list_comb (chead, cargs)  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1217  | 
end"  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1218  | 
"P () 3"}  | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1219  | 
|
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1220  | 
  \begin{readmore}
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1221  | 
  For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see 
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1222  | 
  the files @{ML_file "Pure/thm.ML"}, @{ML_file "Pure/more_thm.ML"} and
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1223  | 
  @{ML_file "Pure/drule.ML"}.
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1224  | 
  \end{readmore}
 | 
| 
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
1225  | 
*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1226  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1227  | 
section {* Theorems *}
 | 
| 
 
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  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1230  | 
  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
 | 
1231  | 
that can only be built by going through interfaces. As a consequence, every proof  | 
| 351 | 1232  | 
in Isabelle is correct by construction. This follows the tradition of the LCF approach.  | 
| 
318
 
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  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1235  | 
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
 | 
1236  | 
statement:  | 
| 
 
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  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1239  | 
lemma  | 
| 351 | 1240  | 
assumes assm\<^isub>1: "\<And>(x::nat). P x \<Longrightarrow> Q x"  | 
1241  | 
and assm\<^isub>2: "P t"  | 
|
1242  | 
shows "Q t"(*<*)oops(*>*)  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1243  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1244  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1245  | 
The corresponding ML-code is as follows:  | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1246  | 
*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1247  | 
|
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1248  | 
ML{*val my_thm = 
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1249  | 
let  | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1250  | 
  val assm1 = @{cprop "\<And>(x::nat). P x \<Longrightarrow> Q x"}
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1251  | 
  val assm2 = @{cprop "(P::nat \<Rightarrow> bool) t"}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1252  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1253  | 
val Pt_implies_Qt =  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1254  | 
assume assm1  | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1255  | 
        |> forall_elim @{cterm "t::nat"}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1256  | 
|
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1257  | 
val Qt = implies_elim Pt_implies_Qt (assume assm2)  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1258  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1259  | 
Qt  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1260  | 
|> implies_intr assm2  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1261  | 
|> implies_intr assm1  | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1262  | 
end*}  | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1263  | 
|
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1264  | 
text {* 
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1265  | 
  If we print out the value of @{ML my_thm} then we see only the 
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1266  | 
final statement of the theorem.  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1267  | 
|
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1268  | 
  @{ML_response_fake [display, gray]
 | 
| 348 | 1269  | 
  "tracing (string_of_thm @{context} my_thm)"
 | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1270  | 
"\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t"}  | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1271  | 
|
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1272  | 
However, internally the code-snippet constructs the following  | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1273  | 
proof.  | 
| 
318
 
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  | 
  \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
 | 
1277  | 
    {\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
 | 
1278  | 
       {\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
 | 
1279  | 
          {\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
 | 
1280  | 
                 {\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
 | 
1281  | 
&  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1282  | 
           \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
 | 
1283  | 
}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1284  | 
}  | 
| 
 
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  | 
\]  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1287  | 
|
| 
339
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1288  | 
While we obtained a theorem as result, this theorem is not  | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1289  | 
yet stored in Isabelle's theorem database. Consequently, it cannot be  | 
| 348 | 1290  | 
referenced on the user level. One way to store it in the theorem database is  | 
| 351 | 1291  | 
  by using the function @{ML_ind note in LocalTheory}.\footnote{\bf FIXME: make sure a pointer
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1292  | 
to the section about local-setup is given earlier.}  | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1293  | 
*}  | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1294  | 
|
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1295  | 
local_setup %gray {*
 | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1296  | 
LocalTheory.note Thm.theoremK  | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1297  | 
     ((@{binding "my_thm"}, []), [my_thm]) #> snd *}
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1298  | 
|
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1299  | 
text {*
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1300  | 
  The fourth argument of @{ML note in LocalTheory} is the list of theorems we
 | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1301  | 
want to store under a name. We can store more than one under a single name.  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1302  | 
  The first argument @{ML_ind theoremK in Thm} is
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1303  | 
a kind indicator, which classifies the theorem. There are several such kind  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1304  | 
  indicators: for a theorem arising from a definition you should use @{ML_ind
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1305  | 
  definitionK in Thm}, for an axiom @{ML_ind axiomK in Thm}, and for
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1306  | 
  ``normal'' theorems the kinds @{ML_ind theoremK in Thm} or @{ML_ind lemmaK
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1307  | 
  in Thm}.  The second argument of @{ML note in LocalTheory} is the name under
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1308  | 
which we store the theorem or theorems. The third argument can contain a  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1309  | 
list of theorem attributes, which we will explain in detail in  | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1310  | 
  Section~\ref{sec:attributes}. Below we just use one such attribute for
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1311  | 
adding the theorem to the simpset:  | 
| 
339
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1312  | 
*}  | 
| 
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1313  | 
|
| 
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1314  | 
local_setup %gray {*
 | 
| 
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1315  | 
LocalTheory.note Thm.theoremK  | 
| 347 | 1316  | 
    ((@{binding "my_thm_simp"}, 
 | 
1317  | 
[Attrib.internal (K Simplifier.simp_add)]), [my_thm]) #> snd *}  | 
|
| 
339
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1318  | 
|
| 
 
c588e8422737
used a better implementation of \index in Latex; added more to the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
338 
diff
changeset
 | 
1319  | 
text {*
 | 
| 348 | 1320  | 
Note that we have to use another name under which the theorem is stored,  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1321  | 
  since Isabelle does not allow us to call  @{ML_ind note in LocalTheory} twice
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1322  | 
  with the same name. The attribute needs to be wrapped inside the function @{ML_ind
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1323  | 
  internal in Attrib} from the structure @{ML_struct Attrib}. If we use the function 
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1324  | 
  @{ML get_thm_names_from_ss} from
 | 
| 348 | 1325  | 
the previous chapter, we can check whether the theorem has actually been  | 
1326  | 
added.  | 
|
1327  | 
||
| 340 | 1328  | 
|
1329  | 
  @{ML_response [display,gray]
 | 
|
1330  | 
"let  | 
|
1331  | 
fun pred s = match_string \"my_thm_simp\" s  | 
|
1332  | 
in  | 
|
1333  | 
  exists pred (get_thm_names_from_ss @{simpset})
 | 
|
1334  | 
end"  | 
|
1335  | 
"true"}  | 
|
1336  | 
||
| 347 | 1337  | 
  The main point of storing the theorems @{thm [source] my_thm} and @{thm
 | 
1338  | 
[source] my_thm_simp} is that they can now also be referenced with the  | 
|
1339  | 
  \isacommand{thm}-command on the user-level of Isabelle
 | 
|
1340  | 
||
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1341  | 
|
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1342  | 
  \begin{isabelle}
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1343  | 
  \isacommand{thm}~@{text "my_thm"}\isanewline
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1344  | 
   @{text ">"}~@{prop "\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t"}
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1345  | 
  \end{isabelle}
 | 
| 
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1346  | 
|
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1347  | 
  or with the @{text "@{thm \<dots>}"}-antiquotation on the ML-level. Otherwise the 
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1348  | 
user has no access to these theorems.  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1349  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1350  | 
  Recall that Isabelle does not let you call @{ML note in LocalTheory} twice
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1351  | 
with the same theorem name. In effect, once a theorem is stored under a name,  | 
| 358 | 1352  | 
this association is fixed. While this is a ``safety-net'' to make sure a  | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1353  | 
theorem name refers to a particular theorem or collection of theorems, it is  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1354  | 
also a bit too restrictive in cases where a theorem name should refer to a  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1355  | 
dynamically expanding list of theorems (like a simpset). Therefore Isabelle  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1356  | 
also implements a mechanism where a theorem name can refer to a custom theorem  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1357  | 
  list. For this you can use the function @{ML_ind add_thms_dynamic in PureThy}. 
 | 
| 358 | 1358  | 
  To see how it works let us assume we defined our own theorem list @{text MyThmList}.
 | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1359  | 
*}  | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1360  | 
|
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1361  | 
ML{*structure MyThmList = GenericDataFun
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1362  | 
(type T = thm list  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1363  | 
val empty = []  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1364  | 
val extend = I  | 
| 356 | 1365  | 
val merge = K (op @))  | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1366  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1367  | 
fun update thm = Context.theory_map (MyThmList.map (fn thms => thm::thms))*}  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1368  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1369  | 
text {*
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1370  | 
  The function @{ML update} allows us to update the theorem list, for example
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1371  | 
  by adding the theorem @{thm [source] TrueI}.
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1372  | 
*}  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1373  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1374  | 
setup %gray {* update @{thm TrueI} *}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1375  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1376  | 
text {*
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1377  | 
We can now install the theorem list so that it is visible to the user and  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1378  | 
can be refered to by a theorem name. For this need to call  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1379  | 
  @{ML_ind add_thms_dynamic in PureThy}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1380  | 
*}  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1381  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1382  | 
setup %gray {* 
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1383  | 
  PureThy.add_thms_dynamic (@{binding "mythmlist"}, MyThmList.get) 
 | 
| 347 | 1384  | 
*}  | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1385  | 
|
| 347 | 1386  | 
text {*
 | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1387  | 
with a name and a function that accesses the theorem list. Now if the  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1388  | 
user issues the command  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1389  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1390  | 
  \begin{isabelle}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1391  | 
  \isacommand{thm}~@{text "mythmlist"}\\
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1392  | 
  @{text "> True"}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1393  | 
  \end{isabelle}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1394  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1395  | 
the current content of the theorem list is displayed. If more theorems are stored in  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1396  | 
the list, say  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1397  | 
*}  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1398  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1399  | 
setup %gray {* update @{thm FalseE} *}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1400  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1401  | 
text {*
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1402  | 
then the same command produces  | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1403  | 
|
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1404  | 
  \begin{isabelle}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1405  | 
  \isacommand{thm}~@{text "mythmlist"}\\
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1406  | 
  @{text "> False \<Longrightarrow> ?P"}\\
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1407  | 
  @{text "> True"}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1408  | 
  \end{isabelle}
 | 
| 
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1409  | 
|
| 358 | 1410  | 
  There is a multitude of functions in the structures @{ML_struct Thm} and @{ML_struct Drule} 
 | 
1411  | 
for managing or manipulating theorems. For example  | 
|
| 348 | 1412  | 
we can test theorems for alpha equality. Suppose you proved the following three  | 
1413  | 
theorems.  | 
|
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1414  | 
*}  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1415  | 
|
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1416  | 
lemma  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1417  | 
shows thm1: "\<forall>x. P x"  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1418  | 
and thm2: "\<forall>y. P y"  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1419  | 
and thm3: "\<forall>y. Q y" sorry  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1420  | 
|
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1421  | 
text {*
 | 
| 348 | 1422  | 
  Testing them for alpha equality using the function @{ML_ind eq_thm in Thm} produces:
 | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1423  | 
|
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1424  | 
  @{ML_response [display,gray]
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1425  | 
"(Thm.eq_thm (@{thm thm1}, @{thm thm2}),
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1426  | 
 Thm.eq_thm (@{thm thm2}, @{thm thm3}))"
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1427  | 
"(true, false)"}  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1428  | 
|
| 340 | 1429  | 
  Many functions destruct theorems into @{ML_type cterm}s. For example
 | 
1430  | 
  the functions @{ML_ind lhs_of in Thm} and @{ML_ind rhs_of in Thm} return 
 | 
|
1431  | 
the left and right-hand side, respectively, of a meta-equality.  | 
|
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1432  | 
|
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1433  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1434  | 
"let  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1435  | 
  val eq = @{thm True_def}
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1436  | 
in  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1437  | 
(Thm.lhs_of eq, Thm.rhs_of eq)  | 
| 348 | 1438  | 
  |> pairself (string_of_cterm @{context})
 | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1439  | 
end"  | 
| 348 | 1440  | 
"(True, (\<lambda>x. x) = (\<lambda>x. x))"}  | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1441  | 
|
| 340 | 1442  | 
Other function produce terms that can be pattern-matched.  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1443  | 
Suppose the following two theorems.  | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1444  | 
*}  | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1445  | 
|
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1446  | 
lemma  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1447  | 
shows foo_test1: "A \<Longrightarrow> B \<Longrightarrow> C"  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1448  | 
and foo_test2: "A \<longrightarrow> B \<longrightarrow> C" sorry  | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1449  | 
|
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1450  | 
text {*
 | 
| 348 | 1451  | 
We can destruct them into premises and conclusions as follows.  | 
| 340 | 1452  | 
|
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1453  | 
 @{ML_response_fake [display,gray]
 | 
| 
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1454  | 
"let  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1455  | 
  val ctxt = @{context}
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1456  | 
fun prems_and_concl thm =  | 
| 348 | 1457  | 
[\"Premises: \" ^ (string_of_terms ctxt (Thm.prems_of thm))] @  | 
1458  | 
[\"Conclusion: \" ^ (string_of_term ctxt (Thm.concl_of thm))]  | 
|
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1459  | 
|> cat_lines  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1460  | 
|> tracing  | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1461  | 
in  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1462  | 
  prems_and_concl @{thm foo_test1};
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1463  | 
  prems_and_concl @{thm foo_test2}
 | 
| 
338
 
3bc732c9f7ff
more on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
337 
diff
changeset
 | 
1464  | 
end"  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1465  | 
"Premises: ?A, ?B  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1466  | 
Conclusion: ?C  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1467  | 
Premises:  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1468  | 
Conclusion: ?A \<longrightarrow> ?B \<longrightarrow> ?C"}  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1469  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1470  | 
Note that in the second case, there is no premise.  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1471  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1472  | 
  \begin{readmore}
 | 
| 358 | 1473  | 
The basic functions for theorems are defined in  | 
| 
337
 
a456a21f608a
a bit more work on the theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
336 
diff
changeset
 | 
1474  | 
  @{ML_file "Pure/thm.ML"}, @{ML_file "Pure/more_thm.ML"} and @{ML_file "Pure/drule.ML"}. 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1475  | 
  \end{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1476  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1477  | 
The simplifier can be used to unfold definition in theorems. To show  | 
| 382 | 1478  | 
  this, we build the theorem @{term "True \<equiv> True"} (Line 1) and then 
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1479  | 
  unfold the constant @{term "True"} according to its definition (Line 2).
 | 
| 347 | 1480  | 
|
1481  | 
  @{ML_response_fake [display,gray,linenos]
 | 
|
1482  | 
  "Thm.reflexive @{cterm \"True\"}
 | 
|
1483  | 
  |> Simplifier.rewrite_rule [@{thm True_def}]
 | 
|
1484  | 
  |> string_of_thm @{context}
 | 
|
1485  | 
|> tracing"  | 
|
1486  | 
"(\<lambda>x. x) = (\<lambda>x. x) \<equiv> (\<lambda>x. x) = (\<lambda>x. x)"}  | 
|
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1487  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1488  | 
Often it is necessary to transform theorems to and from the object  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1489  | 
  logic, that is replacing all @{text "\<longrightarrow>"} and @{text "\<forall>"} by @{text "\<Longrightarrow>"} 
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1490  | 
  and @{text "\<And>"}, or the other way around.  A reason for such a transformation 
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1491  | 
might be stating a definition. The reason is that definitions can only be  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1492  | 
stated using object logic connectives, while theorems using the connectives  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1493  | 
from the meta logic are more convenient for reasoning. Therefore there are  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1494  | 
some build in functions which help with these transformations. The function  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1495  | 
  @{ML_ind rulify in ObjectLogic} 
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1496  | 
replaces all object connectives by equivalents in the meta logic. For example  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1497  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1498  | 
  @{ML_response_fake [display, gray]
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1499  | 
  "ObjectLogic.rulify @{thm foo_test2}"
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1500  | 
"\<lbrakk>?A; ?B\<rbrakk> \<Longrightarrow> ?C"}  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1501  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1502  | 
The transformation in the other direction can be achieved with function  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1503  | 
  @{ML_ind atomize in ObjectLogic} and the following code.
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1504  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1505  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1506  | 
"let  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1507  | 
  val thm = @{thm foo_test1}
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1508  | 
val meta_eq = ObjectLogic.atomize (cprop_of thm)  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1509  | 
in  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1510  | 
MetaSimplifier.rewrite_rule [meta_eq] thm  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1511  | 
end"  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1512  | 
"?A \<longrightarrow> ?B \<longrightarrow> ?C"}  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1513  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1514  | 
  In this code the function @{ML atomize in ObjectLogic} produces 
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1515  | 
a meta-equation between the given theorem and the theorem transformed  | 
| 347 | 1516  | 
into the object logic. The result is the theorem with object logic  | 
1517  | 
connectives. However, in order to completely transform a theorem  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1518  | 
  involving meta variables, such as @{thm [source] list.induct}, which 
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1519  | 
is of the form  | 
| 347 | 1520  | 
|
1521  | 
  @{thm [display] list.induct}
 | 
|
1522  | 
||
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1523  | 
  we have to first abstract over the meta variables @{text "?P"} and 
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1524  | 
  @{text "?list"}. For this we can use the function 
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1525  | 
  @{ML_ind forall_intr_vars in Drule}. This allows us to implement the
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1526  | 
following function for atomizing a theorem.  | 
| 347 | 1527  | 
*}  | 
1528  | 
||
1529  | 
ML{*fun atomize_thm thm =
 | 
|
1530  | 
let  | 
|
1531  | 
val thm' = forall_intr_vars thm  | 
|
1532  | 
val thm'' = ObjectLogic.atomize (cprop_of thm')  | 
|
1533  | 
in  | 
|
1534  | 
MetaSimplifier.rewrite_rule [thm''] thm'  | 
|
1535  | 
end*}  | 
|
1536  | 
||
1537  | 
text {*
 | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1538  | 
  This function produces for the theorem @{thm [source] list.induct}
 | 
| 347 | 1539  | 
|
1540  | 
  @{ML_response_fake [display, gray]
 | 
|
1541  | 
  "atomize_thm @{thm list.induct}"
 | 
|
1542  | 
"\<forall>P list. P [] \<longrightarrow> (\<forall>a list. P list \<longrightarrow> P (a # list)) \<longrightarrow> P list"}  | 
|
1543  | 
||
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1544  | 
Theorems can also be produced from terms by giving an explicit proof.  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1545  | 
  One way to achieve this is by using the function @{ML_ind prove in Goal}
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1546  | 
  in the structure @{ML_struct Goal}. For example below we use this function
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1547  | 
  to prove the term @{term "P \<Longrightarrow> P"}.
 | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1548  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1549  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1550  | 
"let  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1551  | 
  val trm = @{term \"P \<Longrightarrow> P::bool\"}
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1552  | 
val tac = K (atac 1)  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1553  | 
in  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1554  | 
  Goal.prove @{context} [\"P\"] [] trm tac
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1555  | 
end"  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1556  | 
"?P \<Longrightarrow> ?P"}  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1557  | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1558  | 
This function takes first a context and second a list of strings. This list  | 
| 359 | 1559  | 
specifies which variables should be turned into schematic variables once the term  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1560  | 
is proved. The fourth argument is the term to be proved. The fifth is a  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1561  | 
corresponding proof given in form of a tactic (we explain tactics in  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1562  | 
  Chapter~\ref{chp:tactical}). In the code above, the tactic proves the theorem
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1563  | 
by assumption. As before this code will produce a theorem, but the theorem  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1564  | 
is not yet stored in the theorem database.  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1565  | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1566  | 
Theorems also contain auxiliary data, such as the name of the theorem, its  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1567  | 
kind, the names for cases and so on. This data is stored in a string-string  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1568  | 
  list and can be retrieved with the function @{ML_ind get_tags in
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1569  | 
Thm}. Assume you prove the following lemma.  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1570  | 
*}  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1571  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1572  | 
lemma foo_data:  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1573  | 
shows "P \<Longrightarrow> P \<Longrightarrow> P" by assumption  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1574  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1575  | 
text {*
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1576  | 
The auxiliary data of this lemma can be retrieved using the function  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1577  | 
  @{ML_ind get_tags in Thm}. So far the the auxiliary data of this lemma is 
 | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1578  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1579  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1580  | 
  "Thm.get_tags @{thm foo_data}"
 | 
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1581  | 
"[(\"name\", \"General.foo_data\"), (\"kind\", \"lemma\")]"}  | 
| 
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1582  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1583  | 
consisting of a name and a kind. When we store lemmas in the theorem database,  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1584  | 
we might want to explicitly extend this data by attaching case names to the  | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
1585  | 
  two premises of the lemma.  This can be done with the function @{ML_ind name in Rule_Cases}
 | 
| 
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
1586  | 
  from the structure @{ML_struct Rule_Cases}.
 | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1587  | 
*}  | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1588  | 
|
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1589  | 
local_setup %gray {*
 | 
| 
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1590  | 
LocalTheory.note Thm.lemmaK  | 
| 
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1591  | 
    ((@{binding "foo_data'"}, []), 
 | 
| 
375
 
92f7328dc5cc
added type work and updated to Isabelle and poly 5.3
 
Christian Urban <urbanc@in.tum.de> 
parents: 
374 
diff
changeset
 | 
1592  | 
[(Rule_Cases.name ["foo_case_one", "foo_case_two"]  | 
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1593  | 
        @{thm foo_data})]) #> snd *}
 | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1594  | 
|
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1595  | 
text {*
 | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1596  | 
  The data of the theorem @{thm [source] foo_data'} is then as follows:
 | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1597  | 
|
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1598  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1599  | 
  "Thm.get_tags @{thm foo_data'}"
 | 
| 
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1600  | 
"[(\"name\", \"General.foo_data'\"), (\"kind\", \"lemma\"),  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1601  | 
(\"case_names\", \"foo_case_one;foo_case_two\")]"}  | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1602  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1603  | 
You can observe the case names of this lemma on the user level when using  | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1604  | 
  the proof methods @{text cases} and @{text induct}. In the proof below
 | 
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1605  | 
*}  | 
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1606  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1607  | 
lemma  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1608  | 
shows "Q \<Longrightarrow> Q \<Longrightarrow> Q"  | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1609  | 
proof (cases rule: foo_data')  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1610  | 
|
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1611  | 
(*<*)oops(*>*)  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1612  | 
text_raw{*
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1613  | 
\begin{tabular}{@ {}l}
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1614  | 
\isacommand{print\_cases}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1615  | 
@{text "> cases:"}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1616  | 
@{text ">   foo_case_one:"}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1617  | 
@{text ">     let \"?case\" = \"?P\""}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1618  | 
@{text ">   foo_case_two:"}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1619  | 
@{text ">     let \"?case\" = \"?P\""}
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1620  | 
\end{tabular}*}
 | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1621  | 
|
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1622  | 
text {*
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1623  | 
  we can proceed by analysing the cases @{text "foo_case_one"} and 
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1624  | 
  @{text "foo_case_two"}. While if the theorem has no names, then
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1625  | 
  the cases have standard names @{text 1}, @{text 2} and so 
 | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1626  | 
on. This can be seen in the proof below.  | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1627  | 
*}  | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1628  | 
|
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1629  | 
lemma  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1630  | 
shows "Q \<Longrightarrow> Q \<Longrightarrow> Q"  | 
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
1631  | 
proof (cases rule: foo_data)  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1632  | 
|
| 
342
 
930b1308fd96
fixed glitch with tocibind
 
Christian Urban <urbanc@in.tum.de> 
parents: 
341 
diff
changeset
 | 
1633  | 
(*<*)oops(*>*)  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1634  | 
text_raw{*
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1635  | 
\begin{tabular}{@ {}l}
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1636  | 
\isacommand{print\_cases}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1637  | 
@{text "> cases:"}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1638  | 
@{text ">   1:"}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1639  | 
@{text ">     let \"?case\" = \"?P\""}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1640  | 
@{text ">   2:"}\\
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1641  | 
@{text ">     let \"?case\" = \"?P\""}
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1642  | 
\end{tabular}*}
 | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1643  | 
|
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1644  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
1645  | 
text {*
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1646  | 
One great feature of Isabelle is its document preparation system, where  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1647  | 
proved theorems can be quoted in documents referencing directly their  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1648  | 
formalisation. This helps tremendously to minimise cut-and-paste errors. However,  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1649  | 
sometimes the verbatim quoting is not what is wanted or what can be shown to  | 
| 354 | 1650  | 
  readers. For such situations Isabelle allows the installation of \emph{\index*{theorem 
 | 
1651  | 
styles}}. These are, roughly speaking, functions from terms to terms. The input  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1652  | 
term stands for the theorem to be presented; the output can be constructed to  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1653  | 
ones wishes. Let us, for example, assume we want to quote theorems without  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1654  | 
  leading @{text \<forall>}-quantifiers. For this we can implement the following function 
 | 
| 358 | 1655  | 
  that strips off @{text "\<forall>"}s.
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1656  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1657  | 
|
| 358 | 1658  | 
ML %linenosgray{*fun strip_allq (Const (@{const_name "All"}, _) $ Abs body) = 
 | 
| 354 | 1659  | 
Term.dest_abs body |> snd |> strip_allq  | 
1660  | 
  | strip_allq (Const (@{const_name "Trueprop"}, _) $ t) = 
 | 
|
1661  | 
strip_allq t  | 
|
1662  | 
| strip_allq t = t*}  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1663  | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1664  | 
text {*
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1665  | 
  We use in Line 2 the function @{ML_ind dest_abs in Term} for deconstructing abstractions,
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1666  | 
since this function deals correctly with potential name clashes. This function produces  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1667  | 
a pair consisting of the variable and the body of the abstraction. We are only interested  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1668  | 
in the body, which we feed into the recursive call. In Line 3 and 4, we also  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1669  | 
  have to explicitly strip of the outermost @{term Trueprop}-coercion. Now we can 
 | 
| 354 | 1670  | 
  install this function as the theorem style named @{text "my_strip_allq"}. 
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1671  | 
*}  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1672  | 
|
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1673  | 
setup %gray {*
 | 
| 354 | 1674  | 
Term_Style.setup "my_strip_allq" (Scan.succeed (K strip_allq))  | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1675  | 
*}  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1676  | 
|
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1677  | 
text {*
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1678  | 
We can test this theorem style with the following theorem  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1679  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1680  | 
|
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1681  | 
theorem style_test:  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1682  | 
shows "\<forall>x y z. (x, x) = (y, z)" sorry  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1683  | 
|
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1684  | 
text {*
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1685  | 
  Now printing out in a document the theorem @{thm [source] style_test} normally
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1686  | 
  using @{text "@{thm \<dots>}"} produces
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1687  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1688  | 
  \begin{isabelle}
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1689  | 
  @{text "@{thm style_test}"}\\
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1690  | 
  @{text ">"}~@{thm style_test}
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1691  | 
  \end{isabelle}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1692  | 
|
| 354 | 1693  | 
  as expected. But with the theorem style @{text "@{thm (my_strip_allq) \<dots>}"} 
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1694  | 
we obtain  | 
| 
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1695  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1696  | 
  \begin{isabelle}
 | 
| 354 | 1697  | 
  @{text "@{thm (my_strip_allq) style_test}"}\\
 | 
1698  | 
  @{text ">"}~@{thm (my_strip_allq) style_test}\\
 | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1699  | 
  \end{isabelle}
 | 
| 
352
 
9f12e53eb121
polished theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
351 
diff
changeset
 | 
1700  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1701  | 
without the leading quantifiers. We can improve this theorem style by explicitly  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1702  | 
giving a list of strings that should be used for the replacement of the  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1703  | 
variables. For this we implement the function which takes a list of strings  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1704  | 
and uses them as name in the outermost abstractions.  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1705  | 
*}  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1706  | 
|
| 358 | 1707  | 
ML{*fun rename_allq [] t = t
 | 
| 354 | 1708  | 
  | rename_allq (x::xs) (Const (@{const_name "All"}, U) $ Abs (_, T, t)) = 
 | 
1709  | 
      Const (@{const_name "All"}, U) $ Abs (x, T, rename_allq xs t)
 | 
|
1710  | 
  | rename_allq xs (Const (@{const_name "Trueprop"}, U) $ t) =
 | 
|
1711  | 
rename_allq xs t  | 
|
1712  | 
| rename_allq _ t = t*}  | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1713  | 
|
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1714  | 
text {*
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1715  | 
We can now install a the modified theorem style as follows  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1716  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1717  | 
|
| 356 | 1718  | 
setup %gray {* let
 | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1719  | 
val parser = Scan.repeat Args.name  | 
| 354 | 1720  | 
fun action xs = K (rename_allq xs #> strip_allq)  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1721  | 
in  | 
| 354 | 1722  | 
Term_Style.setup "my_strip_allq2" (parser >> action)  | 
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1723  | 
end *}  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1724  | 
|
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1725  | 
text {*
 | 
| 
355
 
42a1c230daff
added something about add_thms_dynamic
 
Christian Urban <urbanc@in.tum.de> 
parents: 
354 
diff
changeset
 | 
1726  | 
  The parser reads a list of names. In the function @{text action} we first
 | 
| 354 | 1727  | 
  call @{ML rename_allq} with the parsed list, then we call @{ML strip_allq}
 | 
1728  | 
on the resulting term. We can now suggest, for example, two variables for  | 
|
1729  | 
  stripping off the first two @{text \<forall>}-quantifiers.
 | 
|
1730  | 
||
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1731  | 
|
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1732  | 
  \begin{isabelle}
 | 
| 354 | 1733  | 
  @{text "@{thm (my_strip_allq2 x' x'') style_test}"}\\
 | 
1734  | 
  @{text ">"}~@{thm (my_strip_allq2 x' x'') style_test}
 | 
|
| 
353
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1735  | 
  \end{isabelle}
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1736  | 
|
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1737  | 
Such theorem styles allow one to print out theorems in documents formatted to  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1738  | 
ones heart content. Next we explain theorem attributes, which is another  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1739  | 
mechanism for dealing with theorems.  | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1740  | 
|
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1741  | 
  \begin{readmore}
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1742  | 
  Theorem styles are implemented in @{ML_file "Pure/Thy/term_style.ML"}.
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1743  | 
  \end{readmore}
 | 
| 
 
e73ccbed776e
completed section on theorems
 
Christian Urban <urbanc@in.tum.de> 
parents: 
352 
diff
changeset
 | 
1744  | 
*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1745  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1746  | 
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
 | 
1747  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1748  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1749  | 
  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
 | 
1750  | 
  "[simp]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
 | 
| 356 | 1751  | 
annotated to theorems, but functions that do further processing of  | 
1752  | 
theorems. In particular, it is not possible to find out  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1753  | 
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
 | 
1754  | 
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
 | 
1755  | 
data structure.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1756  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1757  | 
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
 | 
1758  | 
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
 | 
1759  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1760  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1761  | 
  \isacommand{print\_attributes}\\
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1762  | 
  @{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
 | 
1763  | 
  @{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
 | 
1764  | 
  @{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
 | 
1765  | 
  @{text "> \<dots>"}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1766  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1767  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1768  | 
The theorem attributes fall roughly into two categories: the first category manipulates  | 
| 356 | 1769  | 
  theorems (for example @{text "[symmetric]"} and @{text "[THEN \<dots>]"}), and the second
 | 
1770  | 
  stores theorems somewhere as data (for example @{text "[simp]"}, which adds
 | 
|
1771  | 
theorems to the current simpset).  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1772  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1773  | 
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
 | 
1774  | 
  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
 | 
1775  | 
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
 | 
1776  | 
this attribute is  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1777  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1778  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1779  | 
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
 | 
1780  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1781  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1782  | 
  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
 | 
1783  | 
  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
 | 
1784  | 
  returns another theorem (namely @{text thm} resolved with the theorem 
 | 
| 363 | 1785  | 
  @{thm [source] sym}: @{thm sym[no_vars]}; the function @{ML_ind RS in Drule} 
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1786  | 
  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
 | 
1787  | 
  @{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
 | 
1788  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1789  | 
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
 | 
1790  | 
  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
 | 
1791  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1792  | 
|
| 356 | 1793  | 
attribute_setup %gray my_sym =  | 
1794  | 
  {* Scan.succeed my_symmetric *} "applying the sym rule"
 | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1795  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1796  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1797  | 
  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
 | 
1798  | 
for the theorem attribute. Since the attribute does not expect any further  | 
| 356 | 1799  | 
  arguments (unlike @{text "[THEN \<dots>]"}, for instance), we use the parser @{ML
 | 
1800  | 
  Scan.succeed}. An example for the attribute @{text "[my_sym]"} is the proof
 | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1801  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1802  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1803  | 
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
 | 
1804  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1805  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1806  | 
  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
 | 
1807  | 
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
 | 
1808  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1809  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1810  | 
  \isacommand{thm}~@{text "test"}\\
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1811  | 
  @{text "> "}~@{thm test}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1812  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1813  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1814  | 
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
 | 
1815  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1816  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1817  | 
  \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
 | 
1818  | 
  @{text "> "}~@{thm test[my_sym]}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1819  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1820  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1821  | 
  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
 | 
1822  | 
  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
 | 
1823  | 
attribute as follows:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1824  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1825  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1826  | 
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
 | 
1827  | 
"applying the sym rule" *}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1828  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1829  | 
text {*
 | 
| 356 | 1830  | 
  This gives a function from @{ML_type "theory -> theory"}, which
 | 
| 361 | 1831  | 
  can be used for example with \isacommand{setup} or with
 | 
| 
368
 
b1a458a03a8e
new parts in the tactical section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
363 
diff
changeset
 | 
1832  | 
  @{ML "Context.>> o Context.map_theory"}.\footnote{\bf FIXME: explain what happens here.}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1833  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1834  | 
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
 | 
1835  | 
  our own version of @{text "[THEN \<dots>]"}. This attribute will take a list of theorems
 | 
| 356 | 1836  | 
as argument and resolve the theorem with this list (one theorem  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1837  | 
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
 | 
1838  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1839  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1840  | 
ML{*fun MY_THEN thms = 
 | 
| 
368
 
b1a458a03a8e
new parts in the tactical section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
363 
diff
changeset
 | 
1841  | 
Thm.rule_attribute  | 
| 
 
b1a458a03a8e
new parts in the tactical section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
363 
diff
changeset
 | 
1842  | 
(fn _ => fn thm => fold (curry ((op RS) o swap)) thms thm)*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1843  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1844  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1845  | 
  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
 | 
1846  | 
  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
 | 
1847  | 
theorems.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1848  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1849  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1850  | 
attribute_setup %gray MY_THEN = {* Attrib.thms >> MY_THEN *} 
 | 
| 356 | 1851  | 
"resolving the list of theorems with the theorem"  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1852  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1853  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1854  | 
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
 | 
1855  | 
meta-equation:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1856  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1857  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1858  | 
  \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
 | 
1859  | 
  @{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
 | 
1860  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1861  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1862  | 
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
 | 
1863  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1864  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1865  | 
  \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
 | 
1866  | 
  @{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
 | 
1867  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1868  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1869  | 
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
 | 
1870  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1871  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1872  | 
  \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
 | 
1873  | 
  @{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
 | 
1874  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1875  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1876  | 
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
 | 
1877  | 
of theorem attributes shows through: since theorem attributes can be  | 
| 329 | 1878  | 
arbitrary functions, they do not commute in general. If you try  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1879  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1880  | 
  \begin{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1881  | 
  \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
 | 
1882  | 
  @{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
 | 
1883  | 
  \end{isabelle}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1884  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1885  | 
  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
 | 
1886  | 
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
 | 
1887  | 
|
| 329 | 1888  | 
  The purpose of @{ML_ind rule_attribute in Thm} is to directly manipulate
 | 
1889  | 
theorems. Another usage of theorem attributes is to add and delete theorems  | 
|
1890  | 
  from stored data.  For example the theorem attribute @{text "[simp]"} adds
 | 
|
1891  | 
or deletes a theorem from the current simpset. For these applications, you  | 
|
1892  | 
  can use @{ML_ind declaration_attribute in Thm}. To illustrate this function,
 | 
|
1893  | 
let us introduce a theorem list.  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1894  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1895  | 
|
| 329 | 1896  | 
ML{*structure MyThms = Named_Thms
 | 
1897  | 
(val name = "attr_thms"  | 
|
1898  | 
val description = "Theorems for an Attribute") *}  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1899  | 
|
| 329 | 1900  | 
text {* 
 | 
1901  | 
We are going to modify this list by adding and deleting theorems.  | 
|
1902  | 
  For this we use the two functions @{ML MyThms.add_thm} and
 | 
|
1903  | 
  @{ML MyThms.del_thm}. You can turn them into attributes 
 | 
|
1904  | 
with the code  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1905  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1906  | 
|
| 329 | 1907  | 
ML{*val my_add = Thm.declaration_attribute MyThms.add_thm
 | 
1908  | 
val my_del = Thm.declaration_attribute MyThms.del_thm *}  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1909  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1910  | 
text {* 
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1911  | 
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
 | 
1912  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1913  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1914  | 
attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
 | 
| 329 | 1915  | 
"maintaining a list of my_thms"  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1916  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1917  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1918  | 
  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
 | 
1919  | 
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
 | 
1920  | 
  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
 | 
1921  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1922  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1923  | 
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
 | 
1924  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1925  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1926  | 
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
 | 
1927  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1928  | 
  @{ML_response_fake [display,gray]
 | 
| 329 | 1929  | 
  "MyThms.get @{context}" 
 | 
1930  | 
"[\"True\"]"}  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1931  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1932  | 
  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
 | 
1933  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1934  | 
|
| 329 | 1935  | 
declare test[my_thms] trueI_2[my_thms add]  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1936  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1937  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1938  | 
  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
 | 
1939  | 
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
 | 
1940  | 
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
 | 
1941  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1942  | 
  @{ML_response_fake [display,gray]
 | 
| 329 | 1943  | 
  "MyThms.get @{context}"
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1944  | 
"[\"True\", \"Suc (Suc 0) = 2\"]"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1945  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1946  | 
  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
 | 
1947  | 
  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
 | 
1948  | 
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
 | 
1949  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1950  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1951  | 
declare test[my_thms del]  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1952  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1953  | 
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
 | 
1954  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1955  | 
  @{ML_response_fake [display,gray]
 | 
| 329 | 1956  | 
  "MyThms.get @{context}"
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1957  | 
"[\"True\"]"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1958  | 
|
| 329 | 1959  | 
  We used in this example two functions declared as @{ML_ind
 | 
1960  | 
declaration_attribute in Thm}, but there can be any number of them. We just  | 
|
1961  | 
have to change the parser for reading the arguments accordingly.  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1962  | 
|
| 329 | 1963  | 
  \footnote{\bf FIXME What are: @{text "theory_attributes"}, @{text "proof_attributes"}?}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1964  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1965  | 
  \begin{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1966  | 
  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
 | 
1967  | 
  @{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
 | 
1968  | 
parsing.  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1969  | 
  \end{readmore}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1970  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
1971  | 
|
| 358 | 1972  | 
section {* Theories\label{sec:theories} (TBD) *}
 | 
1973  | 
||
1974  | 
||
1975  | 
text {*
 | 
|
1976  | 
Theories are the most fundamental building blocks in Isabelle. They  | 
|
1977  | 
contain definitions, syntax declarations, axioms, proofs etc. If a definition  | 
|
1978  | 
is stated, it must be stored in a theory in order to be usable later  | 
|
1979  | 
on. Similar with proofs: once a proof is finished, the proved theorem  | 
|
1980  | 
needs to be stored in the theorem database of the theory in order to  | 
|
1981  | 
be usable. All relevant data of a theort can be querried as follows.  | 
|
1982  | 
||
1983  | 
  \begin{isabelle}
 | 
|
1984  | 
  \isacommand{print\_theory}\\
 | 
|
1985  | 
  @{text "> names: Pure Code_Generator HOL \<dots>"}\\
 | 
|
1986  | 
  @{text "> classes: Inf < type \<dots>"}\\
 | 
|
1987  | 
  @{text "> default sort: type"}\\
 | 
|
1988  | 
  @{text "> syntactic types: #prop \<dots>"}\\
 | 
|
1989  | 
  @{text "> logical types: 'a \<times> 'b \<dots>"}\\
 | 
|
1990  | 
  @{text "> type arities: * :: (random, random) random \<dots>"}\\
 | 
|
1991  | 
  @{text "> logical constants: == :: 'a \<Rightarrow> 'a \<Rightarrow> prop \<dots>"}\\
 | 
|
1992  | 
  @{text "> abbreviations: \<dots>"}\\
 | 
|
1993  | 
  @{text "> axioms: \<dots>"}\\
 | 
|
1994  | 
  @{text "> oracles: \<dots>"}\\
 | 
|
1995  | 
  @{text "> definitions: \<dots>"}\\
 | 
|
1996  | 
  @{text "> theorems: \<dots>"}
 | 
|
1997  | 
  \end{isabelle}
 | 
|
1998  | 
||
1999  | 
  \begin{center}
 | 
|
2000  | 
  \begin{tikzpicture}
 | 
|
2001  | 
  \node[top color=white, bottom color=gray!30, draw=black!100, drop shadow] {A};
 | 
|
2002  | 
  \end{tikzpicture}
 | 
|
2003  | 
  \end{center}
 | 
|
2004  | 
||
2005  | 
||
2006  | 
In contrast to an ordinary theory, which simply consists of a type  | 
|
2007  | 
signature, as well as tables for constants, axioms and theorems, a local  | 
|
2008  | 
theory contains additional context information, such as locally fixed  | 
|
2009  | 
variables and local assumptions that may be used by the package. The type  | 
|
2010  | 
  @{ML_type local_theory} is identical to the type of \emph{proof contexts}
 | 
|
2011  | 
  @{ML_type "Proof.context"}, although not every proof context constitutes a
 | 
|
2012  | 
valid local theory.  | 
|
| 361 | 2013  | 
|
2014  | 
  @{ML "Context.>> o Context.map_theory"}
 | 
|
| 358 | 2015  | 
*}  | 
2016  | 
||
| 348 | 2017  | 
section {* Setups (TBD) *}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2018  | 
|
| 348 | 2019  | 
text {*
 | 
| 361 | 2020  | 
  @{ML Sign.declare_const}
 | 
2021  | 
||
| 348 | 2022  | 
  In the previous section we used \isacommand{setup} in order to make
 | 
2023  | 
a theorem attribute known to Isabelle. What happens behind the scenes  | 
|
2024  | 
  is that \isacommand{setup} expects a function of type 
 | 
|
2025  | 
  @{ML_type "theory -> theory"}: the input theory is the current theory and the 
 | 
|
2026  | 
output the theory where the theory attribute has been stored.  | 
|
2027  | 
||
2028  | 
This is a fundamental principle in Isabelle. A similar situation occurs  | 
|
2029  | 
for example with declaring constants. The function that declares a  | 
|
2030  | 
  constant on the ML-level is @{ML_ind  add_consts_i in Sign}. 
 | 
|
2031  | 
  If you write\footnote{Recall that ML-code  needs to be 
 | 
|
2032  | 
  enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
 | 
|
2033  | 
*}  | 
|
2034  | 
||
2035  | 
ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
 | 
|
2036  | 
||
2037  | 
text {*
 | 
|
2038  | 
  for declaring the constant @{text "BAR"} with type @{typ nat} and 
 | 
|
2039  | 
run the code, then you indeed obtain a theory as result. But if you  | 
|
2040  | 
  query the constant on the Isabelle level using the command \isacommand{term}
 | 
|
2041  | 
||
2042  | 
  \begin{isabelle}
 | 
|
2043  | 
  \isacommand{term}~@{text [quotes] "BAR"}\\
 | 
|
2044  | 
  @{text "> \"BAR\" :: \"'a\""}
 | 
|
2045  | 
  \end{isabelle}
 | 
|
2046  | 
||
2047  | 
  you do not obtain a constant of type @{typ nat}, but a free variable (printed in 
 | 
|
2048  | 
blue) of polymorphic type. The problem is that the ML-expression above did  | 
|
2049  | 
not register the declaration with the current theory. This is what the command  | 
|
2050  | 
  \isacommand{setup} is for. The constant is properly declared with
 | 
|
2051  | 
*}  | 
|
2052  | 
||
2053  | 
setup %gray {* Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] *}
 | 
|
2054  | 
||
2055  | 
text {* 
 | 
|
2056  | 
Now  | 
|
2057  | 
||
2058  | 
  \begin{isabelle}
 | 
|
2059  | 
  \isacommand{term}~@{text [quotes] "BAR"}\\
 | 
|
2060  | 
  @{text "> \"BAR\" :: \"nat\""}
 | 
|
2061  | 
  \end{isabelle}
 | 
|
2062  | 
||
2063  | 
  returns a (black) constant with the type @{typ nat}.
 | 
|
2064  | 
||
2065  | 
  A similar command is \isacommand{local\_setup}, which expects a function
 | 
|
2066  | 
  of type @{ML_type "local_theory -> local_theory"}. Later on we will also
 | 
|
2067  | 
  use the commands \isacommand{method\_setup} for installing methods in the
 | 
|
2068  | 
  current theory and \isacommand{simproc\_setup} for adding new simprocs to
 | 
|
2069  | 
the current simpset.  | 
|
2070  | 
*}  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2071  | 
|
| 
341
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
2072  | 
section {* Contexts (TBD) *}
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
2073  | 
|
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
2074  | 
section {* Local Theories (TBD) *}
 | 
| 
 
62dea749d5ed
more work on theorem section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
340 
diff
changeset
 | 
2075  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2076  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2077  | 
(*  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2078  | 
setup {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2079  | 
 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
 | 
2080  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2081  | 
lemma "bar = (1::nat)"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2082  | 
oops  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2083  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2084  | 
setup {*   
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2085  | 
  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
 | 
2086  | 
 #> 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
 | 
2087  | 
       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
 | 
2088  | 
#> snd  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2089  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2090  | 
*)  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2091  | 
(*  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2092  | 
lemma "foo = (1::nat)"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2093  | 
apply(simp add: foo_def)  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2094  | 
done  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2095  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2096  | 
thm foo_def  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2097  | 
*)  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2098  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2099  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2100  | 
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
 | 
2101  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2102  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2103  | 
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
 | 
2104  | 
  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
 | 
2105  | 
sufficient for ``quick-and-dirty'' printouts. For something more  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2106  | 
sophisticated, Isabelle includes an infrastructure for properly formatting  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2107  | 
  text. Most of its functions do not operate on @{ML_type string}s, but on
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2108  | 
instances of the pretty type:  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2109  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2110  | 
  @{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
 | 
2111  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2112  | 
  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
 | 
2113  | 
type. For example  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2114  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2115  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2116  | 
"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
 | 
2117  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2118  | 
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
 | 
2119  | 
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
 | 
2120  | 
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
 | 
2121  | 
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
 | 
2122  | 
  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
 | 
2123  | 
  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
 | 
2124  | 
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
 | 
2125  | 
type:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2126  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2127  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2128  | 
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
 | 
2129  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2130  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2131  | 
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
 | 
2132  | 
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
 | 
2133  | 
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
 | 
2134  | 
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
 | 
2135  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2136  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2137  | 
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
 | 
2138  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2139  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2140  | 
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
 | 
2141  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2142  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2143  | 
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
 | 
2144  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2145  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2146  | 
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
 | 
2147  | 
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
 | 
2148  | 
we obtain the ugly output:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2149  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2150  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2151  | 
"tracing test_str"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2152  | 
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2153  | 
ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2154  | 
aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2155  | 
oooooooooooooobaaaaaaaaaaaar"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2156  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2157  | 
  We obtain the same if we just use the function @{ML pprint}.
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2158  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2159  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2160  | 
"pprint (Pretty.str test_str)"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2161  | 
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2162  | 
ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2163  | 
aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2164  | 
oooooooooooooobaaaaaaaaaaaar"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2165  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2166  | 
However by using pretty types you have the ability to indicate possible  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2167  | 
linebreaks for example at each whitespace. You can achieve this with the  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2168  | 
  function @{ML_ind breaks in Pretty}, which expects a list of pretty types
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2169  | 
and inserts a possible line break in between every two elements in this  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2170  | 
list. To print this list of pretty types as a single string, we concatenate  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2171  | 
  them with the function @{ML_ind blk in Pretty} as follows:
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2172  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2173  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2174  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2175  | 
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
 | 
2176  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2177  | 
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
 | 
2178  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2179  | 
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2180  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2181  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2182  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2183  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2184  | 
  Here the layout of @{ML test_str} is much more pleasing to the 
 | 
| 360 | 2185  | 
  eye. The @{ML "0"} in @{ML_ind  blk in Pretty} stands for no hanging 
 | 
2186  | 
indentation of the printed string. You can increase the indentation  | 
|
2187  | 
and obtain  | 
|
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2188  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2189  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2190  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2191  | 
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
 | 
2192  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2193  | 
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
 | 
2194  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2195  | 
"fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2196  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2197  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2198  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2199  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2200  | 
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
 | 
2201  | 
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
 | 
2202  | 
  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
 | 
2203  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2204  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2205  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2206  | 
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
 | 
2207  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2208  | 
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
 | 
2209  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2210  | 
" fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2211  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2212  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2213  | 
fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2214  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2215  | 
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
 | 
2216  | 
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
 | 
2217  | 
  @{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
 | 
2218  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2219  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2220  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2221  | 
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
 | 
2222  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2223  | 
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
 | 
2224  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2225  | 
"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
 | 
2226  | 
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
 | 
2227  | 
100016, 100017, 100018, 100019, 100020"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2228  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2229  | 
  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
 | 
2230  | 
  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
 | 
2231  | 
  @{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
 | 
2232  | 
  @{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
 | 
2233  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2234  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2235  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2236  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2237  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2238  | 
"let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2239  | 
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
 | 
2240  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2241  | 
  pprint (Pretty.enum \",\" \"{\" \"}\" ptrs)
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2242  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2243  | 
"{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
 | 
2244  | 
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
 | 
2245  | 
100016, 100017, 100018, 100019, 100020}"}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2246  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2247  | 
As can be seen, this function prints out the ``set'' so that starting  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2248  | 
from the second, each new line has an indentation of 2.  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2249  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2250  | 
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
 | 
2251  | 
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
 | 
2252  | 
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
 | 
2253  | 
  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
 | 
2254  | 
write the function  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2255  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2256  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2257  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2258  | 
ML %linenosgray{*fun and_list [] = []
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2259  | 
| and_list [x] = [x]  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2260  | 
| and_list xs =  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2261  | 
let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2262  | 
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
 | 
2263  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2264  | 
(Pretty.commas front) @  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2265  | 
[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
 | 
2266  | 
end *}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2267  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2268  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2269  | 
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
 | 
2270  | 
  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
 | 
2271  | 
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
 | 
2272  | 
  @{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
 | 
2273  | 
  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
 | 
2274  | 
for example  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2275  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2276  | 
@{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2277  | 
"let  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2278  | 
val ptrs1 = map (Pretty.str o string_of_int) (1 upto 22)  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2279  | 
val ptrs2 = map (Pretty.str o string_of_int) (10 upto 28)  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2280  | 
in  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2281  | 
pprint (Pretty.blk (0, and_list ptrs1));  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2282  | 
pprint (Pretty.blk (0, and_list ptrs2))  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2283  | 
end"  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2284  | 
"1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2285  | 
and 22  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2286  | 
10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27 and  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2287  | 
28"}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2288  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2289  | 
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
 | 
2290  | 
  function @{text "tell_type"}.
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2291  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2292  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2293  | 
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
 | 
2294  | 
let  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2295  | 
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
 | 
2296  | 
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
 | 
2297  | 
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
 | 
2298  | 
in  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2299  | 
pprint (Pretty.blk (0,  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2300  | 
(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
 | 
2301  | 
@ [pty, Pretty.str "."])))  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2302  | 
end*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2303  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2304  | 
text {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2305  | 
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
 | 
2306  | 
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
 | 
2307  | 
  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
 | 
2308  | 
  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
 | 
2309  | 
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
 | 
2310  | 
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
 | 
2311  | 
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
 | 
2312  | 
Now you can write  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2313  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2314  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2315  | 
  "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
 | 
2316  | 
"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
 | 
2317  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2318  | 
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
 | 
2319  | 
and type. For example:  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2320  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2321  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2322  | 
  "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
 | 
2323  | 
"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
 | 
2324  | 
\"((((('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
 | 
2325  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2326  | 
  The function @{ML_ind big_list in Pretty} helps with printing long lists.
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2327  | 
  It is used for example in the command \isacommand{print\_theorems}. An
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2328  | 
example is as follows.  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2329  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2330  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2331  | 
"let  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2332  | 
val pstrs = map (Pretty.str o string_of_int) (4 upto 10)  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2333  | 
in  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2334  | 
pprint (Pretty.big_list \"header\" pstrs)  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2335  | 
end"  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2336  | 
"header  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2337  | 
4  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2338  | 
5  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2339  | 
6  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2340  | 
7  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2341  | 
8  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2342  | 
9  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2343  | 
10"}  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2344  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2345  | 
  Like @{ML blk in Pretty}, the function @{ML_ind chunks in Pretty} prints out 
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2346  | 
a list of items, but automatically inserts forced breaks between each item.  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2347  | 
Compare  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2348  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2349  | 
  @{ML_response_fake [display,gray]
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2350  | 
"let  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2351  | 
val a_and_b = [Pretty.str \"a\", Pretty.str \"b\"]  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2352  | 
in  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2353  | 
pprint (Pretty.blk (0, a_and_b));  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2354  | 
pprint (Pretty.chunks a_and_b)  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2355  | 
end"  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2356  | 
"ab  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2357  | 
a  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2358  | 
b"}  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2359  | 
|
| 356 | 2360  | 
  \footnote{\bf FIXME: What happens with printing big formulae?}
 | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2361  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2362  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2363  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2364  | 
  \begin{readmore}
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2365  | 
The general infrastructure for pretty-printing is implemented in the file  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2366  | 
  @{ML_file "Pure/General/pretty.ML"}. The file @{ML_file "Pure/Syntax/syntax.ML"}
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2367  | 
contains pretty-printing functions for terms, types, theorems and so on.  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2368  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2369  | 
  @{ML_file "Pure/General/markup.ML"}
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2370  | 
  \end{readmore}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2371  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2372  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2373  | 
(*  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2374  | 
text {*
 | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2375  | 
printing into the goal buffer as part of the proof state  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2376  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2377  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2378  | 
text {* writing into the goal buffer *}
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2379  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2380  | 
ML {* fun my_hook interactive state =
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2381  | 
(interactive ? Proof.goal_message (fn () => Pretty.str  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2382  | 
"foo")) state  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2383  | 
*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2384  | 
|
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2385  | 
setup %gray {* Context.theory_map (Specification.add_theorem_hook my_hook) *}
 | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2386  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2387  | 
lemma "False"  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2388  | 
oops  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2389  | 
*)  | 
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2390  | 
|
| 
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2391  | 
(*  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2392  | 
ML {*
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2393  | 
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
 | 
2394  | 
setmp show_all_types  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2395  | 
(! 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
 | 
2396  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2397  | 
val ctxt = @{context};
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2398  | 
val t1 = @{term "undefined::nat"};
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2399  | 
val t2 = @{term "Suc y"};
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2400  | 
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
 | 
2401  | 
[(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
 | 
2402  | 
(Syntax.pretty_term ctxt) t1,  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2403  | 
Pretty.str "=", Syntax.pretty_term ctxt t2]);  | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2404  | 
pty |> Pretty.string_of  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2405  | 
*}  | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2406  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2407  | 
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
 | 
2408  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2409  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2410  | 
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
 | 
2411  | 
ML {* (Pretty.str "bar") |> Pretty.string_of |> tracing *}
 | 
| 
336
 
a12bb28fe2bd
polished on the pretty printing section
 
Christian Urban <urbanc@in.tum.de> 
parents: 
335 
diff
changeset
 | 
2412  | 
*)  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2413  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2414  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2415  | 
section {* Misc (TBD) *}
 | 
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2416  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2417  | 
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
 | 
2418  | 
|
| 
319
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2419  | 
text {* 
 | 
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2420  | 
FIXME: association lists:  | 
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2421  | 
@{ML_file "Pure/General/alist.ML"}
 | 
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2422  | 
|
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2423  | 
FIXME: calling the ML-compiler  | 
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2424  | 
|
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2425  | 
*}  | 
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2426  | 
|
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2427  | 
|
| 
 
6bce4acf7f2a
added file for producing a keyword file
 
Christian Urban <urbanc@in.tum.de> 
parents: 
318 
diff
changeset
 | 
2428  | 
|
| 
335
 
163ac0662211
reorganised the certified terms section; tuned
 
Christian Urban <urbanc@in.tum.de> 
parents: 
329 
diff
changeset
 | 
2429  | 
section {* Managing Name Spaces (TBD) *}
 | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2430  | 
|
| 360 | 2431  | 
ML {* Sign.intern_type @{theory} "list" *}
 | 
2432  | 
ML {* Sign.intern_const @{theory} "prod_fun" *}
 | 
|
2433  | 
||
| 
349
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2434  | 
section {* Summary *}
 | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2435  | 
|
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2436  | 
text {*
 | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2437  | 
  \begin{conventions}
 | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2438  | 
  \begin{itemize}
 | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2439  | 
\item Start with a proper context and then pass it around  | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2440  | 
through all your functions.  | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2441  | 
  \end{itemize}
 | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2442  | 
  \end{conventions}
 | 
| 
 
9e374cd891e1
updated to Isabelle changes
 
Christian Urban <urbanc@in.tum.de> 
parents: 
348 
diff
changeset
 | 
2443  | 
*}  | 
| 
318
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2444  | 
|
| 
 
efb5fff99c96
split up the first-steps section into two chapters
 
Christian Urban <urbanc@in.tum.de> 
parents:  
diff
changeset
 | 
2445  | 
end  |