| author | Andreas Lochbihler |
| Tue, 01 Dec 2015 12:35:11 +0100 | |
| changeset 61766 | 507b39df1a57 |
| parent 61656 | cfabbc083977 |
| child 61854 | 38b049cd3aad |
| permissions | -rw-r--r-- |
| 61656 | 1 |
(*:maxLineLen=78:*) |
2 |
||
| 46295 | 3 |
theory Eq |
4 |
imports Base |
|
5 |
begin |
|
6 |
||
| 58618 | 7 |
chapter \<open>Equational reasoning\<close> |
| 46295 | 8 |
|
| 58618 | 9 |
text \<open>Equality is one of the most fundamental concepts of |
| 46295 | 10 |
mathematics. The Isabelle/Pure logic (\chref{ch:logic}) provides a
|
| 61493 | 11 |
builtin relation \<open>\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop\<close> that expresses equality |
| 46295 | 12 |
of arbitrary terms (or propositions) at the framework level, as |
13 |
expressed by certain basic inference rules (\secref{sec:eq-rules}).
|
|
14 |
||
15 |
Equational reasoning means to replace equals by equals, using |
|
16 |
reflexivity and transitivity to form chains of replacement steps, |
|
17 |
and congruence rules to access sub-structures. Conversions |
|
18 |
(\secref{sec:conv}) provide a convenient framework to compose basic
|
|
19 |
equational steps to build specific equational reasoning tools. |
|
20 |
||
21 |
Higher-order matching is able to provide suitable instantiations for |
|
22 |
giving equality rules, which leads to the versatile concept of |
|
| 61493 | 23 |
\<open>\<lambda>\<close>-term rewriting (\secref{sec:rewriting}). Internally
|
| 46295 | 24 |
this is based on the general-purpose Simplifier engine of Isabelle, |
25 |
which is more specific and more efficient than plain conversions. |
|
26 |
||
27 |
Object-logics usually introduce specific notions of equality or |
|
28 |
equivalence, and relate it with the Pure equality. This enables to |
|
29 |
re-use the Pure tools for equational reasoning for particular |
|
30 |
object-logic connectives as well. |
|
| 58618 | 31 |
\<close> |
| 46295 | 32 |
|
33 |
||
| 58618 | 34 |
section \<open>Basic equality rules \label{sec:eq-rules}\<close>
|
| 46295 | 35 |
|
| 61493 | 36 |
text \<open>Isabelle/Pure uses \<open>\<equiv>\<close> for equality of arbitrary |
| 50085 | 37 |
terms, which includes equivalence of propositions of the logical |
| 61493 | 38 |
framework. The conceptual axiomatization of the constant \<open>\<equiv> |
39 |
:: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow> prop\<close> is given in \figref{fig:pure-equality}. The
|
|
| 50085 | 40 |
inference kernel presents slightly different equality rules, which |
41 |
may be understood as derived rules from this minimal axiomatization. |
|
42 |
The Pure theory also provides some theorems that express the same |
|
43 |
reasoning schemes as theorems that can be composed like object-level |
|
44 |
rules as explained in \secref{sec:obj-rules}.
|
|
45 |
||
46 |
For example, @{ML Thm.symmetric} as Pure inference is an ML function
|
|
| 61493 | 47 |
that maps a theorem \<open>th\<close> stating \<open>t \<equiv> u\<close> to one |
48 |
stating \<open>u \<equiv> t\<close>. In contrast, @{thm [source]
|
|
| 50085 | 49 |
Pure.symmetric} as Pure theorem expresses the same reasoning in |
| 61493 | 50 |
declarative form. If used like \<open>th [THEN Pure.symmetric]\<close> |
| 50085 | 51 |
in Isar source notation, it achieves a similar effect as the ML |
52 |
inference function, although the rule attribute @{attribute THEN} or
|
|
53 |
ML operator @{ML "op RS"} involve the full machinery of higher-order
|
|
| 61493 | 54 |
unification (modulo \<open>\<beta>\<eta>\<close>-conversion) and lifting of \<open>\<And>/\<Longrightarrow>\<close> contexts.\<close> |
| 50085 | 55 |
|
| 58618 | 56 |
text %mlref \<open> |
| 50085 | 57 |
\begin{mldecls}
|
58 |
@{index_ML Thm.reflexive: "cterm -> thm"} \\
|
|
59 |
@{index_ML Thm.symmetric: "thm -> thm"} \\
|
|
60 |
@{index_ML Thm.transitive: "thm -> thm -> thm"} \\
|
|
61 |
@{index_ML Thm.abstract_rule: "string -> cterm -> thm -> thm"} \\
|
|
62 |
@{index_ML Thm.combination: "thm -> thm -> thm"} \\[0.5ex]
|
|
63 |
@{index_ML Thm.equal_intr: "thm -> thm -> thm"} \\
|
|
64 |
@{index_ML Thm.equal_elim: "thm -> thm -> thm"} \\
|
|
65 |
\end{mldecls}
|
|
66 |
||
| 53982 | 67 |
See also @{file "~~/src/Pure/thm.ML" } for further description of
|
| 61493 | 68 |
these inference rules, and a few more for primitive \<open>\<beta>\<close> and |
69 |
\<open>\<eta>\<close> conversions. Note that \<open>\<alpha>\<close> conversion is |
|
| 50085 | 70 |
implicit due to the representation of terms with de-Bruijn indices |
| 58618 | 71 |
(\secref{sec:terms}).\<close>
|
| 46295 | 72 |
|
73 |
||
| 58618 | 74 |
section \<open>Conversions \label{sec:conv}\<close>
|
| 46295 | 75 |
|
| 58618 | 76 |
text \<open> |
| 52422 | 77 |
%FIXME |
| 50122 | 78 |
|
79 |
The classic article that introduces the concept of conversion (for |
|
| 58555 | 80 |
Cambridge LCF) is @{cite "paulson:1983"}.
|
| 58618 | 81 |
\<close> |
| 46295 | 82 |
|
83 |
||
| 58618 | 84 |
section \<open>Rewriting \label{sec:rewriting}\<close>
|
| 46295 | 85 |
|
| 58618 | 86 |
text \<open>Rewriting normalizes a given term (theorem or goal) by |
| 61493 | 87 |
replacing instances of given equalities \<open>t \<equiv> u\<close> in subterms. |
| 46295 | 88 |
Rewriting continues until no rewrites are applicable to any subterm. |
| 61493 | 89 |
This may be used to unfold simple definitions of the form \<open>f |
90 |
x\<^sub>1 \<dots> x\<^sub>n \<equiv> u\<close>, but is slightly more general than that. |
|
| 58618 | 91 |
\<close> |
| 46295 | 92 |
|
| 58618 | 93 |
text %mlref \<open> |
| 46295 | 94 |
\begin{mldecls}
|
|
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53982
diff
changeset
|
95 |
@{index_ML rewrite_rule: "Proof.context -> thm list -> thm -> thm"} \\
|
|
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53982
diff
changeset
|
96 |
@{index_ML rewrite_goals_rule: "Proof.context -> thm list -> thm -> thm"} \\
|
|
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53982
diff
changeset
|
97 |
@{index_ML rewrite_goal_tac: "Proof.context -> thm list -> int -> tactic"} \\
|
|
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53982
diff
changeset
|
98 |
@{index_ML rewrite_goals_tac: "Proof.context -> thm list -> tactic"} \\
|
|
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
53982
diff
changeset
|
99 |
@{index_ML fold_goals_tac: "Proof.context -> thm list -> tactic"} \\
|
| 46295 | 100 |
\end{mldecls}
|
101 |
||
| 61493 | 102 |
\<^descr> @{ML rewrite_rule}~\<open>ctxt rules thm\<close> rewrites the whole
|
| 46486 | 103 |
theorem by the given rules. |
104 |
||
| 61493 | 105 |
\<^descr> @{ML rewrite_goals_rule}~\<open>ctxt rules thm\<close> rewrites the
|
| 46486 | 106 |
outer premises of the given theorem. Interpreting the same as a |
107 |
goal state (\secref{sec:tactical-goals}) it means to rewrite all
|
|
108 |
subgoals (in the same manner as @{ML rewrite_goals_tac}).
|
|
109 |
||
| 61493 | 110 |
\<^descr> @{ML rewrite_goal_tac}~\<open>ctxt rules i\<close> rewrites subgoal
|
111 |
\<open>i\<close> by the given rewrite rules. |
|
| 46295 | 112 |
|
| 61493 | 113 |
\<^descr> @{ML rewrite_goals_tac}~\<open>ctxt rules\<close> rewrites all subgoals
|
| 46295 | 114 |
by the given rewrite rules. |
115 |
||
| 61493 | 116 |
\<^descr> @{ML fold_goals_tac}~\<open>ctxt rules\<close> essentially uses @{ML
|
117 |
rewrite_goals_tac} with the symmetric form of each member of \<open>rules\<close>, re-ordered to fold longer expression first. This supports |
|
| 46295 | 118 |
to idea to fold primitive definitions that appear in expended form |
119 |
in the proof state. |
|
| 58618 | 120 |
\<close> |
| 46295 | 121 |
|
122 |
end |