| author | wenzelm | 
| Wed, 20 Dec 2023 22:32:57 +0100 | |
| changeset 79324 | b03e22697999 | 
| parent 78150 | 2963ea647c2a | 
| 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 Manual  | 
| 
66453
 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 
wenzelm 
parents: 
62969 
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  | 
chapter \<open>The method command\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
  The @{command_def method} command provides the ability to write proof
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
methods by combining existing ones with their usual syntax. Specifically it  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
allows compound proof methods to be named, and to extend the name space of  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
basic methods accordingly. Method definitions may abstract over parameters:  | 
| 78150 | 14  | 
terms, facts, or other methods. They may also provide an optional text  | 
15  | 
  description for display in @{command_ref print_methods}.
 | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
|
| 61417 | 17  | 
\<^medskip>  | 
| 61576 | 18  | 
The syntax diagram below refers to some syntactic categories that are  | 
| 76987 | 19  | 
further defined in \<^cite>\<open>"isabelle-isar-ref"\<close>.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
|
| 69597 | 21  | 
\<^rail>\<open>  | 
| 78150 | 22  | 
    @@{command method} name args description @'=' method
 | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
23  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
24  | 
args: term_args? method_args? \<newline> fact_args? decl_args?  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
    term_args: @'for' @{syntax "fixes"}
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
28  | 
method_args: @'methods' (name+)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
fact_args: @'uses' (name+)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
decl_args: @'declares' (name+)  | 
| 78150 | 33  | 
;  | 
34  | 
    description: @{syntax text}?
 | 
|
| 69597 | 35  | 
\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
36  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
38  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
39  | 
section \<open>Basic method definitions\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
42  | 
Consider the following proof that makes use of usual Isar method  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
43  | 
combinators.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
44  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
46  | 
lemma "P \<and> Q \<longrightarrow> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
by ((rule impI, (erule conjE)?) | assumption)+  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
It is clear that this compound method will be applicable in more cases than  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
  this proof alone. With the @{command method} command we can define a proof
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
method that makes the above functionality available generally.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
53  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
method prop_solver\<^sub>1 =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
((rule impI, (erule conjE)?) | assumption)+  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
lemma "P \<and> Q \<and> R \<longrightarrow> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
by prop_solver\<^sub>1  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
60  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
61  | 
text \<open>  | 
| 61576 | 62  | 
In this example, the facts \<open>impI\<close> and \<open>conjE\<close> are static. They are evaluated  | 
63  | 
once when the method is defined and cannot be changed later. This makes the  | 
|
64  | 
method stable in the sense of \<^emph>\<open>static scoping\<close>: naming another fact \<open>impI\<close>  | 
|
65  | 
in a later context won't affect the behaviour of \<open>prop_solver\<^sub>1\<close>.  | 
|
| 78150 | 66  | 
|
67  | 
The following example defines the same method and gives it a description for the  | 
|
68  | 
  @{command_ref print_methods} command.
 | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
|
| 78150 | 71  | 
method prop_solver\<^sub>2 \<open>solver for propositional formulae\<close> =  | 
72  | 
((rule impI, (erule conjE)?) | assumption)+  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
section \<open>Term abstraction\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
76  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
  Methods can also abstract over terms using the @{keyword_def "for"} keyword,
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
optionally providing type constraints. For instance, the following proof  | 
| 69597 | 79  | 
method \<open>intro_ex\<close> takes a term \<^term>\<open>y\<close> of any type, which it uses to  | 
80  | 
instantiate the \<^term>\<open>x\<close>-variable of \<open>exI\<close> (existential introduction)  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
81  | 
before applying the result as a rule. The instantiation is performed here by  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
  Isar's @{attribute_ref "where"} attribute. If the current subgoal is to find
 | 
| 69597 | 83  | 
a witness for the given predicate \<^term>\<open>Q\<close>, then this has the effect of  | 
84  | 
committing to \<^term>\<open>y\<close>.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
85  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
86  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
87  | 
method intro_ex for Q :: "'a \<Rightarrow> bool" and y :: 'a =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
(rule exI ["where" P = Q and x = y])  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
91  | 
text \<open>  | 
| 69597 | 92  | 
The term parameters \<^term>\<open>y\<close> and \<^term>\<open>Q\<close> can be used arbitrarily inside  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
the method body, as part of attribute applications or arguments to other  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
methods. The expression is type-checked as far as possible when the method  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
is defined, however dynamic type errors can still occur when it is invoked  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
(e.g.\ when terms are instantiated in a parameterized fact). Actual term  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
arguments are supplied positionally, in the same order as in the method  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
definition.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
lemma "P a \<Longrightarrow> \<exists>x. P x"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
by (intro_ex P a)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
section \<open>Fact abstraction\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
subsection \<open>Named theorems\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
text \<open>  | 
| 61576 | 110  | 
A \<^emph>\<open>named theorem\<close> is a fact whose contents are produced dynamically within  | 
111  | 
  the current proof context. The Isar command @{command_ref "named_theorems"}
 | 
|
| 75160 | 112  | 
declares a dynamic fact with a corresponding \<^emph>\<open>attribute\<close> of the same  | 
113  | 
name. This allows to maintain a collection of facts in the context as  | 
|
114  | 
follows:  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
named_theorems intros  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
text \<open>  | 
| 61493 | 120  | 
So far \<open>intros\<close> refers to the empty fact. Using the Isar command  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
121  | 
  @{command_ref "declare"} we may apply declaration attributes to the context.
 | 
| 61576 | 122  | 
Below we declare both \<open>conjI\<close> and \<open>impI\<close> as \<open>intros\<close>, adding them to the  | 
123  | 
named theorem slot.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
124  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
125  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
126  | 
declare conjI [intros] and impI [intros]  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
127  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
We can refer to named theorems as dynamic facts within a particular proof  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
130  | 
context, which are evaluated whenever the method is invoked. Instead of  | 
| 61493 | 131  | 
having facts hard-coded into the method, as in \<open>prop_solver\<^sub>1\<close>, we can  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
132  | 
instead refer to these named theorems.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
133  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
134  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
named_theorems elims  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
declare conjE [elims]  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
method prop_solver\<^sub>3 =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
((rule intros, (erule elims)?) | assumption)+  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
140  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
lemma "P \<and> Q \<longrightarrow> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
by prop_solver\<^sub>3  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
145  | 
Often these named theorems need to be augmented on the spot, when a method  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
146  | 
  is invoked. The @{keyword_def "declares"} keyword in the signature of
 | 
| 61576 | 147  | 
  @{command method} adds the common method syntax \<open>method decl: facts\<close> for
 | 
148  | 
each named theorem \<open>decl\<close>.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
151  | 
method prop_solver\<^sub>4 declares intros elims =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
152  | 
((rule intros, (erule elims)?) | assumption)+  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
153  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
154  | 
lemma "P \<and> (P \<longrightarrow> Q) \<longrightarrow> Q \<and> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
155  | 
by (prop_solver\<^sub>4 elims: impE intros: conjI)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
156  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
157  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
158  | 
subsection \<open>Simple fact abstraction\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
160  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
  The @{keyword "declares"} keyword requires that a corresponding dynamic fact
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
162  | 
  has been declared with @{command_ref named_theorems}. This is useful for
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
managing collections of facts which are to be augmented with declarations,  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
164  | 
but is overkill if we simply want to pass a fact to a method.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
166  | 
  We may use the @{keyword_def "uses"} keyword in the method header to provide
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
167  | 
  a simple fact parameter. In contrast to @{keyword "declares"}, these facts
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
168  | 
are always implicitly empty unless augmented when the method is invoked.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
170  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
171  | 
method rule_twice uses my_rule =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
172  | 
(rule my_rule, rule my_rule)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
173  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
174  | 
lemma "P \<Longrightarrow> Q \<Longrightarrow> (P \<and> Q) \<and> Q"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
175  | 
by (rule_twice my_rule: conjI)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
176  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
177  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
178  | 
section \<open>Higher-order methods\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
179  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
180  | 
text \<open>  | 
| 61576 | 181  | 
The \<^emph>\<open>structured concatenation\<close> combinator ``\<open>method\<^sub>1 ; method\<^sub>2\<close>'' was  | 
| 75160 | 182  | 
introduced in Isabelle2015, motivated by the development of Eisbach. It is  | 
183  | 
similar to ``\<open>method\<^sub>1, method\<^sub>2\<close>'', but \<open>method\<^sub>2\<close> is invoked on \<^emph>\<open>all\<close>  | 
|
| 61576 | 184  | 
subgoals that have newly emerged from \<open>method\<^sub>1\<close>. This is useful to handle  | 
185  | 
cases where the number of subgoals produced by a method is determined  | 
|
186  | 
dynamically at run-time.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
187  | 
\<close>  | 
| 61417 | 188  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
method conj_with uses rule =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
(intro conjI ; intro rule)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
assumes A: "P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
194  | 
shows "P \<and> P \<and> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
195  | 
by (conj_with rule: A)  | 
| 61417 | 196  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
198  | 
Method definitions may take other methods as arguments, and thus implement  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
199  | 
method combinators with prefix syntax. For example, to more usefully exploit  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
200  | 
Isabelle's backtracking, the explicit requirement that a method solve all  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
201  | 
produced subgoals is frequently useful. This can easily be written as a  | 
| 61576 | 202  | 
  \<^emph>\<open>higher-order method\<close> using ``\<open>;\<close>''. The @{keyword "methods"} keyword
 | 
203  | 
denotes method parameters that are other proof methods to be invoked by the  | 
|
204  | 
method being defined.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
206  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
method solve methods m = (m ; fail)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
text \<open>  | 
| 61576 | 210  | 
Given some method-argument \<open>m\<close>, \<open>solve \<open>m\<close>\<close> applies the method \<open>m\<close> and then  | 
211  | 
fails whenever \<open>m\<close> produces any new unsolved subgoals --- i.e. when \<open>m\<close>  | 
|
212  | 
fails to completely discharge the goal it was applied to.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
215  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
216  | 
section \<open>Example\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
217  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
218  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
219  | 
With these simple features we are ready to write our first non-trivial proof  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
220  | 
method. Returning to the first-order logic example, the following method  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
221  | 
definition applies various rules with their canonical methods.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
222  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
223  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
224  | 
named_theorems subst  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
226  | 
method prop_solver declares intros elims subst =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
227  | 
(assumption |  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
(rule intros) | erule elims |  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
229  | 
subst subst | subst (asm) subst |  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
230  | 
(erule notE ; solve \<open>prop_solver\<close>))+  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
231  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
text \<open>  | 
| 61576 | 233  | 
The only non-trivial part above is the final alternative \<open>(erule notE ;  | 
234  | 
solve \<open>prop_solver\<close>)\<close>. Here, in the case that all other alternatives fail,  | 
|
| 69597 | 235  | 
the method takes one of the assumptions \<^term>\<open>\<not> P\<close> of the current goal  | 
| 61576 | 236  | 
and eliminates it with the rule \<open>notE\<close>, causing the goal to be proved to  | 
| 69597 | 237  | 
become \<^term>\<open>P\<close>. The method then recursively invokes itself on the  | 
| 61576 | 238  | 
remaining goals. The job of the recursive call is to demonstrate that there  | 
| 69597 | 239  | 
is a contradiction in the original assumptions (i.e.\ that \<^term>\<open>P\<close> can be  | 
| 61576 | 240  | 
derived from them). Note this recursive invocation is applied with the  | 
241  | 
  @{method solve} method combinator to ensure that a contradiction will indeed
 | 
|
242  | 
be shown. In the case where a contradiction cannot be found, backtracking  | 
|
| 69597 | 243  | 
will occur and a different assumption \<^term>\<open>\<not> Q\<close> will be chosen for  | 
| 61576 | 244  | 
elimination.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
245  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
246  | 
  Note that the recursive call to @{method prop_solver} does not have any
 | 
| 61576 | 247  | 
parameters passed to it. Recall that fact parameters, e.g.\ \<open>intros\<close>,  | 
248  | 
\<open>elims\<close>, and \<open>subst\<close>, are managed by declarations in the current proof  | 
|
249  | 
  context. They will therefore be passed to any recursive call to @{method
 | 
|
250  | 
prop_solver} and, more generally, any invocation of a method which declares  | 
|
251  | 
these named theorems.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
252  | 
|
| 61417 | 253  | 
\<^medskip>  | 
254  | 
  After declaring some standard rules to the context, the @{method
 | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
prop_solver} becomes capable of solving non-trivial propositional  | 
| 61576 | 256  | 
tautologies.  | 
257  | 
\<close>  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
258  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
259  | 
lemmas [intros] =  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
260  | 
      conjI  \<comment> \<open>@{thm conjI}\<close>
 | 
| 
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
261  | 
      impI  \<comment> \<open>@{thm impI}\<close>
 | 
| 
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
262  | 
      disjCI  \<comment> \<open>@{thm disjCI}\<close>
 | 
| 
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
263  | 
      iffI  \<comment> \<open>@{thm iffI}\<close>
 | 
| 
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
264  | 
      notI  \<comment> \<open>@{thm notI}\<close>
 | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
265  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
266  | 
lemmas [elims] =  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
267  | 
      impCE  \<comment> \<open>@{thm impCE}\<close>
 | 
| 
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
268  | 
      conjE  \<comment> \<open>@{thm conjE}\<close>
 | 
| 
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
66453 
diff
changeset
 | 
269  | 
      disjE  \<comment> \<open>@{thm disjE}\<close>
 | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
271  | 
lemma "(A \<or> B) \<and> (A \<longrightarrow> C) \<and> (B \<longrightarrow> C) \<longrightarrow> C"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
272  | 
by prop_solver  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
273  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
274  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
275  | 
chapter \<open>The match method \label{s:matching}\<close>
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
276  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
277  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
278  | 
So far we have seen methods defined as simple combinations of other methods.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
279  | 
Some familiar programming language concepts have been introduced (i.e.\  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
280  | 
abstraction and recursion). The only control flow has been implicitly the  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
281  | 
result of backtracking. When designing more sophisticated proof methods this  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
282  | 
proves too restrictive and difficult to manage conceptually.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
283  | 
|
| 61576 | 284  | 
  To address this, we introduce the @{method_def match} method, which provides
 | 
285  | 
more direct access to the higher-order matching facility at the core of  | 
|
286  | 
Isabelle. It is implemented as a separate proof method (in Isabelle/ML), and  | 
|
287  | 
thus can be directly applied to proofs, however it is most useful when  | 
|
288  | 
applied in the context of writing Eisbach method definitions.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
289  | 
|
| 61417 | 290  | 
\<^medskip>  | 
| 61576 | 291  | 
The syntax diagram below refers to some syntactic categories that are  | 
| 76987 | 292  | 
further defined in \<^cite>\<open>"isabelle-isar-ref"\<close>.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
293  | 
|
| 69597 | 294  | 
\<^rail>\<open>  | 
| 
61912
 
ad710de5c576
more standard nesting of sub-language: Parse.text allows atomic entities without quotes;
 
wenzelm 
parents: 
61853 
diff
changeset
 | 
295  | 
    @@{method match} kind @'in' (pattern '\<Rightarrow>' @{syntax text} + '\<bar>')
 | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
296  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
297  | 
kind:  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
298  | 
      (@'conclusion' | @'premises' ('(' 'local' ')')? |
 | 
| 62969 | 299  | 
       '(' term ')' | @{syntax thms})
 | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
300  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
301  | 
pattern: fact_name? term args? \<newline> (@'for' fixes)?  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
302  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
303  | 
    fact_name: @{syntax name} @{syntax attributes}? ':'
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
304  | 
;  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
305  | 
    args: '(' (('multi' | 'cut' nat?) + ',') ')'
 | 
| 69597 | 306  | 
\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
307  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
308  | 
Matching allows methods to introspect the goal state, and to implement more  | 
| 61576 | 309  | 
explicit control flow. In the basic case, a term or fact \<open>ts\<close> is given to  | 
310  | 
match against as a \<^emph>\<open>match target\<close>, along with a collection of  | 
|
311  | 
pattern-method pairs \<open>(p, m)\<close>: roughly speaking, when the pattern \<open>p\<close>  | 
|
312  | 
matches any member of \<open>ts\<close>, the \<^emph>\<open>inner\<close> method \<open>m\<close> will be executed.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
313  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
314  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
315  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
316  | 
assumes X:  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
317  | 
"Q \<longrightarrow> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
318  | 
"Q"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
319  | 
shows P  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
320  | 
by (match X in I: "Q \<longrightarrow> P" and I': "Q" \<Rightarrow> \<open>insert mp [OF I I']\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
321  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
322  | 
text \<open>  | 
| 61576 | 323  | 
In this example we have a structured Isar proof, with the named assumption  | 
| 69597 | 324  | 
\<open>X\<close> and a conclusion \<^term>\<open>P\<close>. With the match method we can find the  | 
325  | 
local facts \<^term>\<open>Q \<longrightarrow> P\<close> and \<^term>\<open>Q\<close>, binding them to separately as  | 
|
| 61576 | 326  | 
  \<open>I\<close> and \<open>I'\<close>. We then specialize the modus-ponens rule @{thm mp [of Q P]} to
 | 
327  | 
these facts to solve the goal.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
328  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
329  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
331  | 
section \<open>Subgoal focus\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
332  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
333  | 
text\<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
334  | 
In the previous example we were able to match against an assumption out of  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
335  | 
the Isar proof state. In general, however, proof subgoals can be  | 
| 61477 | 336  | 
\<^emph>\<open>unstructured\<close>, with goal parameters and premises arising from rule  | 
| 61576 | 337  | 
  application. To address this, @{method match} uses \<^emph>\<open>subgoal focusing\<close> to
 | 
338  | 
produce structured goals out of unstructured ones. In place of fact or term,  | 
|
339  | 
  we may give the keyword @{keyword_def "premises"} as the match target. This
 | 
|
340  | 
causes a subgoal focus on the first subgoal, lifting local goal parameters  | 
|
341  | 
to fixed term variables and premises into hypothetical theorems. The match  | 
|
342  | 
is performed against these theorems, naming them and binding them as  | 
|
343  | 
  appropriate. Similarly giving the keyword @{keyword_def "conclusion"}
 | 
|
344  | 
matches against the conclusion of the first subgoal.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
345  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
346  | 
An unstructured version of the previous example can then be similarly solved  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
347  | 
through focusing.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
348  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
349  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
350  | 
lemma "Q \<longrightarrow> P \<Longrightarrow> Q \<Longrightarrow> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
351  | 
by (match premises in  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
352  | 
I: "Q \<longrightarrow> P" and I': "Q" \<Rightarrow> \<open>insert mp [OF I I']\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
353  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
354  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
355  | 
  Match variables may be specified by giving a list of @{keyword_ref
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
356  | 
"for"}-fixes after the pattern description. This marks those terms as bound  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
357  | 
variables, which may be used in the method body.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
358  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
359  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
360  | 
lemma "Q \<longrightarrow> P \<Longrightarrow> Q \<Longrightarrow> P"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
361  | 
by (match premises in I: "Q \<longrightarrow> A" and I': "Q" for A \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
362  | 
\<open>match conclusion in A \<Rightarrow> \<open>insert mp [OF I I']\<close>\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
363  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
364  | 
text \<open>  | 
| 69597 | 365  | 
In this example \<^term>\<open>A\<close> is a match variable which is bound to \<^term>\<open>P\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
366  | 
  upon a successful match. The inner @{method match} then matches the
 | 
| 69597 | 367  | 
now-bound \<^term>\<open>A\<close> (bound to \<^term>\<open>P\<close>) against the conclusion (also \<^term>\<open>P\<close>), finally applying the specialized rule to solve the goal.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
368  | 
|
| 61576 | 369  | 
Schematic terms like \<open>?P\<close> may also be used to specify match variables, but  | 
370  | 
the result of the match is not bound, and thus cannot be used in the inner  | 
|
371  | 
method body.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
372  | 
|
| 61417 | 373  | 
\<^medskip>  | 
| 61576 | 374  | 
In the following example we extract the predicate of an existentially  | 
375  | 
quantified conclusion in the current subgoal and search the current premises  | 
|
376  | 
for a matching fact. If both matches are successful, we then instantiate the  | 
|
377  | 
existential introduction rule with both the witness and predicate, solving  | 
|
378  | 
with the matched premise.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
379  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
380  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
381  | 
method solve_ex =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
382  | 
(match conclusion in "\<exists>x. Q x" for Q \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
383  | 
\<open>match premises in U: "Q y" for y \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
384  | 
\<open>rule exI [where P = Q and x = y, OF U]\<close>\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
385  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
386  | 
text \<open>  | 
| 69597 | 387  | 
  The first @{method match} matches the pattern \<^term>\<open>\<exists>x. Q x\<close> against the
 | 
388  | 
current conclusion, binding the term \<^term>\<open>Q\<close> in the inner match. Next  | 
|
| 61576 | 389  | 
the pattern \<open>Q y\<close> is matched against all premises of the current subgoal. In  | 
| 69597 | 390  | 
this case \<^term>\<open>Q\<close> is fixed and \<^term>\<open>y\<close> may be instantiated. Once a  | 
| 61576 | 391  | 
match is found, the local fact \<open>U\<close> is bound to the matching premise and the  | 
| 69597 | 392  | 
variable \<^term>\<open>y\<close> is bound to the matching witness. The existential  | 
393  | 
  introduction rule \<open>exI:\<close>~@{thm exI} is then instantiated with \<^term>\<open>y\<close> as
 | 
|
394  | 
the witness and \<^term>\<open>Q\<close> as the predicate, with its proof obligation  | 
|
| 61576 | 395  | 
  solved by the local fact U (using the Isar attribute @{attribute OF}). The
 | 
396  | 
following example is a trivial use of this method.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
397  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
398  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
399  | 
lemma "halts p \<Longrightarrow> \<exists>x. halts x"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
400  | 
by solve_ex  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
401  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
402  | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
403  | 
subsection \<open>Operating within a focus\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
404  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
405  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
406  | 
Subgoal focusing provides a structured form of a subgoal, allowing for more  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
407  | 
expressive introspection of the goal state. This requires some consideration  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
408  | 
  in order to be used effectively. When the keyword @{keyword "premises"} is
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
409  | 
given as the match target, the premises of the subgoal are lifted into  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
410  | 
hypothetical theorems, which can be found and named via match patterns.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
411  | 
Additionally these premises are stripped from the subgoal, leaving only the  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
412  | 
conclusion. This renders them inaccessible to standard proof methods which  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
413  | 
  operate on the premises, such as @{method frule} or @{method erule}. Naive
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
414  | 
usage of these methods within a match will most likely not function as the  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
415  | 
method author intended.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
416  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
417  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
418  | 
method my_allE_bad for y :: 'a =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
419  | 
(match premises in I: "\<forall>x :: 'a. ?Q x" \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
420  | 
\<open>erule allE [where x = y]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
421  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
422  | 
text \<open>  | 
| 69597 | 423  | 
Here we take a single parameter \<^term>\<open>y\<close> and specialize the universal  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
424  | 
  elimination rule (@{thm allE}) to it, then attempt to apply this specialized
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
425  | 
  rule with @{method erule}. The method @{method erule} will attempt to unify
 | 
| 69597 | 426  | 
  with a universal quantifier in the premises that matches the type of \<^term>\<open>y\<close>. Since @{keyword "premises"} causes a focus, however, there are no
 | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
427  | 
  subgoal premises to be found and thus @{method my_allE_bad} will always
 | 
| 61576 | 428  | 
fail. If focusing instead left the premises in place, using methods like  | 
429  | 
  @{method erule} would lead to unintended behaviour, specifically during
 | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
430  | 
  backtracking. In our example, @{method erule} could choose an alternate
 | 
| 61576 | 431  | 
premise while backtracking, while leaving \<open>I\<close> bound to the original match.  | 
432  | 
In the case of more complex inner methods, where either \<open>I\<close> or bound terms  | 
|
433  | 
are used, this would almost certainly not be the intended behaviour.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
434  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
435  | 
An alternative implementation would be to specialize the elimination rule to  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
436  | 
the bound term and apply it directly.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
437  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
438  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
439  | 
method my_allE_almost for y :: 'a =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
440  | 
(match premises in I: "\<forall>x :: 'a. ?Q x" \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
441  | 
\<open>rule allE [where x = y, OF I]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
442  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
443  | 
lemma "\<forall>x. P x \<Longrightarrow> P y"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
444  | 
by (my_allE_almost y)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
445  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
446  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
447  | 
This method will insert a specialized duplicate of a universally quantified  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
448  | 
premise. Although this will successfully apply in the presence of such a  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
449  | 
premise, it is not likely the intended behaviour. Repeated application of  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
450  | 
this method will produce an infinite stream of duplicate specialized  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
451  | 
premises, due to the original premise never being removed. To address this,  | 
| 61576 | 452  | 
  matched premises may be declared with the @{attribute thin} attribute. This
 | 
453  | 
will hide the premise from subsequent inner matches, and remove it from the  | 
|
454  | 
list of premises when the inner method has finished and the subgoal is  | 
|
| 61493 | 455  | 
unfocused. It can be considered analogous to the existing \<open>thin_tac\<close>.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
456  | 
|
| 61576 | 457  | 
To complete our example, the correct implementation of the method will  | 
458  | 
  @{attribute thin} the premise from the match and then apply it to the
 | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
459  | 
specialized elimination rule.\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
461  | 
method my_allE for y :: 'a =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
462  | 
(match premises in I [thin]: "\<forall>x :: 'a. ?Q x" \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
463  | 
\<open>rule allE [where x = y, OF I]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
464  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
465  | 
lemma "\<forall>x. P x \<Longrightarrow> \<forall>x. Q x \<Longrightarrow> P y \<and> Q y"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
466  | 
by (my_allE y)+ (rule conjI)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
467  | 
|
| 61576 | 468  | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
469  | 
subsubsection \<open>Inner focusing\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
470  | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
471  | 
text \<open>  | 
| 61576 | 472  | 
Premises are \<^emph>\<open>accumulated\<close> for the purposes of subgoal focusing. In  | 
473  | 
  contrast to using standard methods like @{method frule} within focused
 | 
|
474  | 
  match, another @{method match} will have access to all the premises of the
 | 
|
475  | 
outer focus.  | 
|
| 61477 | 476  | 
\<close>  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
477  | 
|
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
478  | 
lemma "A \<Longrightarrow> B \<Longrightarrow> A \<and> B"  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
479  | 
by (match premises in H: A \<Rightarrow> \<open>intro conjI, rule H,  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
480  | 
match premises in H': B \<Rightarrow> \<open>rule H'\<close>\<close>)  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
481  | 
|
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
482  | 
text \<open>  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
483  | 
  In this example, the inner @{method match} can find the focused premise
 | 
| 69597 | 484  | 
  \<^term>\<open>B\<close>. In contrast, the @{method assumption} method would fail here due
 | 
485  | 
to \<^term>\<open>B\<close> not being logically accessible.  | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
486  | 
\<close>  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
487  | 
|
| 61576 | 488  | 
lemma "A \<Longrightarrow> A \<and> (B \<longrightarrow> B)"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
489  | 
by (match premises in H: A \<Rightarrow> \<open>intro conjI, rule H, rule impI,  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
490  | 
match premises (local) in A \<Rightarrow> \<open>fail\<close>  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
491  | 
\<bar> H': B \<Rightarrow> \<open>rule H'\<close>\<close>)  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
492  | 
|
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
493  | 
text \<open>  | 
| 69597 | 494  | 
In this example, the only premise that exists in the first focus is \<^term>\<open>A\<close>. Prior to the inner match, the rule \<open>impI\<close> changes the goal \<^term>\<open>B \<longrightarrow>  | 
495  | 
B\<close> into \<^term>\<open>B \<Longrightarrow> B\<close>. A standard premise match would also include \<^term>\<open>A\<close> as an original premise of the outer match. The \<open>local\<close> argument limits  | 
|
| 61576 | 496  | 
the match to newly focused premises.  | 
497  | 
\<close>  | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
498  | 
|
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
499  | 
|
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
500  | 
section \<open>Attributes\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
501  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
502  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
503  | 
Attributes may throw errors when applied to a given fact. For example, rule  | 
| 75160 | 504  | 
instantiation will fail if there is a type mismatch or if a given variable  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
505  | 
doesn't exist. Within a match or a method definition, it isn't generally  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
506  | 
possible to guarantee that applied attributes won't fail. For example, in  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
507  | 
the following method there is no guarantee that the two provided facts will  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
508  | 
necessarily compose.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
509  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
510  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
511  | 
method my_compose uses rule1 rule2 =  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
512  | 
(rule rule1 [OF rule2])  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
513  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
514  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
515  | 
  Some attributes (like @{attribute OF}) have been made partially
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
516  | 
Eisbach-aware. This means that they are able to form a closure despite not  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
517  | 
  necessarily always being applicable. In the case of @{attribute OF}, it is
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
518  | 
up to the proof author to guard attribute application with an appropriate  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
519  | 
  @{method match}, but there are still no static guarantees.
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
520  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
521  | 
  In contrast to @{attribute OF}, the @{attribute "where"} and @{attribute of}
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
522  | 
attributes attempt to provide static guarantees that they will apply  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
523  | 
whenever possible.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
524  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
525  | 
Within a match pattern for a fact, each outermost quantifier specifies the  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
526  | 
requirement that a matching fact must have a schematic variable at that  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
527  | 
point. This gives a corresponding name to this ``slot'' for the purposes of  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
528  | 
  forming a static closure, allowing the @{attribute "where"} attribute to
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
529  | 
perform an instantiation at run-time.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
530  | 
\<close>  | 
| 61417 | 531  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
532  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
533  | 
assumes A: "Q \<Longrightarrow> False"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
534  | 
shows "\<not> Q"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
535  | 
by (match intros in X: "\<And>P. (P \<Longrightarrow> False) \<Longrightarrow> \<not> P" \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
536  | 
\<open>rule X [where P = Q, OF A]\<close>)  | 
| 61417 | 537  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
538  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
539  | 
Subgoal focusing converts the outermost quantifiers of premises into  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
540  | 
schematics when lifting them to hypothetical facts. This allows us to  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
541  | 
  instantiate them with @{attribute "where"} when using an appropriate match
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
542  | 
pattern.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
543  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
544  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
545  | 
lemma "(\<And>x :: 'a. A x \<Longrightarrow> B x) \<Longrightarrow> A y \<Longrightarrow> B y"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
546  | 
by (match premises in I: "\<And>x :: 'a. ?P x \<Longrightarrow> ?Q x" \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
547  | 
\<open>rule I [where x = y]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
549  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
550  | 
  The @{attribute of} attribute behaves similarly. It is worth noting,
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
551  | 
  however, that the positional instantiation of @{attribute of} occurs against
 | 
| 61576 | 552  | 
the position of the variables as they are declared \<^emph>\<open>in the match pattern\<close>.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
553  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
554  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
555  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
556  | 
fixes A B and x :: 'a and y :: 'b  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
557  | 
assumes asm: "(\<And>x y. A y x \<Longrightarrow> B x y )"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
558  | 
shows "A y x \<Longrightarrow> B x y"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
559  | 
by (match asm in I: "\<And>(x :: 'a) (y :: 'b). ?P x y \<Longrightarrow> ?Q x y" \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
560  | 
\<open>rule I [of x y]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
561  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
562  | 
text \<open>  | 
| 61576 | 563  | 
In this example, the order of schematics in \<open>asm\<close> is actually \<open>?y ?x\<close>, but  | 
564  | 
we instantiate our matched rule in the opposite order. This is because the  | 
|
| 69597 | 565  | 
effective rule \<^term>\<open>I\<close> was bound from the match, which declared the \<^typ>\<open>'a\<close> slot first and the \<^typ>\<open>'b\<close> slot second.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
566  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
567  | 
  To get the dynamic behaviour of @{attribute of} we can choose to invoke it
 | 
| 61576 | 568  | 
\<^emph>\<open>unchecked\<close>. This avoids trying to do any type inference for the provided  | 
569  | 
parameters, instead storing them as their most general type and doing type  | 
|
570  | 
  matching at run-time. This, like @{attribute OF}, will throw errors if the
 | 
|
571  | 
expected slots don't exist or there is a type mismatch.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
572  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
573  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
574  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
575  | 
fixes A B and x :: 'a and y :: 'b  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
576  | 
assumes asm: "\<And>x y. A y x \<Longrightarrow> B x y"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
577  | 
shows "A y x \<Longrightarrow> B x y"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
578  | 
by (match asm in I: "PROP ?P" \<Rightarrow> \<open>rule I [of (unchecked) y x]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
579  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
580  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
581  | 
Attributes may be applied to matched facts directly as they are matched. Any  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
582  | 
declarations will therefore be applied in the context of the inner method,  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
583  | 
as well as any transformations to the rule.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
584  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
585  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
586  | 
lemma "(\<And>x :: 'a. A x \<Longrightarrow> B x) \<Longrightarrow> A y \<longrightarrow> B y"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
587  | 
by (match premises in I [of y, intros]: "\<And>x :: 'a. ?P x \<Longrightarrow> ?Q x" \<Rightarrow>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
588  | 
\<open>prop_solver\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
589  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
590  | 
text \<open>  | 
| 61576 | 591  | 
In this example, the pattern \<open>\<And>x :: 'a. ?P x \<Longrightarrow> ?Q x\<close> matches against the  | 
| 69597 | 592  | 
only premise, giving an appropriately typed slot for \<^term>\<open>y\<close>. After the  | 
593  | 
match, the resulting rule is instantiated to \<^term>\<open>y\<close> and then declared as  | 
|
| 61576 | 594  | 
  an @{attribute intros} rule. This is then picked up by @{method prop_solver}
 | 
595  | 
to solve the goal.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
596  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
597  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
598  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
599  | 
section \<open>Multi-match \label{sec:multi}\<close>
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
600  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
601  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
602  | 
  In all previous examples, @{method match} was only ever searching for a
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
603  | 
single rule or premise. Each local fact would therefore always have a length  | 
| 61576 | 604  | 
of exactly one. We may, however, wish to find \<^emph>\<open>all\<close> matching results. To  | 
605  | 
achieve this, we can simply mark a given pattern with the \<open>(multi)\<close>  | 
|
606  | 
argument.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
607  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
608  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
609  | 
lemma  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
610  | 
assumes asms: "A \<Longrightarrow> B" "A \<Longrightarrow> D"  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
611  | 
shows "(A \<longrightarrow> B) \<and> (A \<longrightarrow> D)"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
612  | 
apply (match asms in I [intros]: "?P \<Longrightarrow> ?Q" \<Rightarrow> \<open>solves \<open>prop_solver\<close>\<close>)?  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
613  | 
apply (match asms in I [intros]: "?P \<Longrightarrow> ?Q" (multi) \<Rightarrow> \<open>prop_solver\<close>)  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
614  | 
done  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
615  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
616  | 
text \<open>  | 
| 69597 | 617  | 
  In the first @{method match}, without the \<open>(multi)\<close> argument, \<^term>\<open>I\<close> is
 | 
| 61576 | 618  | 
only ever be bound to one of the members of \<open>asms\<close>. This backtracks over  | 
619  | 
both possibilities (see next section), however neither assumption in  | 
|
620  | 
  isolation is sufficient to solve to goal. The use of the @{method solves}
 | 
|
621  | 
  combinator ensures that @{method prop_solver} has no effect on the goal when
 | 
|
622  | 
it doesn't solve it, and so the first match leaves the goal unchanged. In  | 
|
623  | 
  the second @{method match}, \<open>I\<close> is bound to all of \<open>asms\<close>, declaring both
 | 
|
624  | 
  results as \<open>intros\<close>. With these rules @{method prop_solver} is capable of
 | 
|
625  | 
solving the goal.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
626  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
627  | 
Using for-fixed variables in patterns imposes additional constraints on the  | 
| 61576 | 628  | 
results. In all previous examples, the choice of using \<open>?P\<close> or a for-fixed  | 
| 69597 | 629  | 
\<^term>\<open>P\<close> only depended on whether or not \<^term>\<open>P\<close> was mentioned in another  | 
| 61576 | 630  | 
pattern or the inner method. When using a multi-match, however, all  | 
631  | 
for-fixed terms must agree in the results.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
632  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
633  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
634  | 
lemma  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
635  | 
assumes asms: "A \<Longrightarrow> B" "A \<Longrightarrow> D" "D \<Longrightarrow> B"  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
636  | 
shows "(A \<longrightarrow> B) \<and> (A \<longrightarrow> D)"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
637  | 
apply (match asms in I [intros]: "?P \<Longrightarrow> Q" (multi) for Q \<Rightarrow>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
638  | 
\<open>solves \<open>prop_solver\<close>\<close>)?  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
639  | 
apply (match asms in I [intros]: "P \<Longrightarrow> ?Q" (multi) for P \<Rightarrow>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
640  | 
\<open>prop_solver\<close>)  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
641  | 
done  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
642  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
643  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
644  | 
  Here we have two seemingly-equivalent applications of @{method match},
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
645  | 
however only the second one is capable of solving the goal. The first  | 
| 61576 | 646  | 
  @{method match} selects the first and third members of \<open>asms\<close> (those that
 | 
647  | 
  agree on their conclusion), which is not sufficient. The second @{method
 | 
|
648  | 
match} selects the first and second members of \<open>asms\<close> (those that agree on  | 
|
649  | 
  their assumption), which is enough for @{method prop_solver} to solve the
 | 
|
650  | 
goal.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
651  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
652  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
653  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
654  | 
section \<open>Dummy patterns\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
655  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
656  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
657  | 
Dummy patterns may be given as placeholders for unique schematics in  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
658  | 
patterns. They implicitly receive all currently bound variables as  | 
| 69597 | 659  | 
arguments, and are coerced into the \<^typ>\<open>prop\<close> type whenever possible. For  | 
| 61576 | 660  | 
example, the trivial dummy pattern \<open>_\<close> will match any proposition. In  | 
| 69597 | 661  | 
contrast, by default the pattern \<open>?P\<close> is considered to have type \<^typ>\<open>bool\<close>. It will not bind anything with meta-logical connectives (e.g. \<open>_ \<Longrightarrow> _\<close>  | 
| 61576 | 662  | 
or \<open>_ &&& _\<close>).  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
663  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
664  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
665  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
666  | 
assumes asms: "A &&& B \<Longrightarrow> D"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
667  | 
shows "(A \<and> B \<longrightarrow> D)"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
668  | 
by (match asms in I: _ \<Rightarrow> \<open>prop_solver intros: I conjunctionI\<close>)  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
669  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
670  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
671  | 
section \<open>Backtracking\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
672  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
673  | 
text \<open>  | 
| 61576 | 674  | 
Patterns are considered top-down, executing the inner method \<open>m\<close> of the  | 
675  | 
first pattern which is satisfied by the current match target. By default,  | 
|
676  | 
matching performs extensive backtracking by attempting all valid variable  | 
|
677  | 
and fact bindings according to the given pattern. In particular, all  | 
|
678  | 
unifiers for a given pattern will be explored, as well as each matching  | 
|
| 61493 | 679  | 
fact. The inner method \<open>m\<close> will be re-executed for each different  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
680  | 
variable/fact binding during backtracking. A successful match is considered  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
681  | 
a cut-point for backtracking. Specifically, once a match is made no other  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
682  | 
pattern-method pairs will be considered.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
683  | 
|
| 61576 | 684  | 
The method \<open>foo\<close> below fails for all goals that are conjunctions. Any such  | 
685  | 
goal will match the first pattern, causing the second pattern (that would  | 
|
686  | 
otherwise match all goals) to never be considered.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
687  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
688  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
689  | 
method foo =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
690  | 
(match conclusion in "?P \<and> ?Q" \<Rightarrow> \<open>fail\<close> \<bar> "?R" \<Rightarrow> \<open>prop_solver\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
691  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
692  | 
text \<open>  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
693  | 
The failure of an inner method that is executed after a successful match  | 
| 61576 | 694  | 
will cause the entire match to fail. This distinction is important due to  | 
695  | 
the pervasive use of backtracking. When a method is used in a combinator  | 
|
696  | 
chain, its failure becomes significant because it signals previously applied  | 
|
697  | 
  methods to move to the next result. Therefore, it is necessary for @{method
 | 
|
698  | 
match} to not mask such failure. One can always rewrite a match using the  | 
|
699  | 
combinators ``\<open>?\<close>'' and ``\<open>|\<close>'' to try subsequent patterns in the case of an  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
700  | 
inner-method failure. The following proof method, for example, always  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
701  | 
  invokes @{method prop_solver} for all goals because its first alternative
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
702  | 
either never matches or (if it does match) always fails.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
703  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
704  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
705  | 
method foo\<^sub>1 =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
706  | 
(match conclusion in "?P \<and> ?Q" \<Rightarrow> \<open>fail\<close>) |  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
707  | 
(match conclusion in "?R" \<Rightarrow> \<open>prop_solver\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
708  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
709  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
710  | 
subsection \<open>Cut\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
711  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
712  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
713  | 
Backtracking may be controlled more precisely by marking individual patterns  | 
| 61576 | 714  | 
as \<open>cut\<close>. This causes backtracking to not progress beyond this pattern: once  | 
715  | 
a match is found no others will be considered.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
716  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
717  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
718  | 
method foo\<^sub>2 =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
719  | 
(match premises in I: "P \<and> Q" (cut) and I': "P \<longrightarrow> ?U" for P Q \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
720  | 
\<open>rule mp [OF I' I [THEN conjunct1]]\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
721  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
722  | 
text \<open>  | 
| 69597 | 723  | 
In this example, once a conjunction is found (\<^term>\<open>P \<and> Q\<close>), all possible  | 
724  | 
implications of \<^term>\<open>P\<close> in the premises are considered, evaluating the  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
725  | 
  inner @{method rule} with each consequent. No other conjunctions will be
 | 
| 61576 | 726  | 
considered, with method failure occurring once all implications of the form  | 
727  | 
\<open>P \<longrightarrow> ?U\<close> have been explored. Here the left-right processing of individual  | 
|
728  | 
patterns is important, as all patterns after of the cut will maintain their  | 
|
729  | 
usual backtracking behaviour.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
730  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
731  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
732  | 
lemma "A \<and> B \<Longrightarrow> A \<longrightarrow> D \<Longrightarrow> A \<longrightarrow> C \<Longrightarrow> C"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
733  | 
by foo\<^sub>2  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
734  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
735  | 
lemma "C \<and> D \<Longrightarrow> A \<and> B \<Longrightarrow> A \<longrightarrow> C \<Longrightarrow> C"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
736  | 
by (foo\<^sub>2 | prop_solver)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
737  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
738  | 
text \<open>  | 
| 61576 | 739  | 
In this example, the first lemma is solved by \<open>foo\<^sub>2\<close>, by first picking  | 
| 69597 | 740  | 
\<^term>\<open>A \<longrightarrow> D\<close> for \<open>I'\<close>, then backtracking and ultimately succeeding after  | 
741  | 
picking \<^term>\<open>A \<longrightarrow> C\<close>. In the second lemma, however, \<^term>\<open>C \<and> D\<close> is  | 
|
| 61576 | 742  | 
matched first, the second pattern in the match cannot be found and so the  | 
743  | 
  method fails, falling through to @{method prop_solver}.
 | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
744  | 
|
| 61576 | 745  | 
More precise control is also possible by giving a positive number \<open>n\<close> as an  | 
746  | 
argument to \<open>cut\<close>. This will limit the number of backtracking results of  | 
|
747  | 
that match to be at most \<open>n\<close>. The match argument \<open>(cut 1)\<close> is the same as  | 
|
748  | 
simply \<open>(cut)\<close>.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
749  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
750  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
751  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
752  | 
subsection \<open>Multi-match revisited\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
753  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
754  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
755  | 
A multi-match will produce a sequence of potential bindings for for-fixed  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
756  | 
variables, where each binding environment is the result of matching against  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
757  | 
at least one element from the match target. For each environment, the match  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
758  | 
result will be all elements of the match target which agree with the pattern  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
759  | 
under that environment. This can result in unexpected behaviour when giving  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
760  | 
very general patterns.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
761  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
762  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
763  | 
lemma  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
764  | 
assumes asms: "\<And>x. A x \<and> B x" "\<And>y. A y \<and> C y" "\<And>z. B z \<and> C z"  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
765  | 
shows "A x \<and> C x"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
766  | 
by (match asms in I: "\<And>x. P x \<and> ?Q x" (multi) for P \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
767  | 
\<open>match (P) in "A" \<Rightarrow> \<open>fail\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
768  | 
\<bar> _ \<Rightarrow> \<open>match I in "\<And>x. A x \<and> B x" \<Rightarrow> \<open>fail\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
769  | 
\<bar> _ \<Rightarrow> \<open>rule I\<close>\<close>\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
770  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
771  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
772  | 
Intuitively it seems like this proof should fail to check. The first match  | 
| 69597 | 773  | 
result, which binds \<^term>\<open>I\<close> to the first two members of \<open>asms\<close>, fails the  | 
774  | 
second inner match due to binding \<^term>\<open>P\<close> to \<^term>\<open>A\<close>. Backtracking then  | 
|
775  | 
attempts to bind \<^term>\<open>I\<close> to the third member of \<open>asms\<close>. This passes all  | 
|
| 61576 | 776  | 
  inner matches, but fails when @{method rule} cannot successfully apply this
 | 
777  | 
to the current goal. After this, a valid match that is produced by the  | 
|
| 69597 | 778  | 
unifier is one which binds \<^term>\<open>P\<close> to simply \<open>\<lambda>a. A ?x\<close>. The first inner  | 
779  | 
match succeeds because \<open>\<lambda>a. A ?x\<close> does not match \<^term>\<open>A\<close>. The next inner  | 
|
780  | 
match succeeds because \<^term>\<open>I\<close> has only been bound to the first member of  | 
|
| 61576 | 781  | 
  \<open>asms\<close>. This is due to @{method match} considering \<open>\<lambda>a. A ?x\<close> and \<open>\<lambda>a. A ?y\<close>
 | 
782  | 
as distinct terms.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
783  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
784  | 
The simplest way to address this is to explicitly disallow term bindings  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
785  | 
which we would consider invalid.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
786  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
787  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
788  | 
method abs_used for P =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
789  | 
(match (P) in "\<lambda>a. ?P" \<Rightarrow> \<open>fail\<close> \<bar> _ \<Rightarrow> \<open>-\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
790  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
791  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
792  | 
This method has no effect on the goal state, but instead serves as a filter  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
793  | 
on the environment produced from match.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
794  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
795  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
796  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
797  | 
section \<open>Uncurrying\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
798  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
799  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
800  | 
  The @{method match} method is not aware of the logical content of match
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
801  | 
targets. Each pattern is simply matched against the shallow structure of a  | 
| 61576 | 802  | 
fact or term. Most facts are in \<^emph>\<open>normal form\<close>, which curries premises via  | 
803  | 
meta-implication \<open>_ \<Longrightarrow> _\<close>.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
804  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
805  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
806  | 
lemma  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
807  | 
assumes asms: "D \<Longrightarrow> B \<Longrightarrow> C" "D \<Longrightarrow> A"  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
808  | 
shows "D \<Longrightarrow> B \<Longrightarrow> C \<and> A"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
809  | 
by (match asms in H: "D \<Longrightarrow> _" (multi) \<Rightarrow> \<open>prop_solver elims: H\<close>)  | 
| 61417 | 810  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
811  | 
text \<open>  | 
| 61493 | 812  | 
For the first member of \<open>asms\<close> the dummy pattern successfully matches  | 
| 69597 | 813  | 
against \<^term>\<open>B \<Longrightarrow> C\<close> and so the proof is successful.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
814  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
815  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
816  | 
lemma  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
817  | 
assumes asms: "A \<Longrightarrow> B \<Longrightarrow> C" "D \<Longrightarrow> C"  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
818  | 
shows "D \<or> (A \<and> B) \<Longrightarrow> C"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
819  | 
apply (match asms in H: "_ \<Longrightarrow> C" (multi) \<Rightarrow> \<open>prop_solver elims: H\<close>)(*<*)?  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
820  | 
apply (prop_solver elims: asms)  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
821  | 
done(*>*)  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
822  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
823  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
824  | 
This proof will fail to solve the goal. Our match pattern will only match  | 
| 69597 | 825  | 
rules which have a single premise, and conclusion \<^term>\<open>C\<close>, so the first  | 
| 61576 | 826  | 
member of \<open>asms\<close> is not bound and thus the proof fails. Matching a pattern  | 
| 69597 | 827  | 
of the form \<^term>\<open>P \<Longrightarrow> Q\<close> against this fact will bind \<^term>\<open>P\<close> to  | 
828  | 
\<^term>\<open>A\<close> and \<^term>\<open>Q\<close> to \<^term>\<open>B \<Longrightarrow> C\<close>. Our pattern, with a concrete  | 
|
829  | 
\<^term>\<open>C\<close> in the conclusion, will fail to match this fact.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
830  | 
|
| 61576 | 831  | 
To express our desired match, we may \<^emph>\<open>uncurry\<close> our rules before matching  | 
832  | 
against them. This forms a meta-conjunction of all premises in a fact, so  | 
|
833  | 
that only one implication remains. For example the uncurried version of  | 
|
| 69597 | 834  | 
\<^term>\<open>A \<Longrightarrow> B \<Longrightarrow> C\<close> is \<^term>\<open>A &&& B \<Longrightarrow> C\<close>. This will now match our  | 
| 61576 | 835  | 
desired pattern \<open>_ \<Longrightarrow> C\<close>, and can be \<^emph>\<open>curried\<close> after the match to put it  | 
836  | 
back into normal form.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
837  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
838  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
839  | 
lemma  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
840  | 
assumes asms: "A \<Longrightarrow> B \<Longrightarrow> C" "D \<Longrightarrow> C"  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
841  | 
shows "D \<or> (A \<and> B) \<Longrightarrow> C"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
842  | 
by (match asms [uncurry] in H [curry]: "_ \<Longrightarrow> C" (multi) \<Rightarrow>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
843  | 
\<open>prop_solver elims: H\<close>)  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
844  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
845  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
846  | 
section \<open>Reverse matching\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
847  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
848  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
849  | 
  The @{method match} method only attempts to perform matching of the pattern
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
850  | 
against the match target. Specifically this means that it will not  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
851  | 
instantiate schematic terms in the match target.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
852  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
853  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
854  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
855  | 
assumes asms: "\<And>x :: 'a. A x"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
856  | 
shows "A y"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
857  | 
apply (match asms in H: "A y" \<Rightarrow> \<open>rule H\<close>)?  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
858  | 
apply (match asms in H: P for P \<Rightarrow>  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
859  | 
          \<open>match ("A y") in P \<Rightarrow> \<open>rule H\<close>\<close>)
 | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
860  | 
done  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
861  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
862  | 
text \<open>  | 
| 61576 | 863  | 
  In the first @{method match} we attempt to find a member of \<open>asms\<close> which
 | 
864  | 
matches our goal precisely. This fails due to no such member existing. The  | 
|
865  | 
second match reverses the role of the fact in the match, by first giving a  | 
|
| 69597 | 866  | 
general pattern \<^term>\<open>P\<close>. This bound pattern is then matched against \<^term>\<open>A y\<close>. In this case, \<^term>\<open>P\<close> is bound to \<open>A ?x\<close> and so it successfully  | 
| 61576 | 867  | 
matches.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
868  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
869  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
870  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
871  | 
section \<open>Type matching\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
872  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
873  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
874  | 
  The rule instantiation attributes @{attribute "where"} and @{attribute "of"}
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
875  | 
attempt to guarantee type-correctness wherever possible. This can require  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
876  | 
  additional invocations of @{method match} in order to statically ensure that
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
877  | 
instantiation will succeed.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
878  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
879  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
880  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
881  | 
assumes asms: "\<And>x :: 'a. A x"  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
882  | 
shows "A y"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
883  | 
by (match asms in H: "\<And>z :: 'b. P z" for P \<Rightarrow>  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
884  | 
\<open>match (y) in "y :: 'b" for y \<Rightarrow> \<open>rule H [where z = y]\<close>\<close>)  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
885  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
886  | 
text \<open>  | 
| 61576 | 887  | 
In this example the type \<open>'b\<close> is matched to \<open>'a\<close>, however statically they  | 
888  | 
are formally distinct types. The first match binds \<open>'b\<close> while the inner  | 
|
| 69597 | 889  | 
match serves to coerce \<^term>\<open>y\<close> into having the type \<open>'b\<close>. This allows the  | 
| 61576 | 890  | 
rule instantiation to successfully apply.  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
891  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
892  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
893  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
894  | 
chapter \<open>Method development\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
895  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
896  | 
section \<open>Tracing methods\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
897  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
898  | 
text \<open>  | 
| 69597 | 899  | 
  Method tracing is supported by auxiliary print methods provided by \<^theory>\<open>HOL-Eisbach.Eisbach_Tools\<close>. These include @{method print_fact}, @{method
 | 
| 68484 | 900  | 
  print_term} and @{method print_type}. Whenever a print method is evaluated
 | 
901  | 
it leaves the goal unchanged and writes its argument as tracing output.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
902  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
903  | 
  Print methods can be combined with the @{method fail} method to investigate
 | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
904  | 
the backtracking behaviour of a method.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
905  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
906  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
907  | 
lemma  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
908  | 
assumes asms: A B C D  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
909  | 
shows D  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
910  | 
apply (match asms in H: _ \<Rightarrow> \<open>print_fact H, fail\<close>)(*<*)?  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
911  | 
apply (simp add: asms)  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
912  | 
done(*>*)  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
913  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
914  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
915  | 
This proof will fail, but the tracing output will show the order that the  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
916  | 
assumptions are attempted.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
917  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
918  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
919  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
920  | 
section \<open>Integrating with Isabelle/ML\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
921  | 
|
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
922  | 
subsubsection \<open>Attributes\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
923  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
924  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
925  | 
A custom rule attribute is a simple way to extend the functionality of  | 
| 61576 | 926  | 
Eisbach methods. The dummy rule attribute notation (\<open>[[ _ ]]\<close>) invokes the  | 
927  | 
given attribute against a dummy fact and evaluates to the result of that  | 
|
928  | 
attribute. When used as a match target, this can serve as an effective  | 
|
929  | 
auxiliary function.  | 
|
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
930  | 
\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
931  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
932  | 
attribute_setup get_split_rule =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
933  | 
\<open>Args.term >> (fn t =>  | 
| 
61853
 
fb7756087101
rule_attribute and declaration_attribute implicitly support abstract closure, but mixed_attribute implementations need to be aware of Thm.is_free_dummy;
 
wenzelm 
parents: 
61656 
diff
changeset
 | 
934  | 
Thm.rule_attribute [] (fn context => fn _ =>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
935  | 
(case get_split_rule (Context.proof_of context) t of  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
936  | 
SOME thm => thm  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
937  | 
| NONE => Drule.dummy_thm)))\<close>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
938  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
939  | 
text \<open>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
940  | 
  In this example, the new attribute @{attribute get_split_rule} lifts the ML
 | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
941  | 
function of the same name into an attribute. When applied to a case  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
942  | 
distinction over a datatype, it retrieves its corresponding split rule.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
943  | 
|
| 75160 | 944  | 
We can then integrate this into a method that applies the split rule, first  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
945  | 
matching to ensure that fetching the rule was successful.  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
946  | 
\<close>  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
947  | 
(*<*)declare TrueI [intros](*>*)  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
948  | 
method splits =  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
949  | 
(match conclusion in "?P f" for f \<Rightarrow>  | 
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
950  | 
\<open>match [[get_split_rule f]] in U: "(_ :: bool) = _" \<Rightarrow>  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
951  | 
\<open>rule U [THEN iffD2]\<close>\<close>)  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
952  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
953  | 
lemma "L \<noteq> [] \<Longrightarrow> case L of [] \<Rightarrow> False | _ \<Rightarrow> True"  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
954  | 
apply splits  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
955  | 
apply (prop_solver intros: allI)  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
956  | 
done  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
957  | 
|
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
958  | 
text \<open>  | 
| 
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
959  | 
  Here the new @{method splits} method transforms the goal to use only logical
 | 
| 69597 | 960  | 
connectives: \<^term>\<open>L = [] \<longrightarrow> False \<and> (\<forall>x y. L = x # y \<longrightarrow> True)\<close>. This goal  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
961  | 
  is then in a form solvable by @{method prop_solver} when given the universal
 | 
| 61493 | 962  | 
quantifier introduction rule \<open>allI\<close>.  | 
| 
60298
 
7c278b692aae
updated Eisbach manual, using version 3149f9146eb5 of its Bitbucket repository;
 
wenzelm 
parents: 
60288 
diff
changeset
 | 
963  | 
\<close>  | 
| 
60288
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
964  | 
|
| 
 
d7f636331176
added Eisbach manual, using version 8845c4cb28b6 of its Bitbucket repository;
 
wenzelm 
parents:  
diff
changeset
 | 
965  | 
end  |