author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 76987 | 4c275405faae |
permissions | -rw-r--r-- |
61656 | 1 |
(*:maxLineLen=78:*) |
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
2 |
|
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
3 |
theory Preface |
66453
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents:
61656
diff
changeset
|
4 |
imports Base "HOL-Eisbach.Eisbach_Tools" |
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
5 |
begin |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
6 |
|
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
7 |
text \<open> |
61576 | 8 |
\<^emph>\<open>Eisbach\<close> is a collection of tools which form the basis for defining new |
76987 | 9 |
proof methods in Isabelle/Isar~\<^cite>\<open>"Wenzel-PhD"\<close>. It can be thought of as |
61576 | 10 |
a ``proof method language'', but is more precisely an infrastructure for |
11 |
defining new proof methods out of existing ones. |
|
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
12 |
|
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
13 |
The core functionality of Eisbach is provided by the Isar @{command method} |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
14 |
command. Here users may define new methods by combining existing ones with |
60298
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
wenzelm
parents:
60288
diff
changeset
|
15 |
the usual Isar syntax. These methods can be abstracted over terms, facts and |
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
16 |
other methods, as one might expect in any higher-order functional language. |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
17 |
|
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
18 |
Additional functionality is provided by extending the space of methods and |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
19 |
attributes. The new @{method match} method allows for explicit control-flow, |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
20 |
by taking a match target and a list of pattern-method pairs. By using the |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
21 |
functionality provided by Eisbach, additional support methods can be easily |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
22 |
written. For example, the @{method catch} method, which provides basic |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
23 |
try-catch functionality, only requires a few lines of ML. |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
24 |
|
75160 | 25 |
Eisbach enables users to implement automated proof tools conveniently via |
26 |
Isabelle/Isar syntax. This is in contrast to the traditional approach to use |
|
27 |
Isabelle/ML (via @{command_ref method_setup}), which poses a higher |
|
28 |
barrier-to-entry for casual users. |
|
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
29 |
|
61417 | 30 |
\<^medskip> |
75160 | 31 |
This manual is written for readers familiar with Isabelle/Isar, but not |
61576 | 32 |
necessarily Isabelle/ML. It covers the usage of the @{command method} as |
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
33 |
well as the @{method match} method, as well as discussing their integration |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
34 |
with existing Isar concepts such as @{command named_theorems}. |
68480 | 35 |
|
69597 | 36 |
These commands are provided by theory \<^theory>\<open>HOL-Eisbach.Eisbach\<close>: it |
75160 | 37 |
needs to be imported by all Eisbach applications. Theory |
38 |
\<^theory>\<open>HOL-Eisbach.Eisbach_Tools\<close> provides additional proof methods and |
|
68484 | 39 |
attributes that are occasionally useful. |
60288
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
40 |
\<close> |
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
wenzelm
parents:
diff
changeset
|
41 |
|
67399 | 42 |
end |