68 |
68 |
69 Be careful if you pretty-print terms. Consider pretty-printing the abstraction |
69 Be careful if you pretty-print terms. Consider pretty-printing the abstraction |
70 term shown above: |
70 term shown above: |
71 |
71 |
72 @{ML_matchresult_fake [display, gray] |
72 @{ML_matchresult_fake [display, gray] |
73 "@{term \"\<lambda>x y. x y\"} |
73 \<open>@{term "\<lambda>x y. x y"} |
74 |> pretty_term @{context} |
74 |> pretty_term @{context} |
75 |> pwriteln" |
75 |> pwriteln\<close> |
76 "\<lambda>x. x"} |
76 \<open>\<lambda>x. x\<close>} |
77 |
77 |
78 This is one common source for puzzlement in Isabelle, which has |
78 This is one common source for puzzlement in Isabelle, which has |
79 tacitly eta-contracted the output. You obtain a similar result |
79 tacitly eta-contracted the output. You obtain a similar result |
80 with beta-redexes |
80 with beta-redexes |
81 |
81 |
82 @{ML_matchresult_fake [display, gray] |
82 @{ML_matchresult_fake [display, gray] |
83 "let |
83 \<open>let |
84 val redex = @{term \"(\<lambda>(x::int) (y::int). x)\"} |
84 val redex = @{term "(\<lambda>(x::int) (y::int). x)"} |
85 val arg1 = @{term \"1::int\"} |
85 val arg1 = @{term "1::int"} |
86 val arg2 = @{term \"2::int\"} |
86 val arg2 = @{term "2::int"} |
87 in |
87 in |
88 pretty_term @{context} (redex $ arg1 $ arg2) |
88 pretty_term @{context} (redex $ arg1 $ arg2) |
89 |> pwriteln |
89 |> pwriteln |
90 end" |
90 end\<close> |
91 "1"} |
91 \<open>1\<close>} |
92 |
92 |
93 There is a dedicated configuration value for switching off tacit |
93 There is a dedicated configuration value for switching off tacit |
94 eta-contractions, namely @{ML_ind eta_contract in Syntax} (see Section |
94 eta-contractions, namely @{ML_ind eta_contract in Syntax} (see Section |
95 \ref{sec:printing}), but none for beta-contractions. However you can avoid |
95 \ref{sec:printing}), but none for beta-contractions. However you can avoid |
96 the beta-contractions by switching off abbreviations using the configuration |
96 the beta-contractions by switching off abbreviations using the configuration |
97 value @{ML_ind show_abbrevs in Syntax}. For example |
97 value @{ML_ind show_abbrevs in Syntax}. For example |
98 |
98 |
99 |
99 |
100 @{ML_matchresult_fake [display, gray] |
100 @{ML_matchresult_fake [display, gray] |
101 "let |
101 \<open>let |
102 val redex = @{term \"(\<lambda>(x::int) (y::int). x)\"} |
102 val redex = @{term "(\<lambda>(x::int) (y::int). x)"} |
103 val arg1 = @{term \"1::int\"} |
103 val arg1 = @{term "1::int"} |
104 val arg2 = @{term \"2::int\"} |
104 val arg2 = @{term "2::int"} |
105 val ctxt = Config.put show_abbrevs false @{context} |
105 val ctxt = Config.put show_abbrevs false @{context} |
106 in |
106 in |
107 pretty_term ctxt (redex $ arg1 $ arg2) |
107 pretty_term ctxt (redex $ arg1 $ arg2) |
108 |> pwriteln |
108 |> pwriteln |
109 end" |
109 end\<close> |
110 "(\<lambda>x y. x) 1 2"} |
110 \<open>(\<lambda>x y. x) 1 2\<close>} |
111 |
111 |
112 Isabelle makes a distinction between \emph{free} variables (term-constructor |
112 Isabelle makes a distinction between \emph{free} variables (term-constructor |
113 @{ML Free} and written on the user level in blue colour) and |
113 @{ML Free} and written on the user level in blue colour) and |
114 \emph{schematic} variables (term-constructor @{ML Var} and written with a |
114 \emph{schematic} variables (term-constructor @{ML Var} and written with a |
115 leading question mark). Consider the following two examples |
115 leading question mark). Consider the following two examples |
116 |
116 |
117 @{ML_matchresult_fake [display, gray] |
117 @{ML_matchresult_fake [display, gray] |
118 "let |
118 \<open>let |
119 val v1 = Var ((\"x\", 3), @{typ bool}) |
119 val v1 = Var (("x", 3), @{typ bool}) |
120 val v2 = Var ((\"x1\", 3), @{typ bool}) |
120 val v2 = Var (("x1", 3), @{typ bool}) |
121 val v3 = Free (\"x\", @{typ bool}) |
121 val v3 = Free ("x", @{typ bool}) |
122 in |
122 in |
123 pretty_terms @{context} [v1, v2, v3] |
123 pretty_terms @{context} [v1, v2, v3] |
124 |> pwriteln |
124 |> pwriteln |
125 end" |
125 end\<close> |
126 "?x3, ?x1.3, x"} |
126 \<open>?x3, ?x1.3, x\<close>} |
127 |
127 |
128 When constructing terms, you are usually concerned with free |
128 When constructing terms, you are usually concerned with free |
129 variables (as mentioned earlier, you cannot construct schematic |
129 variables (as mentioned earlier, you cannot construct schematic |
130 variables using the built-in antiquotation \mbox{\<open>@{term |
130 variables using the built-in antiquotation \mbox{\<open>@{term |
131 \<dots>}\<close>}). If you deal with theorems, you have to, however, observe the |
131 \<dots>}\<close>}). If you deal with theorems, you have to, however, observe the |
184 text \<open> |
184 text \<open> |
185 The antiquotation \<open>@{prop \<dots>}\<close> constructs terms by inserting the |
185 The antiquotation \<open>@{prop \<dots>}\<close> constructs terms by inserting the |
186 usually invisible \<open>Trueprop\<close>-coercions whenever necessary. |
186 usually invisible \<open>Trueprop\<close>-coercions whenever necessary. |
187 Consider for example the pairs |
187 Consider for example the pairs |
188 |
188 |
189 @{ML_matchresult [display,gray] "(@{term \"P x\"}, @{prop \"P x\"})" |
189 @{ML_matchresult [display,gray] \<open>(@{term "P x"}, @{prop "P x"})\<close> |
190 "(Free (\"P\", _) $ Free (\"x\", _), |
190 \<open>(Free ("P", _) $ Free ("x", _), |
191 Const (\"HOL.Trueprop\", _) $ (Free (\"P\", _) $ Free (\"x\", _)))"} |
191 Const ("HOL.Trueprop", _) $ (Free ("P", _) $ Free ("x", _)))\<close>} |
192 |
192 |
193 where a coercion is inserted in the second component and |
193 where a coercion is inserted in the second component and |
194 |
194 |
195 @{ML_matchresult [display,gray] "(@{term \"P x \<Longrightarrow> Q x\"}, @{prop \"P x \<Longrightarrow> Q x\"})" |
195 @{ML_matchresult [display,gray] \<open>(@{term "P x \<Longrightarrow> Q x"}, @{prop "P x \<Longrightarrow> Q x"})\<close> |
196 "(Const (\"Pure.imp\", _) $ _ $ _, |
196 \<open>(Const ("Pure.imp", _) $ _ $ _, |
197 Const (\"Pure.imp\", _) $ _ $ _)"} |
197 Const ("Pure.imp", _) $ _ $ _)\<close>} |
198 |
198 |
199 where it is not (since it is already constructed by a meta-implication). |
199 where it is not (since it is already constructed by a meta-implication). |
200 The purpose of the \<open>Trueprop\<close>-coercion is to embed formulae of |
200 The purpose of the \<open>Trueprop\<close>-coercion is to embed formulae of |
201 an object logic, for example HOL, into the meta-logic of Isabelle. The coercion |
201 an object logic, for example HOL, into the meta-logic of Isabelle. The coercion |
202 is needed whenever a term is constructed that will be proved as a theorem. |
202 is needed whenever a term is constructed that will be proved as a theorem. |
203 |
203 |
204 As already seen above, types can be constructed using the antiquotation |
204 As already seen above, types can be constructed using the antiquotation |
205 \<open>@{typ _}\<close>. For example: |
205 \<open>@{typ _}\<close>. For example: |
206 |
206 |
207 @{ML_matchresult_fake [display,gray] "@{typ \"bool \<Rightarrow> nat\"}" "bool \<Rightarrow> nat"} |
207 @{ML_matchresult_fake [display,gray] \<open>@{typ "bool \<Rightarrow> nat"}\<close> \<open>bool \<Rightarrow> nat\<close>} |
208 |
208 |
209 The corresponding datatype is |
209 The corresponding datatype is |
210 \<close> |
210 \<close> |
211 |
211 |
212 ML_val %grayML\<open>datatype typ = |
212 ML_val %grayML\<open>datatype typ = |
214 | TFree of string * sort |
214 | TFree of string * sort |
215 | TVar of indexname * sort\<close> |
215 | TVar of indexname * sort\<close> |
216 |
216 |
217 text \<open> |
217 text \<open> |
218 Like with terms, there is the distinction between free type |
218 Like with terms, there is the distinction between free type |
219 variables (term-constructor @{ML "TFree"}) and schematic |
219 variables (term-constructor @{ML \<open>TFree\<close>}) and schematic |
220 type variables (term-constructor @{ML "TVar"} and printed with |
220 type variables (term-constructor @{ML \<open>TVar\<close>} and printed with |
221 a leading question mark). A type constant, |
221 a leading question mark). A type constant, |
222 like @{typ "int"} or @{typ bool}, are types with an empty list |
222 like @{typ "int"} or @{typ bool}, are types with an empty list |
223 of argument types. However, it needs a bit of effort to show an |
223 of argument types. However, it needs a bit of effort to show an |
224 example, because Isabelle always pretty prints types (unlike terms). |
224 example, because Isabelle always pretty prints types (unlike terms). |
225 Using just the antiquotation \<open>@{typ "bool"}\<close> we only see |
225 Using just the antiquotation \<open>@{typ "bool"}\<close> we only see |
226 |
226 |
227 @{ML_matchresult [display, gray] |
227 @{ML_matchresult [display, gray] |
228 "@{typ \"bool\"}" |
228 \<open>@{typ "bool"}\<close> |
229 "bool"} |
229 \<open>bool\<close>} |
230 which is the pretty printed version of \<open>bool\<close>. However, in PolyML |
230 which is the pretty printed version of \<open>bool\<close>. However, in PolyML |
231 (version \<open>\<ge>\<close>5.3) it is easy to install your own pretty printer. With the |
231 (version \<open>\<ge>\<close>5.3) it is easy to install your own pretty printer. With the |
232 function below we mimic the behaviour of the usual pretty printer for |
232 function below we mimic the behaviour of the usual pretty printer for |
233 datatypes (it uses pretty-printing functions which will be explained in more |
233 datatypes (it uses pretty-printing functions which will be explained in more |
234 detail in Section~\ref{sec:pretty}). |
234 detail in Section~\ref{sec:pretty}). |
262 ML \<open>@{typ "bool"}\<close> |
262 ML \<open>@{typ "bool"}\<close> |
263 text \<open> |
263 text \<open> |
264 Now the type bool is printed out in full detail. |
264 Now the type bool is printed out in full detail. |
265 |
265 |
266 @{ML_matchresult [display,gray] |
266 @{ML_matchresult [display,gray] |
267 "@{typ \"bool\"}" |
267 \<open>@{typ "bool"}\<close> |
268 "Type (\"HOL.bool\", [])"} |
268 \<open>Type ("HOL.bool", [])\<close>} |
269 |
269 |
270 When printing out a list-type |
270 When printing out a list-type |
271 |
271 |
272 @{ML_matchresult [display,gray] |
272 @{ML_matchresult [display,gray] |
273 "@{typ \"'a list\"}" |
273 \<open>@{typ "'a list"}\<close> |
274 "Type (\"List.list\", [TFree (\"'a\", [\"HOL.type\"])])"} |
274 \<open>Type ("List.list", [TFree ("'a", ["HOL.type"])])\<close>} |
275 |
275 |
276 we can see the full name of the type is actually \<open>List.list\<close>, indicating |
276 we can see the full name of the type is actually \<open>List.list\<close>, indicating |
277 that it is defined in the theory \<open>List\<close>. However, one has to be |
277 that it is defined in the theory \<open>List\<close>. However, one has to be |
278 careful with names of types, because even if |
278 careful with names of types, because even if |
279 \<open>fun\<close> is defined in the theory \<open>HOL\<close>, it is |
279 \<open>fun\<close> is defined in the theory \<open>HOL\<close>, it is |
280 still represented by its simple name. |
280 still represented by its simple name. |
281 |
281 |
282 @{ML_matchresult [display,gray] |
282 @{ML_matchresult [display,gray] |
283 "@{typ \"bool \<Rightarrow> nat\"}" |
283 \<open>@{typ "bool \<Rightarrow> nat"}\<close> |
284 "Type (\"fun\", [Type (\"HOL.bool\", []), Type (\"Nat.nat\", [])])"} |
284 \<open>Type ("fun", [Type ("HOL.bool", []), Type ("Nat.nat", [])])\<close>} |
285 |
285 |
286 We can restore the usual behaviour of Isabelle's pretty printer |
286 We can restore the usual behaviour of Isabelle's pretty printer |
287 with the code |
287 with the code |
288 \<close> |
288 \<close> |
289 |
289 |
336 text \<open> |
336 text \<open> |
337 To see this, apply \<open>@{term S}\<close> and \<open>@{term T}\<close> |
337 To see this, apply \<open>@{term S}\<close> and \<open>@{term T}\<close> |
338 to both functions. With @{ML make_imp} you obtain the intended term involving |
338 to both functions. With @{ML make_imp} you obtain the intended term involving |
339 the given arguments |
339 the given arguments |
340 |
340 |
341 @{ML_matchresult [display,gray] "make_imp @{term S} @{term T}" |
341 @{ML_matchresult [display,gray] \<open>make_imp @{term S} @{term T}\<close> |
342 "Const _ $ |
342 \<open>Const _ $ |
343 Abs (\"x\", Type (\"Nat.nat\",[]), |
343 Abs ("x", Type ("Nat.nat",[]), |
344 Const _ $ (Free (\"S\",_) $ _) $ (Free (\"T\",_) $ _))"} |
344 Const _ $ (Free ("S",_) $ _) $ (Free ("T",_) $ _))\<close>} |
345 |
345 |
346 whereas with @{ML make_wrong_imp} you obtain a term involving the @{term "P"} |
346 whereas with @{ML make_wrong_imp} you obtain a term involving the @{term "P"} |
347 and \<open>Q\<close> from the antiquotation. |
347 and \<open>Q\<close> from the antiquotation. |
348 |
348 |
349 @{ML_matchresult [display,gray] "make_wrong_imp @{term S} @{term T}" |
349 @{ML_matchresult [display,gray] \<open>make_wrong_imp @{term S} @{term T}\<close> |
350 "Const _ $ |
350 \<open>Const _ $ |
351 Abs (\"x\", _, |
351 Abs ("x", _, |
352 Const _ $ (Const _ $ (Free (\"P\",_) $ _)) $ |
352 Const _ $ (Const _ $ (Free ("P",_) $ _)) $ |
353 (Const _ $ (Free (\"Q\",_) $ _)))"} |
353 (Const _ $ (Free ("Q",_) $ _)))\<close>} |
354 |
354 |
355 There are a number of handy functions that are frequently used for |
355 There are a number of handy functions that are frequently used for |
356 constructing terms. One is the function @{ML_ind list_comb in Term}, which |
356 constructing terms. One is the function @{ML_ind list_comb in Term}, which |
357 takes as argument a term and a list of terms, and produces as output the |
357 takes as argument a term and a list of terms, and produces as output the |
358 term list applied to the term. For example |
358 term list applied to the term. For example |
359 |
359 |
360 |
360 |
361 @{ML_matchresult_fake [display,gray] |
361 @{ML_matchresult_fake [display,gray] |
362 "let |
362 \<open>let |
363 val trm = @{term \"P::bool \<Rightarrow> bool \<Rightarrow> bool\"} |
363 val trm = @{term "P::bool \<Rightarrow> bool \<Rightarrow> bool"} |
364 val args = [@{term \"True\"}, @{term \"False\"}] |
364 val args = [@{term "True"}, @{term "False"}] |
365 in |
365 in |
366 list_comb (trm, args) |
366 list_comb (trm, args) |
367 end" |
367 end\<close> |
368 "Free (\"P\", \"bool \<Rightarrow> bool \<Rightarrow> bool\") |
368 \<open>Free ("P", "bool \<Rightarrow> bool \<Rightarrow> bool") |
369 $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"} |
369 $ Const ("True", "bool") $ Const ("False", "bool")\<close>} |
370 |
370 |
371 Another handy function is @{ML_ind lambda in Term}, which abstracts a variable |
371 Another handy function is @{ML_ind lambda in Term}, which abstracts a variable |
372 in a term. For example |
372 in a term. For example |
373 |
373 |
374 @{ML_matchresult_fake [display,gray] |
374 @{ML_matchresult_fake [display,gray] |
375 "let |
375 \<open>let |
376 val x_nat = @{term \"x::nat\"} |
376 val x_nat = @{term "x::nat"} |
377 val trm = @{term \"(P::nat \<Rightarrow> bool) x\"} |
377 val trm = @{term "(P::nat \<Rightarrow> bool) x"} |
378 in |
378 in |
379 lambda x_nat trm |
379 lambda x_nat trm |
380 end" |
380 end\<close> |
381 "Abs (\"x\", \"Nat.nat\", Free (\"P\", \"bool \<Rightarrow> bool\") $ Bound 0)"} |
381 \<open>Abs ("x", "Nat.nat", Free ("P", "bool \<Rightarrow> bool") $ Bound 0)\<close>} |
382 |
382 |
383 In this example, @{ML lambda} produces a de Bruijn index (i.e.\ @{ML "Bound 0"}), |
383 In this example, @{ML lambda} produces a de Bruijn index (i.e.\ @{ML \<open>Bound 0\<close>}), |
384 and an abstraction, where it also records the type of the abstracted |
384 and an abstraction, where it also records the type of the abstracted |
385 variable and for printing purposes also its name. Note that because of the |
385 variable and for printing purposes also its name. Note that because of the |
386 typing annotation on \<open>P\<close>, the variable \<open>x\<close> in \<open>P x\<close> |
386 typing annotation on \<open>P\<close>, the variable \<open>x\<close> in \<open>P x\<close> |
387 is of the same type as the abstracted variable. If it is of different type, |
387 is of the same type as the abstracted variable. If it is of different type, |
388 as in |
388 as in |
389 |
389 |
390 @{ML_matchresult_fake [display,gray] |
390 @{ML_matchresult_fake [display,gray] |
391 "let |
391 \<open>let |
392 val x_int = @{term \"x::int\"} |
392 val x_int = @{term "x::int"} |
393 val trm = @{term \"(P::nat \<Rightarrow> bool) x\"} |
393 val trm = @{term "(P::nat \<Rightarrow> bool) x"} |
394 in |
394 in |
395 lambda x_int trm |
395 lambda x_int trm |
396 end" |
396 end\<close> |
397 "Abs (\"x\", \"int\", Free (\"P\", \"nat \<Rightarrow> bool\") $ Free (\"x\", \"nat\"))"} |
397 \<open>Abs ("x", "int", Free ("P", "nat \<Rightarrow> bool") $ Free ("x", "nat"))\<close>} |
398 |
398 |
399 then the variable \<open>Free ("x", "nat")\<close> is \emph{not} abstracted. |
399 then the variable \<open>Free ("x", "nat")\<close> is \emph{not} abstracted. |
400 This is a fundamental principle |
400 This is a fundamental principle |
401 of Church-style typing, where variables with the same name still differ, if they |
401 of Church-style typing, where variables with the same name still differ, if they |
402 have different type. |
402 have different type. |
405 replaced by other terms. For example below, we will replace in @{term |
405 replaced by other terms. For example below, we will replace in @{term |
406 "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by |
406 "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by |
407 @{term y}, and @{term x} by @{term True}. |
407 @{term y}, and @{term x} by @{term True}. |
408 |
408 |
409 @{ML_matchresult_fake [display,gray] |
409 @{ML_matchresult_fake [display,gray] |
410 "let |
410 \<open>let |
411 val sub1 = (@{term \"(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0\"}, @{term \"y::nat \<Rightarrow> nat\"}) |
411 val sub1 = (@{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"}, @{term "y::nat \<Rightarrow> nat"}) |
412 val sub2 = (@{term \"x::nat\"}, @{term \"True\"}) |
412 val sub2 = (@{term "x::nat"}, @{term "True"}) |
413 val trm = @{term \"((f::nat \<Rightarrow> nat \<Rightarrow> nat) 0) x\"} |
413 val trm = @{term "((f::nat \<Rightarrow> nat \<Rightarrow> nat) 0) x"} |
414 in |
414 in |
415 subst_free [sub1, sub2] trm |
415 subst_free [sub1, sub2] trm |
416 end" |
416 end\<close> |
417 "Free (\"y\", \"nat \<Rightarrow> nat\") $ Const (\"True\", \"bool\")"} |
417 \<open>Free ("y", "nat \<Rightarrow> nat") $ Const ("True", "bool")\<close>} |
418 |
418 |
419 As can be seen, @{ML subst_free} does not take typability into account. |
419 As can be seen, @{ML subst_free} does not take typability into account. |
420 However it takes alpha-equivalence into account: |
420 However it takes alpha-equivalence into account: |
421 |
421 |
422 @{ML_matchresult_fake [display, gray] |
422 @{ML_matchresult_fake [display, gray] |
423 "let |
423 \<open>let |
424 val sub = (@{term \"(\<lambda>y::nat. y)\"}, @{term \"x::nat\"}) |
424 val sub = (@{term "(\<lambda>y::nat. y)"}, @{term "x::nat"}) |
425 val trm = @{term \"(\<lambda>x::nat. x)\"} |
425 val trm = @{term "(\<lambda>x::nat. x)"} |
426 in |
426 in |
427 subst_free [sub] trm |
427 subst_free [sub] trm |
428 end" |
428 end\<close> |
429 "Free (\"x\", \"nat\")"} |
429 \<open>Free ("x", "nat")\<close>} |
430 |
430 |
431 Similarly the function @{ML_ind subst_bounds in Term}, replaces lose bound |
431 Similarly the function @{ML_ind subst_bounds in Term}, replaces lose bound |
432 variables with terms. To see how this function works, let us implement a |
432 variables with terms. To see how this function works, let us implement a |
433 function that strips off the outermost forall quantifiers in a term. |
433 function that strips off the outermost forall quantifiers in a term. |
434 \<close> |
434 \<close> |
465 As said above, after calling @{ML strip_alls}, you obtain a term with loose |
465 As said above, after calling @{ML strip_alls}, you obtain a term with loose |
466 bound variables. With the function @{ML subst_bounds}, you can replace these |
466 bound variables. With the function @{ML subst_bounds}, you can replace these |
467 loose @{ML_ind Bound in Term}s with the stripped off variables. |
467 loose @{ML_ind Bound in Term}s with the stripped off variables. |
468 |
468 |
469 @{ML_matchresult_fake [display, gray, linenos] |
469 @{ML_matchresult_fake [display, gray, linenos] |
470 "let |
470 \<open>let |
471 val (vrs, trm) = strip_alls @{term \"\<forall>x y. x = (y::bool)\"} |
471 val (vrs, trm) = strip_alls @{term "\<forall>x y. x = (y::bool)"} |
472 in |
472 in |
473 subst_bounds (rev vrs, trm) |
473 subst_bounds (rev vrs, trm) |
474 |> pretty_term @{context} |
474 |> pretty_term @{context} |
475 |> pwriteln |
475 |> pwriteln |
476 end" |
476 end\<close> |
477 "x = y"} |
477 \<open>x = y\<close>} |
478 |
478 |
479 Note that in Line 4 we had to reverse the list of variables that @{ML |
479 Note that in Line 4 we had to reverse the list of variables that @{ML |
480 strip_alls} returned. The reason is that the head of the list the function |
480 strip_alls} returned. The reason is that the head of the list the function |
481 @{ML subst_bounds} takes is the replacement for @{ML "Bound 0"}, the next |
481 @{ML subst_bounds} takes is the replacement for @{ML \<open>Bound 0\<close>}, the next |
482 element for @{ML "Bound 1"} and so on. |
482 element for @{ML \<open>Bound 1\<close>} and so on. |
483 |
483 |
484 Notice also that this function might introduce name clashes, since we |
484 Notice also that this function might introduce name clashes, since we |
485 substitute just a variable with the name recorded in an abstraction. This |
485 substitute just a variable with the name recorded in an abstraction. This |
486 name is by no means unique. If clashes need to be avoided, then we should |
486 name is by no means unique. If clashes need to be avoided, then we should |
487 use the function @{ML_ind dest_abs in Term}, which returns the body where |
487 use the function @{ML_ind dest_abs in Term}, which returns the body where |
488 the loose de Bruijn index is replaced by a unique free variable. For example |
488 the loose de Bruijn index is replaced by a unique free variable. For example |
489 |
489 |
490 |
490 |
491 @{ML_matchresult_fake [display,gray] |
491 @{ML_matchresult_fake [display,gray] |
492 "let |
492 \<open>let |
493 val body = Bound 0 $ Free (\"x\", @{typ nat}) |
493 val body = Bound 0 $ Free ("x", @{typ nat}) |
494 in |
494 in |
495 Term.dest_abs (\"x\", @{typ \"nat \<Rightarrow> bool\"}, body) |
495 Term.dest_abs ("x", @{typ "nat \<Rightarrow> bool"}, body) |
496 end" |
496 end\<close> |
497 "(\"xa\", Free (\"xa\", \"Nat.nat \<Rightarrow> bool\") $ Free (\"x\", \"Nat.nat\"))"} |
497 \<open>("xa", Free ("xa", "Nat.nat \<Rightarrow> bool") $ Free ("x", "Nat.nat"))\<close>} |
498 |
498 |
499 Sometimes it is necessary to manipulate de Bruijn indices in terms directly. |
499 Sometimes it is necessary to manipulate de Bruijn indices in terms directly. |
500 There are many functions to do this. We describe only two. The first, |
500 There are many functions to do this. We describe only two. The first, |
501 @{ML_ind incr_boundvars in Term}, increases by an integer the indices |
501 @{ML_ind incr_boundvars in Term}, increases by an integer the indices |
502 of the loose bound variables in a term. In the code below |
502 of the loose bound variables in a term. In the code below |
503 |
503 |
504 @{ML_matchresult_fake [display,gray] |
504 @{ML_matchresult_fake [display,gray] |
505 "@{term \"\<forall>x y z u. z = u\"} |
505 \<open>@{term "\<forall>x y z u. z = u"} |
506 |> strip_alls |
506 |> strip_alls |
507 ||> incr_boundvars 2 |
507 ||> incr_boundvars 2 |
508 |> build_alls |
508 |> build_alls |
509 |> pretty_term @{context} |
509 |> pretty_term @{context} |
510 |> pwriteln" |
510 |> pwriteln\<close> |
511 "\<forall>x y z u. x = y"} |
511 \<open>\<forall>x y z u. x = y\<close>} |
512 |
512 |
513 we first strip off the forall-quantified variables (thus creating two loose |
513 we first strip off the forall-quantified variables (thus creating two loose |
514 bound variables in the body); then we increase the indices of the loose |
514 bound variables in the body); then we increase the indices of the loose |
515 bound variables by @{ML 2} and finally re-quantify the variables. As a |
515 bound variables by @{ML 2} and finally re-quantify the variables. As a |
516 result of @{ML incr_boundvars}, we obtain now a term that has the equation |
516 result of @{ML incr_boundvars}, we obtain now a term that has the equation |
518 |
518 |
519 The second function, @{ML_ind loose_bvar1 in Text}, tests whether a term |
519 The second function, @{ML_ind loose_bvar1 in Text}, tests whether a term |
520 contains a loose bound of a certain index. For example |
520 contains a loose bound of a certain index. For example |
521 |
521 |
522 @{ML_matchresult_fake [gray,display] |
522 @{ML_matchresult_fake [gray,display] |
523 "let |
523 \<open>let |
524 val body = snd (strip_alls @{term \"\<forall>x y. x = (y::bool)\"}) |
524 val body = snd (strip_alls @{term "\<forall>x y. x = (y::bool)"}) |
525 in |
525 in |
526 [loose_bvar1 (body, 0), |
526 [loose_bvar1 (body, 0), |
527 loose_bvar1 (body, 1), |
527 loose_bvar1 (body, 1), |
528 loose_bvar1 (body, 2)] |
528 loose_bvar1 (body, 2)] |
529 end" |
529 end\<close> |
530 "[true, true, false]"} |
530 \<open>[true, true, false]\<close>} |
531 |
531 |
532 There are also many convenient functions that construct specific HOL-terms |
532 There are also many convenient functions that construct specific HOL-terms |
533 in the structure @{ML_structure HOLogic}. For example @{ML_ind mk_eq in |
533 in the structure @{ML_structure HOLogic}. For example @{ML_ind mk_eq in |
534 HOLogic} constructs an equality out of two terms. The types needed in this |
534 HOLogic} constructs an equality out of two terms. The types needed in this |
535 equality are calculated from the type of the arguments. For example |
535 equality are calculated from the type of the arguments. For example |
536 |
536 |
537 @{ML_matchresult_fake [gray,display] |
537 @{ML_matchresult_fake [gray,display] |
538 "let |
538 \<open>let |
539 val eq = HOLogic.mk_eq (@{term \"True\"}, @{term \"False\"}) |
539 val eq = HOLogic.mk_eq (@{term "True"}, @{term "False"}) |
540 in |
540 in |
541 eq |> pretty_term @{context} |
541 eq |> pretty_term @{context} |
542 |> pwriteln |
542 |> pwriteln |
543 end" |
543 end\<close> |
544 "True = False"} |
544 \<open>True = False\<close>} |
545 |
545 |
546 \begin{readmore} |
546 \begin{readmore} |
547 There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file |
547 There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file |
548 "Pure/logic.ML"} and @{ML_file "HOL/Tools/hologic.ML"} that make manual |
548 "Pure/logic.ML"} and @{ML_file "HOL/Tools/hologic.ML"} that make manual |
549 constructions of terms and types easier. |
549 constructions of terms and types easier. |
596 | is_nil _ = false\<close> |
596 | is_nil _ = false\<close> |
597 |
597 |
598 text \<open> |
598 text \<open> |
599 Unfortunately, also this function does \emph{not} work as expected, since |
599 Unfortunately, also this function does \emph{not} work as expected, since |
600 |
600 |
601 @{ML_matchresult [display,gray] "is_nil @{term \"Nil\"}" "false"} |
601 @{ML_matchresult [display,gray] \<open>is_nil @{term "Nil"}\<close> \<open>false\<close>} |
602 |
602 |
603 The problem is that on the ML-level the name of a constant is more |
603 The problem is that on the ML-level the name of a constant is more |
604 subtle than you might expect. The function @{ML is_all} worked correctly, |
604 subtle than you might expect. The function @{ML is_all} worked correctly, |
605 because @{term "All"} is such a fundamental constant, which can be referenced |
605 because @{term "All"} is such a fundamental constant, which can be referenced |
606 by @{ML "Const (\"All\", some_type)" for some_type}. However, if you look at |
606 by @{ML \<open>Const ("All", some_type)\<close> for some_type}. However, if you look at |
607 |
607 |
608 @{ML_matchresult [display,gray] "@{term \"Nil\"}" "Const (\"List.list.Nil\", _)"} |
608 @{ML_matchresult [display,gray] \<open>@{term "Nil"}\<close> \<open>Const ("List.list.Nil", _)\<close>} |
609 |
609 |
610 the name of the constant \<open>Nil\<close> depends on the theory in which the |
610 the name of the constant \<open>Nil\<close> depends on the theory in which the |
611 term constructor is defined (\<open>List\<close>) and also in which datatype |
611 term constructor is defined (\<open>List\<close>) and also in which datatype |
612 (\<open>list\<close>). Even worse, some constants have a name involving |
612 (\<open>list\<close>). Even worse, some constants have a name involving |
613 type-classes. Consider for example the constants for @{term "zero"} and |
613 type-classes. Consider for example the constants for @{term "zero"} and |
614 \mbox{@{term "times"}}: |
614 \mbox{@{term "times"}}: |
615 |
615 |
616 @{ML_matchresult [display,gray] "(@{term \"0::nat\"}, @{term \"times\"})" |
616 @{ML_matchresult [display,gray] \<open>(@{term "0::nat"}, @{term "times"})\<close> |
617 "(Const (\"Groups.zero_class.zero\", _), |
617 \<open>(Const ("Groups.zero_class.zero", _), |
618 Const (\"Groups.times_class.times\", _))"} |
618 Const ("Groups.times_class.times", _))\<close>} |
619 |
619 |
620 While you could use the complete name, for example |
620 While you could use the complete name, for example |
621 @{ML "Const (\"List.list.Nil\", some_type)" for some_type}, for referring to or |
621 @{ML \<open>Const ("List.list.Nil", some_type)\<close> for some_type}, for referring to or |
622 matching against \<open>Nil\<close>, this would make the code rather brittle. |
622 matching against \<open>Nil\<close>, this would make the code rather brittle. |
623 The reason is that the theory and the name of the datatype can easily change. |
623 The reason is that the theory and the name of the datatype can easily change. |
624 To make the code more robust, it is better to use the antiquotation |
624 To make the code more robust, it is better to use the antiquotation |
625 \<open>@{const_name \<dots>}\<close>. With this antiquotation you can harness the |
625 \<open>@{const_name \<dots>}\<close>. With this antiquotation you can harness the |
626 variable parts of the constant's name. Therefore a function for |
626 variable parts of the constant's name. Therefore a function for |
881 To print out the result type environment we switch on the printing |
881 To print out the result type environment we switch on the printing |
882 of sort information by setting @{ML_ind show_sorts in Syntax} to |
882 of sort information by setting @{ML_ind show_sorts in Syntax} to |
883 true. This allows us to inspect the typing environment. |
883 true. This allows us to inspect the typing environment. |
884 |
884 |
885 @{ML_matchresult_fake [display,gray] |
885 @{ML_matchresult_fake [display,gray] |
886 "pretty_tyenv @{context} tyenv" |
886 \<open>pretty_tyenv @{context} tyenv\<close> |
887 "[?'a::s1 := ?'a1::{s1, s2}, ?'b::s2 := ?'a1::{s1, s2}]"} |
887 \<open>[?'a::s1 := ?'a1::{s1, s2}, ?'b::s2 := ?'a1::{s1, s2}]\<close>} |
888 |
888 |
889 As can be seen, the type variables \<open>?'a\<close> and \<open>?'b\<close> are instantiated |
889 As can be seen, the type variables \<open>?'a\<close> and \<open>?'b\<close> are instantiated |
890 with a new type variable \<open>?'a1\<close> with sort \<open>{s1, s2}\<close>. Since a new |
890 with a new type variable \<open>?'a1\<close> with sort \<open>{s1, s2}\<close>. Since a new |
891 type variable has been introduced the @{ML index}, originally being \<open>0\<close>, |
891 type variable has been introduced the @{ML index}, originally being \<open>0\<close>, |
892 has been increased to \<open>1\<close>. |
892 has been increased to \<open>1\<close>. |
893 |
893 |
894 @{ML_matchresult [display,gray] |
894 @{ML_matchresult [display,gray] |
895 "index" |
895 \<open>index\<close> |
896 "1"} |
896 \<open>1\<close>} |
897 |
897 |
898 Let us now return to the unification problem \<open>?'a * ?'b\<close> and |
898 Let us now return to the unification problem \<open>?'a * ?'b\<close> and |
899 \<open>?'b list * nat\<close> from the beginning of this section, and the |
899 \<open>?'b list * nat\<close> from the beginning of this section, and the |
900 calculated type environment @{ML tyenv_unif}: |
900 calculated type environment @{ML tyenv_unif}: |
901 |
901 |
902 @{ML_matchresult_fake [display, gray] |
902 @{ML_matchresult_fake [display, gray] |
903 "pretty_tyenv @{context} tyenv_unif" |
903 \<open>pretty_tyenv @{context} tyenv_unif\<close> |
904 "[?'a := ?'b list, ?'b := nat]"} |
904 \<open>[?'a := ?'b list, ?'b := nat]\<close>} |
905 |
905 |
906 Observe that the type environment which the function @{ML typ_unify in |
906 Observe that the type environment which the function @{ML typ_unify in |
907 Sign} returns is \emph{not} an instantiation in fully solved form: while \<open>?'b\<close> is instantiated to @{typ nat}, this is not propagated to the |
907 Sign} returns is \emph{not} an instantiation in fully solved form: while \<open>?'b\<close> is instantiated to @{typ nat}, this is not propagated to the |
908 instantiation for \<open>?'a\<close>. In unification theory, this is often |
908 instantiation for \<open>?'a\<close>. In unification theory, this is often |
909 called an instantiation in \emph{triangular form}. These triangular |
909 called an instantiation in \emph{triangular form}. These triangular |
910 instantiations, or triangular type environments, are used because of |
910 instantiations, or triangular type environments, are used because of |
911 performance reasons. To apply such a type environment to a type, say \<open>?'a * |
911 performance reasons. To apply such a type environment to a type, say \<open>?'a * |
912 ?'b\<close>, you should use the function @{ML_ind norm_type in Envir}: |
912 ?'b\<close>, you should use the function @{ML_ind norm_type in Envir}: |
913 |
913 |
914 @{ML_matchresult_fake [display, gray] |
914 @{ML_matchresult_fake [display, gray] |
915 "Envir.norm_type tyenv_unif @{typ_pat \"?'a * ?'b\"}" |
915 \<open>Envir.norm_type tyenv_unif @{typ_pat "?'a * ?'b"}\<close> |
916 "nat list * nat"} |
916 \<open>nat list * nat\<close>} |
917 |
917 |
918 Matching of types can be done with the function @{ML_ind typ_match in Sign} |
918 Matching of types can be done with the function @{ML_ind typ_match in Sign} |
919 also from the structure @{ML_structure Sign}. This function returns a @{ML_type |
919 also from the structure @{ML_structure Sign}. This function returns a @{ML_type |
920 Type.tyenv} as well, but might raise the exception \<open>TYPE_MATCH\<close> in case |
920 Type.tyenv} as well, but might raise the exception \<open>TYPE_MATCH\<close> in case |
921 of failure. For example |
921 of failure. For example |
1022 In this example we annotated types explicitly because then |
1022 In this example we annotated types explicitly because then |
1023 the type environment is empty and can be ignored. The |
1023 the type environment is empty and can be ignored. The |
1024 resulting term environment is |
1024 resulting term environment is |
1025 |
1025 |
1026 @{ML_matchresult_fake [display, gray] |
1026 @{ML_matchresult_fake [display, gray] |
1027 "pretty_env @{context} fo_env" |
1027 \<open>pretty_env @{context} fo_env\<close> |
1028 "[?X := P, ?Y := \<lambda>a b. Q b a]"} |
1028 \<open>[?X := P, ?Y := \<lambda>a b. Q b a]\<close>} |
1029 |
1029 |
1030 The matcher can be applied to a term using the function @{ML_ind subst_term |
1030 The matcher can be applied to a term using the function @{ML_ind subst_term |
1031 in Envir} (remember the same convention for types applies to terms: @{ML |
1031 in Envir} (remember the same convention for types applies to terms: @{ML |
1032 subst_term in Envir} is for matchers and @{ML norm_term in Envir} for |
1032 subst_term in Envir} is for matchers and @{ML norm_term in Envir} for |
1033 unifiers). The function @{ML subst_term in Envir} expects a type environment, |
1033 unifiers). The function @{ML subst_term in Envir} expects a type environment, |
1034 which is set to empty in the example below, and a term environment. |
1034 which is set to empty in the example below, and a term environment. |
1035 |
1035 |
1036 @{ML_matchresult_fake [display, gray] |
1036 @{ML_matchresult_fake [display, gray] |
1037 "let |
1037 \<open>let |
1038 val trm = @{term_pat \"(?X::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool) ?Y\"} |
1038 val trm = @{term_pat "(?X::(nat\<Rightarrow>nat\<Rightarrow>nat)\<Rightarrow>bool) ?Y"} |
1039 in |
1039 in |
1040 Envir.subst_term (Vartab.empty, fo_env) trm |
1040 Envir.subst_term (Vartab.empty, fo_env) trm |
1041 |> pretty_term @{context} |
1041 |> pretty_term @{context} |
1042 |> pwriteln |
1042 |> pwriteln |
1043 end" |
1043 end\<close> |
1044 "P (\<lambda>a b. Q b a)"} |
1044 \<open>P (\<lambda>a b. Q b a)\<close>} |
1045 |
1045 |
1046 First-order matching is useful for matching against applications and |
1046 First-order matching is useful for matching against applications and |
1047 variables. It can also deal with abstractions and a limited form of |
1047 variables. It can also deal with abstractions and a limited form of |
1048 alpha-equivalence, but this kind of matching should be used with care, since |
1048 alpha-equivalence, but this kind of matching should be used with care, since |
1049 it is not clear whether the result is meaningful. A meaningful example is |
1049 it is not clear whether the result is meaningful. A meaningful example is |
1050 matching \<open>\<lambda>x. P x\<close> against the pattern \<open>\<lambda>y. ?X y\<close>. In this |
1050 matching \<open>\<lambda>x. P x\<close> against the pattern \<open>\<lambda>y. ?X y\<close>. In this |
1051 case, first-order matching produces \<open>[?X := P]\<close>. |
1051 case, first-order matching produces \<open>[?X := P]\<close>. |
1052 |
1052 |
1053 @{ML_matchresult_fake [display, gray] |
1053 @{ML_matchresult_fake [display, gray] |
1054 "let |
1054 \<open>let |
1055 val fo_pat = @{term_pat \"\<lambda>y. (?X::nat\<Rightarrow>bool) y\"} |
1055 val fo_pat = @{term_pat "\<lambda>y. (?X::nat\<Rightarrow>bool) y"} |
1056 val trm = @{term \"\<lambda>x. (P::nat\<Rightarrow>bool) x\"} |
1056 val trm = @{term "\<lambda>x. (P::nat\<Rightarrow>bool) x"} |
1057 val init = (Vartab.empty, Vartab.empty) |
1057 val init = (Vartab.empty, Vartab.empty) |
1058 in |
1058 in |
1059 Pattern.first_order_match @{theory} (fo_pat, trm) init |
1059 Pattern.first_order_match @{theory} (fo_pat, trm) init |
1060 |> snd |
1060 |> snd |
1061 |> pretty_env @{context} |
1061 |> pretty_env @{context} |
1062 end" |
1062 end\<close> |
1063 "[?X := P]"} |
1063 \<open>[?X := P]\<close>} |
1064 \<close> |
1064 \<close> |
1065 |
1065 |
1066 text \<open> |
1066 text \<open> |
1067 Unification of abstractions is more thoroughly studied in the context of |
1067 Unification of abstractions is more thoroughly studied in the context of |
1068 higher-order pattern unification and higher-order pattern matching. A |
1068 higher-order pattern unification and higher-order pattern matching. A |
1074 @{ML_ind pattern in Pattern} in the structure @{ML_structure Pattern} tests |
1074 @{ML_ind pattern in Pattern} in the structure @{ML_structure Pattern} tests |
1075 whether a term satisfies these restrictions under the assumptions |
1075 whether a term satisfies these restrictions under the assumptions |
1076 that it is beta-normal, well-typed and has no loose bound variables. |
1076 that it is beta-normal, well-typed and has no loose bound variables. |
1077 |
1077 |
1078 @{ML_matchresult [display, gray] |
1078 @{ML_matchresult [display, gray] |
1079 "let |
1079 \<open>let |
1080 val trm_list = |
1080 val trm_list = |
1081 [@{term_pat \"?X\"}, @{term_pat \"a\"}, |
1081 [@{term_pat "?X"}, @{term_pat "a"}, |
1082 @{term_pat \"f (\<lambda>a b. ?X a b) c\"}, |
1082 @{term_pat "f (\<lambda>a b. ?X a b) c"}, |
1083 @{term_pat \"\<lambda>a b. (+) a b\"}, |
1083 @{term_pat "\<lambda>a b. (+) a b"}, |
1084 @{term_pat \"\<lambda>a. (+) a ?Y\"}, @{term_pat \"?X ?Y\"}, |
1084 @{term_pat "\<lambda>a. (+) a ?Y"}, @{term_pat "?X ?Y"}, |
1085 @{term_pat \"\<lambda>a b. ?X a b ?Y\"}, @{term_pat \"\<lambda>a. ?X a a\"}, |
1085 @{term_pat "\<lambda>a b. ?X a b ?Y"}, @{term_pat "\<lambda>a. ?X a a"}, |
1086 @{term_pat \"?X a\"}] |
1086 @{term_pat "?X a"}] |
1087 in |
1087 in |
1088 map Pattern.pattern trm_list |
1088 map Pattern.pattern trm_list |
1089 end" |
1089 end\<close> |
1090 "[true, true, true, true, true, false, false, false, false]"} |
1090 \<open>[true, true, true, true, true, false, false, false, false]\<close>} |
1091 |
1091 |
1092 The point of the restriction to patterns is that unification and matching |
1092 The point of the restriction to patterns is that unification and matching |
1093 are decidable and produce most general unifiers, respectively matchers. |
1093 are decidable and produce most general unifiers, respectively matchers. |
1094 Note that \emph{both} terms to be unified have to be higher-order patterns |
1094 Note that \emph{both} terms to be unified have to be higher-order patterns |
1095 for this to work. The exception @{ML_ind Pattern in Pattern} indicates failure |
1095 for this to work. The exception @{ML_ind Pattern in Pattern} indicates failure |
1151 \footnote{\bf FIXME: what is the list of term pairs in the unifier: flex-flex pairs?} |
1151 \footnote{\bf FIXME: what is the list of term pairs in the unifier: flex-flex pairs?} |
1152 |
1152 |
1153 We can print them out as follows. |
1153 We can print them out as follows. |
1154 |
1154 |
1155 @{ML_matchresult_fake [display, gray] |
1155 @{ML_matchresult_fake [display, gray] |
1156 "pretty_env @{context} (Envir.term_env un1); |
1156 \<open>pretty_env @{context} (Envir.term_env un1); |
1157 pretty_env @{context} (Envir.term_env un2); |
1157 pretty_env @{context} (Envir.term_env un2); |
1158 pretty_env @{context} (Envir.term_env un3)" |
1158 pretty_env @{context} (Envir.term_env un3)\<close> |
1159 "[?X := \<lambda>a. a, ?Y := f a] |
1159 \<open>[?X := \<lambda>a. a, ?Y := f a] |
1160 [?X := f, ?Y := a] |
1160 [?X := f, ?Y := a] |
1161 [?X := \<lambda>b. f a]"} |
1161 [?X := \<lambda>b. f a]\<close>} |
1162 |
1162 |
1163 |
1163 |
1164 In case of failure the function @{ML_ind unifiers in Unify} does not raise |
1164 In case of failure the function @{ML_ind unifiers in Unify} does not raise |
1165 an exception, rather returns the empty sequence. For example |
1165 an exception, rather returns the empty sequence. For example |
1166 |
1166 |
1167 @{ML_matchresult [display, gray] |
1167 @{ML_matchresult [display, gray] |
1168 "let |
1168 \<open>let |
1169 val trm1 = @{term \"a\"} |
1169 val trm1 = @{term "a"} |
1170 val trm2 = @{term \"b\"} |
1170 val trm2 = @{term "b"} |
1171 val init = Envir.empty 0 |
1171 val init = Envir.empty 0 |
1172 in |
1172 in |
1173 Unify.unifiers (Context.Proof @{context}, init, [(trm1, trm2)]) |
1173 Unify.unifiers (Context.Proof @{context}, init, [(trm1, trm2)]) |
1174 |> Seq.pull |
1174 |> Seq.pull |
1175 end" |
1175 end\<close> |
1176 "NONE"} |
1176 \<open>NONE\<close>} |
1177 |
1177 |
1178 In order to find a reasonable solution for a unification problem, Isabelle |
1178 In order to find a reasonable solution for a unification problem, Isabelle |
1179 also tries first to solve the problem by higher-order pattern |
1179 also tries first to solve the problem by higher-order pattern |
1180 unification. Only in case of failure full higher-order unification is |
1180 unification. Only in case of failure full higher-order unification is |
1181 called. This function has a built-in bound, which can be accessed and |
1181 called. This function has a built-in bound, which can be accessed and |
1182 manipulated as a configuration value. For example |
1182 manipulated as a configuration value. For example |
1183 |
1183 |
1184 @{ML_matchresult_fake [display,gray] |
1184 @{ML_matchresult_fake [display,gray] |
1185 "Config.get_global @{theory} (Unify.search_bound)" |
1185 \<open>Config.get_global @{theory} (Unify.search_bound)\<close> |
1186 "Int 60"} |
1186 \<open>Int 60\<close>} |
1187 |
1187 |
1188 If this bound is reached during unification, Isabelle prints out the |
1188 If this bound is reached during unification, Isabelle prints out the |
1189 warning message @{text [quotes] "Unification bound exceeded"} and |
1189 warning message @{text [quotes] "Unification bound exceeded"} and |
1190 plenty of diagnostic information (sometimes annoyingly plenty of |
1190 plenty of diagnostic information (sometimes annoyingly plenty of |
1191 information). |
1191 information). |
1353 information) so that it is always clear what the type of a term is. |
1353 information) so that it is always clear what the type of a term is. |
1354 Given a well-typed term, the function @{ML_ind type_of in Term} returns the |
1354 Given a well-typed term, the function @{ML_ind type_of in Term} returns the |
1355 type of a term. Consider for example: |
1355 type of a term. Consider for example: |
1356 |
1356 |
1357 @{ML_matchresult [display,gray] |
1357 @{ML_matchresult [display,gray] |
1358 "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"} |
1358 \<open>type_of (@{term "f::nat \<Rightarrow> bool"} $ @{term "x::nat"})\<close> \<open>bool\<close>} |
1359 |
1359 |
1360 To calculate the type, this function traverses the whole term and will |
1360 To calculate the type, this function traverses the whole term and will |
1361 detect any typing inconsistency. For example changing the type of the variable |
1361 detect any typing inconsistency. For example changing the type of the variable |
1362 @{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message: |
1362 @{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message: |
1363 |
1363 |
1364 @{ML_matchresult_fake [display,gray] |
1364 @{ML_matchresult_fake [display,gray] |
1365 "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" |
1365 \<open>type_of (@{term "f::nat \<Rightarrow> bool"} $ @{term "x::int"})\<close> |
1366 "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"} |
1366 \<open>*** Exception- TYPE ("type_of: type mismatch in application" \<dots>\<close>} |
1367 |
1367 |
1368 Since the complete traversal might sometimes be too costly and |
1368 Since the complete traversal might sometimes be too costly and |
1369 not necessary, there is the function @{ML_ind fastype_of in Term}, which |
1369 not necessary, there is the function @{ML_ind fastype_of in Term}, which |
1370 also returns the type of a term. |
1370 also returns the type of a term. |
1371 |
1371 |
1372 @{ML_matchresult [display,gray] |
1372 @{ML_matchresult [display,gray] |
1373 "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"} |
1373 \<open>fastype_of (@{term "f::nat \<Rightarrow> bool"} $ @{term "x::nat"})\<close> \<open>bool\<close>} |
1374 |
1374 |
1375 However, efficiency is gained on the expense of skipping some tests. You |
1375 However, efficiency is gained on the expense of skipping some tests. You |
1376 can see this in the following example |
1376 can see this in the following example |
1377 |
1377 |
1378 @{ML_matchresult [display,gray] |
1378 @{ML_matchresult [display,gray] |
1379 "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" "bool"} |
1379 \<open>fastype_of (@{term "f::nat \<Rightarrow> bool"} $ @{term "x::int"})\<close> \<open>bool\<close>} |
1380 |
1380 |
1381 where no error is detected. |
1381 where no error is detected. |
1382 |
1382 |
1383 Sometimes it is a bit inconvenient to construct a term with |
1383 Sometimes it is a bit inconvenient to construct a term with |
1384 complete typing annotations, especially in cases where the typing |
1384 complete typing annotations, especially in cases where the typing |
1431 |
1431 |
1432 Certification is always relative to a context. For example you can |
1432 Certification is always relative to a context. For example you can |
1433 write: |
1433 write: |
1434 |
1434 |
1435 @{ML_matchresult_fake [display,gray] |
1435 @{ML_matchresult_fake [display,gray] |
1436 "Thm.cterm_of @{context} @{term \"(a::nat) + b = c\"}" |
1436 \<open>Thm.cterm_of @{context} @{term "(a::nat) + b = c"}\<close> |
1437 "a + b = c"} |
1437 \<open>a + b = c\<close>} |
1438 |
1438 |
1439 This can also be written with an antiquotation: |
1439 This can also be written with an antiquotation: |
1440 |
1440 |
1441 @{ML_matchresult_fake [display,gray] |
1441 @{ML_matchresult_fake [display,gray] |
1442 "@{cterm \"(a::nat) + b = c\"}" |
1442 \<open>@{cterm "(a::nat) + b = c"}\<close> |
1443 "a + b = c"} |
1443 \<open>a + b = c\<close>} |
1444 |
1444 |
1445 Attempting to obtain the certified term for |
1445 Attempting to obtain the certified term for |
1446 |
1446 |
1447 @{ML_matchresult_fake_both [display,gray] |
1447 @{ML_matchresult_fake_both [display,gray] |
1448 "@{cterm \"1 + True\"}" |
1448 \<open>@{cterm "1 + True"}\<close> |
1449 "Type unification failed \<dots>"} |
1449 \<open>Type unification failed \<dots>\<close>} |
1450 |
1450 |
1451 yields an error (since the term is not typable). A slightly more elaborate |
1451 yields an error (since the term is not typable). A slightly more elaborate |
1452 example that type-checks is: |
1452 example that type-checks is: |
1453 |
1453 |
1454 @{ML_matchresult_fake [display,gray] |
1454 @{ML_matchresult_fake [display,gray] |
1455 "let |
1455 \<open>let |
1456 val natT = @{typ \"nat\"} |
1456 val natT = @{typ "nat"} |
1457 val zero = @{term \"0::nat\"} |
1457 val zero = @{term "0::nat"} |
1458 val plus = Const (@{const_name plus}, [natT, natT] ---> natT) |
1458 val plus = Const (@{const_name plus}, [natT, natT] ---> natT) |
1459 in |
1459 in |
1460 Thm.cterm_of @{context} (plus $ zero $ zero) |
1460 Thm.cterm_of @{context} (plus $ zero $ zero) |
1461 end" |
1461 end\<close> |
1462 "0 + 0"} |
1462 \<open>0 + 0\<close>} |
1463 |
1463 |
1464 In Isabelle not just terms need to be certified, but also types. For example, |
1464 In Isabelle not just terms need to be certified, but also types. For example, |
1465 you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on |
1465 you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on |
1466 the ML-level as follows: |
1466 the ML-level as follows: |
1467 |
1467 |
1468 @{ML_matchresult_fake [display,gray] |
1468 @{ML_matchresult_fake [display,gray] |
1469 "Thm.ctyp_of @{context} (@{typ nat} --> @{typ bool})" |
1469 \<open>Thm.ctyp_of @{context} (@{typ nat} --> @{typ bool})\<close> |
1470 "nat \<Rightarrow> bool"} |
1470 \<open>nat \<Rightarrow> bool\<close>} |
1471 |
1471 |
1472 or with the antiquotation: |
1472 or with the antiquotation: |
1473 |
1473 |
1474 @{ML_matchresult_fake [display,gray] |
1474 @{ML_matchresult_fake [display,gray] |
1475 "@{ctyp \"nat \<Rightarrow> bool\"}" |
1475 \<open>@{ctyp "nat \<Rightarrow> bool"}\<close> |
1476 "nat \<Rightarrow> bool"} |
1476 \<open>nat \<Rightarrow> bool\<close>} |
1477 |
1477 |
1478 Since certified terms are, unlike terms, abstract objects, we cannot |
1478 Since certified terms are, unlike terms, abstract objects, we cannot |
1479 pattern-match against them. However, we can construct them. For example |
1479 pattern-match against them. However, we can construct them. For example |
1480 the function @{ML_ind apply in Thm} produces a certified application. |
1480 the function @{ML_ind apply in Thm} produces a certified application. |
1481 |
1481 |
1482 @{ML_matchresult_fake [display,gray] |
1482 @{ML_matchresult_fake [display,gray] |
1483 "Thm.apply @{cterm \"P::nat \<Rightarrow> bool\"} @{cterm \"3::nat\"}" |
1483 \<open>Thm.apply @{cterm "P::nat \<Rightarrow> bool"} @{cterm "3::nat"}\<close> |
1484 "P 3"} |
1484 \<open>P 3\<close>} |
1485 |
1485 |
1486 Similarly the function @{ML_ind list_comb in Drule} from the structure @{ML_structure Drule} |
1486 Similarly the function @{ML_ind list_comb in Drule} from the structure @{ML_structure Drule} |
1487 applies a list of @{ML_type cterm}s. |
1487 applies a list of @{ML_type cterm}s. |
1488 |
1488 |
1489 @{ML_matchresult_fake [display,gray] |
1489 @{ML_matchresult_fake [display,gray] |
1490 "let |
1490 \<open>let |
1491 val chead = @{cterm \"P::unit \<Rightarrow> nat \<Rightarrow> bool\"} |
1491 val chead = @{cterm "P::unit \<Rightarrow> nat \<Rightarrow> bool"} |
1492 val cargs = [@{cterm \"()\"}, @{cterm \"3::nat\"}] |
1492 val cargs = [@{cterm "()"}, @{cterm "3::nat"}] |
1493 in |
1493 in |
1494 Drule.list_comb (chead, cargs) |
1494 Drule.list_comb (chead, cargs) |
1495 end" |
1495 end\<close> |
1496 "P () 3"} |
1496 \<open>P () 3\<close>} |
1497 |
1497 |
1498 \begin{readmore} |
1498 \begin{readmore} |
1499 For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see |
1499 For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see |
1500 the files @{ML_file "Pure/thm.ML"}, @{ML_file "Pure/more_thm.ML"} and |
1500 the files @{ML_file "Pure/thm.ML"}, @{ML_file "Pure/more_thm.ML"} and |
1501 @{ML_file "Pure/drule.ML"}. |
1501 @{ML_file "Pure/drule.ML"}. |
2035 establishing this theorem. We can traverse this proof-tree |
2035 establishing this theorem. We can traverse this proof-tree |
2036 extracting this information. Let us first extract the name of the |
2036 extracting this information. Let us first extract the name of the |
2037 established theorem. This can be done with |
2037 established theorem. This can be done with |
2038 |
2038 |
2039 @{ML_matchresult_fake [display,gray] |
2039 @{ML_matchresult_fake [display,gray] |
2040 "@{thm my_conjIa} |
2040 \<open>@{thm my_conjIa} |
2041 |> Thm.proof_body_of |
2041 |> Thm.proof_body_of |
2042 |> get_names" |
2042 |> get_names\<close> |
2043 "[\"Essential.my_conjIa\"]"} |
2043 \<open>["Essential.my_conjIa"]\<close>} |
2044 |
2044 |
2045 whereby \<open>Essential\<close> refers to the theory name in which |
2045 whereby \<open>Essential\<close> refers to the theory name in which |
2046 we established the theorem @{thm [source] my_conjIa}. The function @{ML_ind |
2046 we established the theorem @{thm [source] my_conjIa}. The function @{ML_ind |
2047 proof_body_of in Thm} returns a part of the data that is stored in a |
2047 proof_body_of in Thm} returns a part of the data that is stored in a |
2048 theorem. Notice that the first proof above references |
2048 theorem. Notice that the first proof above references |
2049 three theorems, namely @{thm [source] conjI}, @{thm [source] conjunct1} |
2049 three theorems, namely @{thm [source] conjI}, @{thm [source] conjunct1} |
2050 and @{thm [source] conjunct2}. We can obtain them by descending into the |
2050 and @{thm [source] conjunct2}. We can obtain them by descending into the |
2051 first level of the proof-tree, as follows. |
2051 first level of the proof-tree, as follows. |
2052 |
2052 |
2053 @{ML_matchresult_fake [display,gray] |
2053 @{ML_matchresult_fake [display,gray] |
2054 "@{thm my_conjIa} |
2054 \<open>@{thm my_conjIa} |
2055 |> Thm.proof_body_of |
2055 |> Thm.proof_body_of |
2056 |> get_pbodies |
2056 |> get_pbodies |
2057 |> map get_names |
2057 |> map get_names |
2058 |> List.concat" |
2058 |> List.concat\<close> |
2059 "[\"HOL.conjunct2\", \"HOL.conjunct1\", \"HOL.conjI\", \"Pure.protectD\", |
2059 \<open>["HOL.conjunct2", "HOL.conjunct1", "HOL.conjI", "Pure.protectD", |
2060 \"Pure.protectI\"]"} |
2060 "Pure.protectI"]\<close>} |
2061 |
2061 |
2062 The theorems @{thm [source] Pure.protectD} and @{thm [source] |
2062 The theorems @{thm [source] Pure.protectD} and @{thm [source] |
2063 Pure.protectI} that are internal theorems are always part of a |
2063 Pure.protectI} that are internal theorems are always part of a |
2064 proof in Isabelle. Note also that applications of \<open>assumption\<close> do not |
2064 proof in Isabelle. Note also that applications of \<open>assumption\<close> do not |
2065 count as a separate theorem, as you can see in the following code. |
2065 count as a separate theorem, as you can see in the following code. |
2066 |
2066 |
2067 @{ML_matchresult_fake [display,gray] |
2067 @{ML_matchresult_fake [display,gray] |
2068 "@{thm my_conjIb} |
2068 \<open>@{thm my_conjIb} |
2069 |> Thm.proof_body_of |
2069 |> Thm.proof_body_of |
2070 |> get_pbodies |
2070 |> get_pbodies |
2071 |> map get_names |
2071 |> map get_names |
2072 |> List.concat" |
2072 |> List.concat\<close> |
2073 "[\"Pure.protectD\", \"Pure.protectI\"]"} |
2073 \<open>["Pure.protectD", "Pure.protectI"]\<close>} |
2074 |
2074 |
2075 Interestingly, but not surprisingly, the proof of @{thm [source] |
2075 Interestingly, but not surprisingly, the proof of @{thm [source] |
2076 my_conjIc} procceeds quite differently from @{thm [source] my_conjIa} |
2076 my_conjIc} procceeds quite differently from @{thm [source] my_conjIa} |
2077 and @{thm [source] my_conjIb}, as can be seen by the theorems that |
2077 and @{thm [source] my_conjIb}, as can be seen by the theorems that |
2078 are returned for @{thm [source] my_conjIc}. |
2078 are returned for @{thm [source] my_conjIc}. |
2079 |
2079 |
2080 @{ML_matchresult_fake [display,gray] |
2080 @{ML_matchresult_fake [display,gray] |
2081 "@{thm my_conjIc} |
2081 \<open>@{thm my_conjIc} |
2082 |> Thm.proof_body_of |
2082 |> Thm.proof_body_of |
2083 |> get_pbodies |
2083 |> get_pbodies |
2084 |> map get_names |
2084 |> map get_names |
2085 |> List.concat" |
2085 |> List.concat\<close> |
2086 "[\"HOL.Eq_TrueI\", \"HOL.simp_thms_25\", \"HOL.eq_reflection\", |
2086 \<open>["HOL.Eq_TrueI", "HOL.simp_thms_25", "HOL.eq_reflection", |
2087 \"HOL.conjunct2\", \"HOL.conjunct1\", \"HOL.TrueI\", \"Pure.protectD\", |
2087 "HOL.conjunct2", "HOL.conjunct1", "HOL.TrueI", "Pure.protectD", |
2088 \"Pure.protectI\"]"} |
2088 "Pure.protectI"]\<close>} |
2089 |
2089 |
2090 Of course we can also descend into the second level of the tree |
2090 Of course we can also descend into the second level of the tree |
2091 ``behind'' @{thm [source] my_conjIa} say, which |
2091 ``behind'' @{thm [source] my_conjIa} say, which |
2092 means we obtain the theorems that are used in order to prove |
2092 means we obtain the theorems that are used in order to prove |
2093 @{thm [source] conjunct1}, @{thm [source] conjunct2} and @{thm [source] conjI}. |
2093 @{thm [source] conjunct1}, @{thm [source] conjunct2} and @{thm [source] conjI}. |
2094 |
2094 |
2095 @{ML_matchresult_fake [display, gray] |
2095 @{ML_matchresult_fake [display, gray] |
2096 "@{thm my_conjIa} |
2096 \<open>@{thm my_conjIa} |
2097 |> Thm.proof_body_of |
2097 |> Thm.proof_body_of |
2098 |> get_pbodies |
2098 |> get_pbodies |
2099 |> map get_pbodies |
2099 |> map get_pbodies |
2100 |> (map o map) get_names |
2100 |> (map o map) get_names |
2101 |> List.concat |
2101 |> List.concat |
2102 |> List.concat |
2102 |> List.concat |
2103 |> duplicates (op=)" |
2103 |> duplicates (op=)\<close> |
2104 "[\"HOL.spec\", \"HOL.and_def\", \"HOL.mp\", \"HOL.impI\", \"Pure.protectD\", |
2104 \<open>["HOL.spec", "HOL.and_def", "HOL.mp", "HOL.impI", "Pure.protectD", |
2105 \"Pure.protectI\"]"} |
2105 "Pure.protectI"]\<close>} |
2106 |
2106 |
2107 \begin{exercise} |
2107 \begin{exercise} |
2108 Have a look at the theorems that are used when a lemma is ``proved'' |
2108 Have a look at the theorems that are used when a lemma is ``proved'' |
2109 by \isacommand{sorry}? |
2109 by \isacommand{sorry}? |
2110 \end{exercise} |
2110 \end{exercise} |
2502 We deliberately chose a large string so that it spans over more than one line. |
2502 We deliberately chose a large string so that it spans over more than one line. |
2503 If we print out the string using the usual ``quick-and-dirty'' method, then |
2503 If we print out the string using the usual ``quick-and-dirty'' method, then |
2504 we obtain the ugly output: |
2504 we obtain the ugly output: |
2505 |
2505 |
2506 @{ML_matchresult_fake [display,gray] |
2506 @{ML_matchresult_fake [display,gray] |
2507 "tracing test_str" |
2507 \<open>tracing test_str\<close> |
2508 "fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo |
2508 \<open>fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo |
2509 ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa |
2509 ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa |
2510 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo |
2510 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo |
2511 oooooooooooooobaaaaaaaaaaaar"} |
2511 oooooooooooooobaaaaaaaaaaaar\<close>} |
2512 |
2512 |
2513 We obtain the same if we just use the function @{ML pprint}. |
2513 We obtain the same if we just use the function @{ML pprint}. |
2514 |
2514 |
2515 @{ML_matchresult_fake [display,gray] |
2515 @{ML_matchresult_fake [display,gray] |
2516 "pprint (Pretty.str test_str)" |
2516 \<open>pprint (Pretty.str test_str)\<close> |
2517 "fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo |
2517 \<open>fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar foooooooooo |
2518 ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa |
2518 ooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaa |
2519 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo |
2519 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo |
2520 oooooooooooooobaaaaaaaaaaaar"} |
2520 oooooooooooooobaaaaaaaaaaaar\<close>} |
2521 |
2521 |
2522 However by using pretty types you have the ability to indicate possible |
2522 However by using pretty types you have the ability to indicate possible |
2523 linebreaks for example at each whitespace. You can achieve this with the |
2523 linebreaks for example at each whitespace. You can achieve this with the |
2524 function @{ML_ind breaks in Pretty}, which expects a list of pretty types |
2524 function @{ML_ind breaks in Pretty}, which expects a list of pretty types |
2525 and inserts a possible line break in between every two elements in this |
2525 and inserts a possible line break in between every two elements in this |
2526 list. To print this list of pretty types as a single string, we concatenate |
2526 list. To print this list of pretty types as a single string, we concatenate |
2527 them with the function @{ML_ind blk in Pretty} as follows: |
2527 them with the function @{ML_ind blk in Pretty} as follows: |
2528 |
2528 |
2529 @{ML_matchresult_fake [display,gray] |
2529 @{ML_matchresult_fake [display,gray] |
2530 "let |
2530 \<open>let |
2531 val ptrs = map Pretty.str (space_explode \" \" test_str) |
2531 val ptrs = map Pretty.str (space_explode " " test_str) |
2532 in |
2532 in |
2533 pprint (Pretty.blk (0, Pretty.breaks ptrs)) |
2533 pprint (Pretty.blk (0, Pretty.breaks ptrs)) |
2534 end" |
2534 end\<close> |
2535 "fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2535 \<open>fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2536 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2536 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2537 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2537 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2538 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"} |
2538 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar\<close>} |
2539 |
2539 |
2540 Here the layout of @{ML test_str} is much more pleasing to the |
2540 Here the layout of @{ML test_str} is much more pleasing to the |
2541 eye. The @{ML "0"} in @{ML_ind blk in Pretty} stands for no hanging |
2541 eye. The @{ML \<open>0\<close>} in @{ML_ind blk in Pretty} stands for no hanging |
2542 indentation of the printed string. You can increase the indentation |
2542 indentation of the printed string. You can increase the indentation |
2543 and obtain |
2543 and obtain |
2544 |
2544 |
2545 @{ML_matchresult_fake [display,gray] |
2545 @{ML_matchresult_fake [display,gray] |
2546 "let |
2546 \<open>let |
2547 val ptrs = map Pretty.str (space_explode \" \" test_str) |
2547 val ptrs = map Pretty.str (space_explode " " test_str) |
2548 in |
2548 in |
2549 pprint (Pretty.blk (3, Pretty.breaks ptrs)) |
2549 pprint (Pretty.blk (3, Pretty.breaks ptrs)) |
2550 end" |
2550 end\<close> |
2551 "fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2551 \<open>fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2552 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2552 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2553 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2553 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2554 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"} |
2554 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar\<close>} |
2555 |
2555 |
2556 where starting from the second line the indent is 3. If you want |
2556 where starting from the second line the indent is 3. If you want |
2557 that every line starts with the same indent, you can use the |
2557 that every line starts with the same indent, you can use the |
2558 function @{ML_ind indent in Pretty} as follows: |
2558 function @{ML_ind indent in Pretty} as follows: |
2559 |
2559 |
2560 @{ML_matchresult_fake [display,gray] |
2560 @{ML_matchresult_fake [display,gray] |
2561 "let |
2561 \<open>let |
2562 val ptrs = map Pretty.str (space_explode \" \" test_str) |
2562 val ptrs = map Pretty.str (space_explode " " test_str) |
2563 in |
2563 in |
2564 pprint (Pretty.indent 10 (Pretty.blk (0, Pretty.breaks ptrs))) |
2564 pprint (Pretty.indent 10 (Pretty.blk (0, Pretty.breaks ptrs))) |
2565 end" |
2565 end\<close> |
2566 " fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2566 \<open> fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2567 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2567 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2568 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2568 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar |
2569 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"} |
2569 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar\<close>} |
2570 |
2570 |
2571 If you want to print out a list of items separated by commas and |
2571 If you want to print out a list of items separated by commas and |
2572 have the linebreaks handled properly, you can use the function |
2572 have the linebreaks handled properly, you can use the function |
2573 @{ML_ind commas in Pretty}. For example |
2573 @{ML_ind commas in Pretty}. For example |
2574 |
2574 |
2575 @{ML_matchresult_fake [display,gray] |
2575 @{ML_matchresult_fake [display,gray] |
2576 "let |
2576 \<open>let |
2577 val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020) |
2577 val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020) |
2578 in |
2578 in |
2579 pprint (Pretty.blk (0, Pretty.commas ptrs)) |
2579 pprint (Pretty.blk (0, Pretty.commas ptrs)) |
2580 end" |
2580 end\<close> |
2581 "99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006, |
2581 \<open>99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006, |
2582 100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015, |
2582 100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015, |
2583 100016, 100017, 100018, 100019, 100020"} |
2583 100016, 100017, 100018, 100019, 100020\<close>} |
2584 |
2584 |
2585 where @{ML upto} generates a list of integers. You can print out this |
2585 where @{ML upto} generates a list of integers. You can print out this |
2586 list as a ``set'', that means enclosed inside @{text [quotes] "{"} and |
2586 list as a ``set'', that means enclosed inside @{text [quotes] "{"} and |
2587 @{text [quotes] "}"}, and separated by commas using the function |
2587 @{text [quotes] "}"}, and separated by commas using the function |
2588 @{ML_ind enum in Pretty}. For example |
2588 @{ML_ind enum in Pretty}. For example |
2589 \<close> |
2589 \<close> |
2590 |
2590 |
2591 text \<open> |
2591 text \<open> |
2592 |
2592 |
2593 @{ML_matchresult_fake [display,gray] |
2593 @{ML_matchresult_fake [display,gray] |
2594 "let |
2594 \<open>let |
2595 val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020) |
2595 val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020) |
2596 in |
2596 in |
2597 pprint (Pretty.enum \",\" \"{\" \"}\" ptrs) |
2597 pprint (Pretty.enum "," "{" "}" ptrs) |
2598 end" |
2598 end\<close> |
2599 "{99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006, |
2599 \<open>{99998, 99999, 100000, 100001, 100002, 100003, 100004, 100005, 100006, |
2600 100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015, |
2600 100007, 100008, 100009, 100010, 100011, 100012, 100013, 100014, 100015, |
2601 100016, 100017, 100018, 100019, 100020}"} |
2601 100016, 100017, 100018, 100019, 100020}\<close>} |
2602 |
2602 |
2603 As can be seen, this function prints out the ``set'' so that starting |
2603 As can be seen, this function prints out the ``set'' so that starting |
2604 from the second, each new line has an indentation of 2. |
2604 from the second, each new line has an indentation of 2. |
2605 |
2605 |
2606 If you print out something that goes beyond the capabilities of the |
2606 If you print out something that goes beyond the capabilities of the |
2621 [Pretty.brk 1, Pretty.str "and", Pretty.brk 1, last] |
2621 [Pretty.brk 1, Pretty.str "and", Pretty.brk 1, last] |
2622 end\<close> |
2622 end\<close> |
2623 |
2623 |
2624 text \<open> |
2624 text \<open> |
2625 where Line 7 prints the beginning of the list and Line |
2625 where Line 7 prints the beginning of the list and Line |
2626 8 the last item. We have to use @{ML "Pretty.brk 1"} in order |
2626 8 the last item. We have to use @{ML \<open>Pretty.brk 1\<close>} in order |
2627 to insert a space (of length 1) before the |
2627 to insert a space (of length 1) before the |
2628 @{text [quotes] "and"}. This space is also a place where a line break |
2628 @{text [quotes] "and"}. This space is also a place where a line break |
2629 can occur. We do the same after the @{text [quotes] "and"}. This gives you |
2629 can occur. We do the same after the @{text [quotes] "and"}. This gives you |
2630 for example |
2630 for example |
2631 |
2631 |
2632 @{ML_matchresult_fake [display,gray] |
2632 @{ML_matchresult_fake [display,gray] |
2633 "let |
2633 \<open>let |
2634 val ptrs1 = map (Pretty.str o string_of_int) (1 upto 22) |
2634 val ptrs1 = map (Pretty.str o string_of_int) (1 upto 22) |
2635 val ptrs2 = map (Pretty.str o string_of_int) (10 upto 28) |
2635 val ptrs2 = map (Pretty.str o string_of_int) (10 upto 28) |
2636 in |
2636 in |
2637 pprint (Pretty.blk (0, and_list ptrs1)); |
2637 pprint (Pretty.blk (0, and_list ptrs1)); |
2638 pprint (Pretty.blk (0, and_list ptrs2)) |
2638 pprint (Pretty.blk (0, and_list ptrs2)) |
2639 end" |
2639 end\<close> |
2640 "1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21 |
2640 \<open>1, 2, 3, 4, 5, 6, 7, 8, 9, 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21 |
2641 and 22 |
2641 and 22 |
2642 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27 and |
2642 10, 11, 12, 13, 14, 15, 16, 17, 18, 19, 20, 21, 22, 23, 24, 25, 26, 27 and |
2643 28"} |
2643 28\<close>} |
2644 |
2644 |
2645 Like @{ML blk in Pretty}, the function @{ML_ind chunks in Pretty} prints out |
2645 Like @{ML blk in Pretty}, the function @{ML_ind chunks in Pretty} prints out |
2646 a list of items, but automatically inserts forced breaks between each item. |
2646 a list of items, but automatically inserts forced breaks between each item. |
2647 Compare |
2647 Compare |
2648 |
2648 |
2649 @{ML_matchresult_fake [display,gray] |
2649 @{ML_matchresult_fake [display,gray] |
2650 "let |
2650 \<open>let |
2651 val a_and_b = [Pretty.str \"a\", Pretty.str \"b\"] |
2651 val a_and_b = [Pretty.str "a", Pretty.str "b"] |
2652 in |
2652 in |
2653 pprint (Pretty.blk (0, a_and_b)); |
2653 pprint (Pretty.blk (0, a_and_b)); |
2654 pprint (Pretty.chunks a_and_b) |
2654 pprint (Pretty.chunks a_and_b) |
2655 end" |
2655 end\<close> |
2656 "ab |
2656 \<open>ab |
2657 a |
2657 a |
2658 b"} |
2658 b\<close>} |
2659 |
2659 |
2660 The function @{ML_ind big_list in Pretty} helps with printing long lists. |
2660 The function @{ML_ind big_list in Pretty} helps with printing long lists. |
2661 It is used for example in the command \isacommand{print\_theorems}. An |
2661 It is used for example in the command \isacommand{print\_theorems}. An |
2662 example is as follows. |
2662 example is as follows. |
2663 |
2663 |
2664 @{ML_matchresult_fake [display,gray] |
2664 @{ML_matchresult_fake [display,gray] |
2665 "let |
2665 \<open>let |
2666 val pstrs = map (Pretty.str o string_of_int) (4 upto 10) |
2666 val pstrs = map (Pretty.str o string_of_int) (4 upto 10) |
2667 in |
2667 in |
2668 pprint (Pretty.big_list \"header\" pstrs) |
2668 pprint (Pretty.big_list "header" pstrs) |
2669 end" |
2669 end\<close> |
2670 "header |
2670 \<open>header |
2671 4 |
2671 4 |
2672 5 |
2672 5 |
2673 6 |
2673 6 |
2674 7 |
2674 7 |
2675 8 |
2675 8 |
2676 9 |
2676 9 |
2677 10"} |
2677 10\<close>} |
2678 |
2678 |
2679 The point of the pretty-printing functions is to conveninetly obtain |
2679 The point of the pretty-printing functions is to conveninetly obtain |
2680 a lay-out of terms and types that is pleasing to the eye. If we print |
2680 a lay-out of terms and types that is pleasing to the eye. If we print |
2681 out the the terms produced by the the function @{ML de_bruijn} from |
2681 out the the terms produced by the the function @{ML de_bruijn} from |
2682 exercise~\ref{ex:debruijn} we obtain the following: |
2682 exercise~\ref{ex:debruijn} we obtain the following: |
2683 |
2683 |
2684 @{ML_matchresult_fake [display,gray] |
2684 @{ML_matchresult_fake [display,gray] |
2685 "pprint (Syntax.pretty_term @{context} (de_bruijn 4))" |
2685 \<open>pprint (Syntax.pretty_term @{context} (de_bruijn 4))\<close> |
2686 "(P 3 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> |
2686 \<open>(P 3 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> |
2687 (P 2 = P 3 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> |
2687 (P 2 = P 3 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> |
2688 (P 1 = P 2 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> |
2688 (P 1 = P 2 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<and> |
2689 (P 1 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<longrightarrow> |
2689 (P 1 = P 4 \<longrightarrow> P 4 \<and> P 3 \<and> P 2 \<and> P 1) \<longrightarrow> |
2690 P 4 \<and> P 3 \<and> P 2 \<and> P 1"} |
2690 P 4 \<and> P 3 \<and> P 2 \<and> P 1\<close>} |
2691 |
2691 |
2692 We use the function @{ML_ind pretty_term in Syntax} for pretty-printing |
2692 We use the function @{ML_ind pretty_term in Syntax} for pretty-printing |
2693 terms. Next we like to pretty-print a term and its type. For this we use the |
2693 terms. Next we like to pretty-print a term and its type. For this we use the |
2694 function \<open>tell_type\<close>. |
2694 function \<open>tell_type\<close>. |
2695 \<close> |
2695 \<close> |