| author | desharna | 
| Tue, 04 Mar 2025 16:38:21 +0100 | |
| changeset 82239 | 1b7dc0728f5c | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 61656 | 1  | 
(*:maxLineLen=78:*)  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
|
| 58889 | 3  | 
section \<open>Example: First-Order Logic\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
theory %visible First_Order_Logic  | 
| 42651 | 6  | 
imports Base (* FIXME Pure!? *)  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
begin  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
|
| 58618 | 9  | 
text \<open>  | 
| 62279 | 10  | 
In order to commence a new object-logic within Isabelle/Pure we introduce  | 
11  | 
abstract syntactic categories \<open>i\<close> for individuals and \<open>o\<close> for  | 
|
12  | 
object-propositions. The latter is embedded into the language of Pure  | 
|
13  | 
propositions by means of a separate judgment.  | 
|
| 58618 | 14  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
15  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
16  | 
typedecl i  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
typedecl o  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
18  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
19  | 
judgment Trueprop :: "o \<Rightarrow> prop" (\<open>_\<close> 5)  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
20  | 
|
| 58618 | 21  | 
text \<open>  | 
| 62279 | 22  | 
Note that the object-logic judgment is implicit in the syntax: writing  | 
| 69593 | 23  | 
\<^prop>\<open>A\<close> produces \<^term>\<open>Trueprop A\<close> internally. From the Pure  | 
24  | 
perspective this means ``\<^prop>\<open>A\<close> is derivable in the object-logic''.  | 
|
| 58618 | 25  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
27  | 
|
| 58618 | 28  | 
subsection \<open>Equational reasoning \label{sec:framework-ex-equal}\<close>
 | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
29  | 
|
| 58618 | 30  | 
text \<open>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
Equality is axiomatized as a binary predicate on individuals, with  | 
| 62279 | 32  | 
reflexivity as introduction, and substitution as elimination principle. Note  | 
33  | 
that the latter is particularly convenient in a framework like Isabelle,  | 
|
34  | 
because syntactic congruences are implicitly produced by unification of  | 
|
35  | 
\<open>B x\<close> against expressions containing occurrences of \<open>x\<close>.  | 
|
| 58618 | 36  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
38  | 
axiomatization equal :: "i \<Rightarrow> i \<Rightarrow> o" (infix \<open>=\<close> 50)  | 
| 62279 | 39  | 
where refl [intro]: "x = x"  | 
40  | 
and subst [elim]: "x = y \<Longrightarrow> B x \<Longrightarrow> B y"  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
41  | 
|
| 58618 | 42  | 
text \<open>  | 
| 62279 | 43  | 
Substitution is very powerful, but also hard to control in full generality.  | 
| 69593 | 44  | 
We derive some common symmetry~/ transitivity schemes of \<^term>\<open>equal\<close> as  | 
| 62279 | 45  | 
particular consequences.  | 
| 58618 | 46  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
47  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
48  | 
theorem sym [sym]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
49  | 
assumes "x = y"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
shows "y = x"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
proof -  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
have "x = x" ..  | 
| 58618 | 53  | 
with \<open>x = y\<close> show "y = x" ..  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
54  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
55  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
56  | 
theorem forw_subst [trans]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
57  | 
assumes "y = x" and "B x"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
58  | 
shows "B y"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
59  | 
proof -  | 
| 58618 | 60  | 
from \<open>y = x\<close> have "x = y" ..  | 
61  | 
from this and \<open>B x\<close> show "B y" ..  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
62  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
63  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
64  | 
theorem back_subst [trans]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
65  | 
assumes "B x" and "x = y"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
66  | 
shows "B y"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
67  | 
proof -  | 
| 58618 | 68  | 
from \<open>x = y\<close> and \<open>B x\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
69  | 
show "B y" ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
70  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
71  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
72  | 
theorem trans [trans]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
73  | 
assumes "x = y" and "y = z"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
74  | 
shows "x = z"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
75  | 
proof -  | 
| 58618 | 76  | 
from \<open>y = z\<close> and \<open>x = y\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
77  | 
show "x = z" ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
78  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
79  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
80  | 
|
| 58618 | 81  | 
subsection \<open>Basic group theory\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
82  | 
|
| 58618 | 83  | 
text \<open>  | 
| 62279 | 84  | 
As an example for equational reasoning we consider some bits of group  | 
85  | 
theory. The subsequent locale definition postulates group operations and  | 
|
86  | 
axioms; we also derive some consequences of this specification.  | 
|
| 58618 | 87  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
88  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
89  | 
locale group =  | 
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
90  | 
fixes prod :: "i \<Rightarrow> i \<Rightarrow> i" (infix \<open>\<circ>\<close> 70)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
91  | 
and inv :: "i \<Rightarrow> i" (\<open>(_\<inverse>)\<close> [1000] 999)  | 
| 
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
92  | 
and unit :: i (\<open>1\<close>)  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
93  | 
assumes assoc: "(x \<circ> y) \<circ> z = x \<circ> (y \<circ> z)"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
94  | 
and left_unit: "1 \<circ> x = x"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
95  | 
and left_inv: "x\<inverse> \<circ> x = 1"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
96  | 
begin  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
97  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
98  | 
theorem right_inv: "x \<circ> x\<inverse> = 1"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
99  | 
proof -  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
100  | 
have "x \<circ> x\<inverse> = 1 \<circ> (x \<circ> x\<inverse>)" by (rule left_unit [symmetric])  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
101  | 
also have "\<dots> = (1 \<circ> x) \<circ> x\<inverse>" by (rule assoc [symmetric])  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
102  | 
also have "1 = (x\<inverse>)\<inverse> \<circ> x\<inverse>" by (rule left_inv [symmetric])  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
103  | 
also have "\<dots> \<circ> x = (x\<inverse>)\<inverse> \<circ> (x\<inverse> \<circ> x)" by (rule assoc)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
104  | 
also have "x\<inverse> \<circ> x = 1" by (rule left_inv)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
105  | 
also have "((x\<inverse>)\<inverse> \<circ> \<dots>) \<circ> x\<inverse> = (x\<inverse>)\<inverse> \<circ> (1 \<circ> x\<inverse>)" by (rule assoc)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
106  | 
also have "1 \<circ> x\<inverse> = x\<inverse>" by (rule left_unit)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
107  | 
also have "(x\<inverse>)\<inverse> \<circ> \<dots> = 1" by (rule left_inv)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
108  | 
finally show "x \<circ> x\<inverse> = 1" .  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
109  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
111  | 
theorem right_unit: "x \<circ> 1 = x"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
112  | 
proof -  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
113  | 
have "1 = x\<inverse> \<circ> x" by (rule left_inv [symmetric])  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
114  | 
also have "x \<circ> \<dots> = (x \<circ> x\<inverse>) \<circ> x" by (rule assoc [symmetric])  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
115  | 
also have "x \<circ> x\<inverse> = 1" by (rule right_inv)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
116  | 
also have "\<dots> \<circ> x = x" by (rule left_unit)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
117  | 
finally show "x \<circ> 1 = x" .  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
118  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
119  | 
|
| 58618 | 120  | 
text \<open>  | 
| 62279 | 121  | 
Reasoning from basic axioms is often tedious. Our proofs work by producing  | 
122  | 
various instances of the given rules (potentially the symmetric form) using  | 
|
123  | 
the pattern ``\<^theory_text>\<open>have eq by (rule r)\<close>'' and composing the chain of results  | 
|
124  | 
via \<^theory_text>\<open>also\<close>/\<^theory_text>\<open>finally\<close>. These steps may involve any of the transitivity  | 
|
125  | 
  rules declared in \secref{sec:framework-ex-equal}, namely @{thm trans} in
 | 
|
126  | 
  combining the first two results in @{thm right_inv} and in the final steps
 | 
|
127  | 
  of both proofs, @{thm forw_subst} in the first combination of @{thm
 | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
128  | 
  right_unit}, and @{thm back_subst} in all other calculational steps.
 | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
129  | 
|
| 62279 | 130  | 
Occasional substitutions in calculations are adequate, but should not be  | 
131  | 
over-emphasized. The other extreme is to compose a chain by plain  | 
|
132  | 
transitivity only, with replacements occurring always in topmost position.  | 
|
133  | 
For example:  | 
|
| 58618 | 134  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
136  | 
(*<*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
137  | 
theorem "\<And>A. PROP A \<Longrightarrow> PROP A"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
138  | 
proof -  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
139  | 
assume [symmetric, defn]: "\<And>x y. (x \<equiv> y) \<equiv> Trueprop (x = y)"  | 
| 56594 | 140  | 
fix x  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
141  | 
(*>*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
142  | 
have "x \<circ> 1 = x \<circ> (x\<inverse> \<circ> x)" unfolding left_inv ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
143  | 
also have "\<dots> = (x \<circ> x\<inverse>) \<circ> x" unfolding assoc ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
144  | 
also have "\<dots> = 1 \<circ> x" unfolding right_inv ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
145  | 
also have "\<dots> = x" unfolding left_unit ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
146  | 
finally have "x \<circ> 1 = x" .  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
147  | 
(*<*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
148  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
149  | 
(*>*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
150  | 
|
| 58618 | 151  | 
text \<open>  | 
| 62279 | 152  | 
Here we have re-used the built-in mechanism for unfolding definitions in  | 
153  | 
order to normalize each equational problem. A more realistic object-logic  | 
|
154  | 
  would include proper setup for the Simplifier (\secref{sec:simplifier}), the
 | 
|
155  | 
main automated tool for equational reasoning in Isabelle. Then ``\<^theory_text>\<open>unfolding  | 
|
156  | 
left_inv ..\<close>'' would become ``\<^theory_text>\<open>by (simp only: left_inv)\<close>'' etc.  | 
|
| 58618 | 157  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
159  | 
end  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
160  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
161  | 
|
| 58618 | 162  | 
subsection \<open>Propositional logic \label{sec:framework-ex-prop}\<close>
 | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
163  | 
|
| 58618 | 164  | 
text \<open>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
165  | 
We axiomatize basic connectives of propositional logic: implication,  | 
| 62279 | 166  | 
disjunction, and conjunction. The associated rules are modeled after  | 
| 76987 | 167  | 
Gentzen's system of Natural Deduction \<^cite>\<open>"Gentzen:1935"\<close>.  | 
| 58618 | 168  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
169  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
170  | 
axiomatization imp :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<longrightarrow>\<close> 25)  | 
| 62279 | 171  | 
where impI [intro]: "(A \<Longrightarrow> B) \<Longrightarrow> A \<longrightarrow> B"  | 
172  | 
and impD [dest]: "(A \<longrightarrow> B) \<Longrightarrow> A \<Longrightarrow> B"  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
173  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
174  | 
axiomatization disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<or>\<close> 30)  | 
| 62279 | 175  | 
where disjI\<^sub>1 [intro]: "A \<Longrightarrow> A \<or> B"  | 
176  | 
and disjI\<^sub>2 [intro]: "B \<Longrightarrow> A \<or> B"  | 
|
177  | 
and disjE [elim]: "A \<or> B \<Longrightarrow> (A \<Longrightarrow> C) \<Longrightarrow> (B \<Longrightarrow> C) \<Longrightarrow> C"  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
178  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
179  | 
axiomatization conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr \<open>\<and>\<close> 35)  | 
| 62279 | 180  | 
where conjI [intro]: "A \<Longrightarrow> B \<Longrightarrow> A \<and> B"  | 
181  | 
and conjD\<^sub>1: "A \<and> B \<Longrightarrow> A"  | 
|
182  | 
and conjD\<^sub>2: "A \<and> B \<Longrightarrow> B"  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
183  | 
|
| 58618 | 184  | 
text \<open>  | 
| 69593 | 185  | 
The conjunctive destructions have the disadvantage that decomposing \<^prop>\<open>A \<and> B\<close> involves an immediate decision which component should be projected.  | 
186  | 
The more convenient simultaneous elimination \<^prop>\<open>A \<and> B \<Longrightarrow> (A \<Longrightarrow> B \<Longrightarrow> C) \<Longrightarrow>  | 
|
187  | 
C\<close> can be derived as follows:  | 
|
| 58618 | 188  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
189  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
190  | 
theorem conjE [elim]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
191  | 
assumes "A \<and> B"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
192  | 
obtains A and B  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
193  | 
proof  | 
| 58618 | 194  | 
from \<open>A \<and> B\<close> show A by (rule conjD\<^sub>1)  | 
195  | 
from \<open>A \<and> B\<close> show B by (rule conjD\<^sub>2)  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
196  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
197  | 
|
| 58618 | 198  | 
text \<open>  | 
| 62279 | 199  | 
Here is an example of swapping conjuncts with a single intermediate  | 
200  | 
elimination step:  | 
|
| 58618 | 201  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
202  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
203  | 
(*<*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
204  | 
lemma "\<And>A. PROP A \<Longrightarrow> PROP A"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
205  | 
proof -  | 
| 62279 | 206  | 
fix A B  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
207  | 
(*>*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
208  | 
assume "A \<and> B"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
209  | 
then obtain B and A ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
210  | 
then have "B \<and> A" ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
211  | 
(*<*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
212  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
213  | 
(*>*)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
214  | 
|
| 58618 | 215  | 
text \<open>  | 
| 62279 | 216  | 
Note that the analogous elimination rule for disjunction ``\<^theory_text>\<open>assumes "A \<or> B"  | 
217  | 
  obtains A \<BBAR> B\<close>'' coincides with the original axiomatization of @{thm
 | 
|
218  | 
disjE}.  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
219  | 
|
| 61421 | 220  | 
\<^medskip>  | 
| 62279 | 221  | 
We continue propositional logic by introducing absurdity with its  | 
222  | 
characteristic elimination. Plain truth may then be defined as a proposition  | 
|
223  | 
that is trivially true.  | 
|
| 58618 | 224  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
225  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
226  | 
axiomatization false :: o (\<open>\<bottom>\<close>)  | 
| 62279 | 227  | 
where falseE [elim]: "\<bottom> \<Longrightarrow> A"  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
228  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
229  | 
definition true :: o (\<open>\<top>\<close>)  | 
| 62279 | 230  | 
where "\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>"  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
231  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
232  | 
theorem trueI [intro]: \<top>  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
233  | 
unfolding true_def ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
234  | 
|
| 58618 | 235  | 
text \<open>  | 
| 61421 | 236  | 
\<^medskip>  | 
237  | 
Now negation represents an implication towards absurdity:  | 
|
| 58618 | 238  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
239  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
240  | 
definition not :: "o \<Rightarrow> o" (\<open>\<not> _\<close> [40] 40)  | 
| 62279 | 241  | 
where "\<not> A \<equiv> A \<longrightarrow> \<bottom>"  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
242  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
243  | 
theorem notI [intro]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
244  | 
assumes "A \<Longrightarrow> \<bottom>"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
245  | 
shows "\<not> A"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
246  | 
unfolding not_def  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
247  | 
proof  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
248  | 
assume A  | 
| 58618 | 249  | 
then show \<bottom> by (rule \<open>A \<Longrightarrow> \<bottom>\<close>)  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
250  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
251  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
252  | 
theorem notE [elim]:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
253  | 
assumes "\<not> A" and A  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
254  | 
shows B  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
255  | 
proof -  | 
| 58618 | 256  | 
from \<open>\<not> A\<close> have "A \<longrightarrow> \<bottom>" unfolding not_def .  | 
257  | 
from \<open>A \<longrightarrow> \<bottom>\<close> and \<open>A\<close> have \<bottom> ..  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
258  | 
then show B ..  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
259  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
260  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
261  | 
|
| 58618 | 262  | 
subsection \<open>Classical logic\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
263  | 
|
| 58618 | 264  | 
text \<open>  | 
| 62279 | 265  | 
Subsequently we state the principle of classical contradiction as a local  | 
266  | 
assumption. Thus we refrain from forcing the object-logic into the classical  | 
|
267  | 
perspective. Within that context, we may derive well-known consequences of  | 
|
268  | 
the classical principle.  | 
|
| 58618 | 269  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
270  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
271  | 
locale classical =  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
272  | 
assumes classical: "(\<not> C \<Longrightarrow> C) \<Longrightarrow> C"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
273  | 
begin  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
274  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
275  | 
theorem double_negation:  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
276  | 
assumes "\<not> \<not> C"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
277  | 
shows C  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
278  | 
proof (rule classical)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
279  | 
assume "\<not> C"  | 
| 58618 | 280  | 
with \<open>\<not> \<not> C\<close> show C ..  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
281  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
282  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
283  | 
theorem tertium_non_datur: "C \<or> \<not> C"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
284  | 
proof (rule double_negation)  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
285  | 
show "\<not> \<not> (C \<or> \<not> C)"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
286  | 
proof  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
287  | 
assume "\<not> (C \<or> \<not> C)"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
288  | 
have "\<not> C"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
289  | 
proof  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
290  | 
assume C then have "C \<or> \<not> C" ..  | 
| 58618 | 291  | 
with \<open>\<not> (C \<or> \<not> C)\<close> show \<bottom> ..  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
292  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
293  | 
then have "C \<or> \<not> C" ..  | 
| 58618 | 294  | 
with \<open>\<not> (C \<or> \<not> C)\<close> show \<bottom> ..  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
295  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
296  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
297  | 
|
| 58618 | 298  | 
text \<open>  | 
| 62279 | 299  | 
These examples illustrate both classical reasoning and non-trivial  | 
300  | 
propositional proofs in general. All three rules characterize classical  | 
|
301  | 
logic independently, but the original rule is already the most convenient to  | 
|
| 69593 | 302  | 
use, because it leaves the conclusion unchanged. Note that \<^prop>\<open>(\<not> C \<Longrightarrow> C)  | 
303  | 
\<Longrightarrow> C\<close> fits again into our format for eliminations, despite the additional  | 
|
| 62279 | 304  | 
  twist that the context refers to the main conclusion. So we may write @{thm
 | 
305  | 
classical} as the Isar statement ``\<^theory_text>\<open>obtains \<not> thesis\<close>''. This also explains  | 
|
306  | 
nicely how classical reasoning really works: whatever the main \<open>thesis\<close>  | 
|
307  | 
might be, we may always assume its negation!  | 
|
| 58618 | 308  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
309  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
310  | 
end  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
311  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
312  | 
|
| 58618 | 313  | 
subsection \<open>Quantifiers \label{sec:framework-ex-quant}\<close>
 | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
314  | 
|
| 58618 | 315  | 
text \<open>  | 
| 62279 | 316  | 
Representing quantifiers is easy, thanks to the higher-order nature of the  | 
317  | 
underlying framework. According to the well-known technique introduced by  | 
|
| 76987 | 318  | 
Church \<^cite>\<open>"church40"\<close>, quantifiers are operators on predicates, which  | 
| 69593 | 319  | 
are syntactically represented as \<open>\<lambda>\<close>-terms of type \<^typ>\<open>i \<Rightarrow> o\<close>. Binder  | 
| 62279 | 320  | 
notation turns \<open>All (\<lambda>x. B x)\<close> into \<open>\<forall>x. B x\<close> etc.  | 
| 58618 | 321  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
322  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
323  | 
axiomatization All :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<forall>\<close> 10)  | 
| 62279 | 324  | 
where allI [intro]: "(\<And>x. B x) \<Longrightarrow> \<forall>x. B x"  | 
325  | 
and allD [dest]: "(\<forall>x. B x) \<Longrightarrow> B a"  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
326  | 
|
| 
80914
 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 
wenzelm 
parents: 
76987 
diff
changeset
 | 
327  | 
axiomatization Ex :: "(i \<Rightarrow> o) \<Rightarrow> o" (binder \<open>\<exists>\<close> 10)  | 
| 62279 | 328  | 
where exI [intro]: "B a \<Longrightarrow> (\<exists>x. B x)"  | 
329  | 
and exE [elim]: "(\<exists>x. B x) \<Longrightarrow> (\<And>x. B x \<Longrightarrow> C) \<Longrightarrow> C"  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
330  | 
|
| 58618 | 331  | 
text \<open>  | 
| 62279 | 332  | 
  The statement of @{thm exE} corresponds to ``\<^theory_text>\<open>assumes "\<exists>x. B x" obtains x
 | 
333  | 
where "B x"\<close>'' in Isar. In the subsequent example we illustrate quantifier  | 
|
334  | 
reasoning involving all four rules:  | 
|
| 58618 | 335  | 
\<close>  | 
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
336  | 
|
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
337  | 
theorem  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
338  | 
assumes "\<exists>x. \<forall>y. R x y"  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
339  | 
shows "\<forall>y. \<exists>x. R x y"  | 
| 61580 | 340  | 
proof \<comment> \<open>\<open>\<forall>\<close> introduction\<close>  | 
341  | 
obtain x where "\<forall>y. R x y" using \<open>\<exists>x. \<forall>y. R x y\<close> .. \<comment> \<open>\<open>\<exists>\<close> elimination\<close>  | 
|
342  | 
fix y have "R x y" using \<open>\<forall>y. R x y\<close> .. \<comment> \<open>\<open>\<forall>\<close> destruction\<close>  | 
|
343  | 
then show "\<exists>x. R x y" .. \<comment> \<open>\<open>\<exists>\<close> introduction\<close>  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
344  | 
qed  | 
| 
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
345  | 
|
| 29734 | 346  | 
|
| 58618 | 347  | 
subsection \<open>Canonical reasoning patterns\<close>  | 
| 29734 | 348  | 
|
| 58618 | 349  | 
text \<open>  | 
| 29734 | 350  | 
The main rules of first-order predicate logic from  | 
| 62279 | 351  | 
  \secref{sec:framework-ex-prop} and \secref{sec:framework-ex-quant} can now
 | 
352  | 
be summarized as follows, using the native Isar statement format of  | 
|
353  | 
  \secref{sec:framework-stmt}.
 | 
|
| 29734 | 354  | 
|
| 61421 | 355  | 
\<^medskip>  | 
| 29734 | 356  | 
  \begin{tabular}{l}
 | 
| 62279 | 357  | 
\<^theory_text>\<open>impI: assumes "A \<Longrightarrow> B" shows "A \<longrightarrow> B"\<close> \\  | 
358  | 
\<^theory_text>\<open>impD: assumes "A \<longrightarrow> B" and A shows B\<close> \\[1ex]  | 
|
| 29734 | 359  | 
|
| 62279 | 360  | 
\<^theory_text>\<open>disjI\<^sub>1: assumes A shows "A \<or> B"\<close> \\  | 
361  | 
\<^theory_text>\<open>disjI\<^sub>2: assumes B shows "A \<or> B"\<close> \\  | 
|
362  | 
\<^theory_text>\<open>disjE: assumes "A \<or> B" obtains A \<BBAR> B\<close> \\[1ex]  | 
|
| 29734 | 363  | 
|
| 62279 | 364  | 
\<^theory_text>\<open>conjI: assumes A and B shows A \<and> B\<close> \\  | 
365  | 
\<^theory_text>\<open>conjE: assumes "A \<and> B" obtains A and B\<close> \\[1ex]  | 
|
| 29734 | 366  | 
|
| 62279 | 367  | 
\<^theory_text>\<open>falseE: assumes \<bottom> shows A\<close> \\  | 
368  | 
\<^theory_text>\<open>trueI: shows \<top>\<close> \\[1ex]  | 
|
| 29734 | 369  | 
|
| 62279 | 370  | 
\<^theory_text>\<open>notI: assumes "A \<Longrightarrow> \<bottom>" shows "\<not> A"\<close> \\  | 
371  | 
\<^theory_text>\<open>notE: assumes "\<not> A" and A shows B\<close> \\[1ex]  | 
|
| 29734 | 372  | 
|
| 62279 | 373  | 
\<^theory_text>\<open>allI: assumes "\<And>x. B x" shows "\<forall>x. B x"\<close> \\  | 
374  | 
\<^theory_text>\<open>allE: assumes "\<forall>x. B x" shows "B a"\<close> \\[1ex]  | 
|
| 29734 | 375  | 
|
| 62279 | 376  | 
\<^theory_text>\<open>exI: assumes "B a" shows "\<exists>x. B x"\<close> \\  | 
377  | 
\<^theory_text>\<open>exE: assumes "\<exists>x. B x" obtains a where "B a"\<close>  | 
|
| 29734 | 378  | 
  \end{tabular}
 | 
| 61421 | 379  | 
\<^medskip>  | 
| 29734 | 380  | 
|
| 62279 | 381  | 
This essentially provides a declarative reading of Pure rules as Isar  | 
382  | 
reasoning patterns: the rule statements tells how a canonical proof outline  | 
|
383  | 
shall look like. Since the above rules have already been declared as  | 
|
384  | 
  @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure)
 | 
|
385  | 
dest} --- each according to its particular shape --- we can immediately  | 
|
386  | 
write Isar proof texts as follows:  | 
|
| 58618 | 387  | 
\<close>  | 
| 29734 | 388  | 
|
389  | 
(*<*)  | 
|
390  | 
theorem "\<And>A. PROP A \<Longrightarrow> PROP A"  | 
|
391  | 
proof -  | 
|
392  | 
(*>*)  | 
|
393  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
394  | 
  text_raw \<open>\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 395  | 
|
396  | 
have "A \<longrightarrow> B"  | 
|
397  | 
proof  | 
|
398  | 
assume A  | 
|
| 62271 | 399  | 
show B \<proof>  | 
| 29734 | 400  | 
qed  | 
401  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
402  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 403  | 
|
| 62271 | 404  | 
have "A \<longrightarrow> B" and A \<proof>  | 
| 29734 | 405  | 
then have B ..  | 
406  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
407  | 
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 408  | 
|
| 62271 | 409  | 
have A \<proof>  | 
| 29734 | 410  | 
then have "A \<or> B" ..  | 
411  | 
||
| 62271 | 412  | 
have B \<proof>  | 
| 29734 | 413  | 
then have "A \<or> B" ..  | 
414  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
415  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 416  | 
|
| 62271 | 417  | 
have "A \<or> B" \<proof>  | 
| 29734 | 418  | 
then have C  | 
419  | 
proof  | 
|
420  | 
assume A  | 
|
| 62271 | 421  | 
then show C \<proof>  | 
| 29734 | 422  | 
next  | 
423  | 
assume B  | 
|
| 62271 | 424  | 
then show C \<proof>  | 
| 29734 | 425  | 
qed  | 
426  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
427  | 
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 428  | 
|
| 62271 | 429  | 
have A and B \<proof>  | 
| 29734 | 430  | 
then have "A \<and> B" ..  | 
431  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
432  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 433  | 
|
| 62271 | 434  | 
have "A \<and> B" \<proof>  | 
| 29734 | 435  | 
then obtain A and B ..  | 
436  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
437  | 
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 438  | 
|
| 62271 | 439  | 
have "\<bottom>" \<proof>  | 
| 29734 | 440  | 
then have A ..  | 
441  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
442  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 443  | 
|
444  | 
have "\<top>" ..  | 
|
445  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
446  | 
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 447  | 
|
448  | 
have "\<not> A"  | 
|
449  | 
proof  | 
|
450  | 
assume A  | 
|
| 62271 | 451  | 
then show "\<bottom>" \<proof>  | 
| 29734 | 452  | 
qed  | 
453  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
454  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 455  | 
|
| 62271 | 456  | 
have "\<not> A" and A \<proof>  | 
| 29734 | 457  | 
then have B ..  | 
458  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
459  | 
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 460  | 
|
461  | 
have "\<forall>x. B x"  | 
|
462  | 
proof  | 
|
463  | 
fix x  | 
|
| 62271 | 464  | 
show "B x" \<proof>  | 
| 29734 | 465  | 
qed  | 
466  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
467  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 468  | 
|
| 62271 | 469  | 
have "\<forall>x. B x" \<proof>  | 
| 29734 | 470  | 
then have "B a" ..  | 
471  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
472  | 
  text_raw \<open>\end{minipage}\\[3ex]\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 473  | 
|
474  | 
have "\<exists>x. B x"  | 
|
475  | 
proof  | 
|
| 62271 | 476  | 
show "B a" \<proof>  | 
| 29734 | 477  | 
qed  | 
478  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
479  | 
  text_raw \<open>\end{minipage}\qquad\begin{minipage}[t]{0.4\textwidth}\<close>(*<*)next(*>*)
 | 
| 29734 | 480  | 
|
| 62271 | 481  | 
have "\<exists>x. B x" \<proof>  | 
| 29734 | 482  | 
then obtain a where "B a" ..  | 
483  | 
||
| 
58999
 
ed09ae4ea2d8
uniform treatment of all document markup commands: 'text' and 'txt' merely differ in LaTeX style;
 
wenzelm 
parents: 
58889 
diff
changeset
 | 
484  | 
  text_raw \<open>\end{minipage}\<close>
 | 
| 29734 | 485  | 
|
486  | 
(*<*)  | 
|
487  | 
qed  | 
|
488  | 
(*>*)  | 
|
489  | 
||
| 58618 | 490  | 
text \<open>  | 
| 61421 | 491  | 
\<^bigskip>  | 
| 62279 | 492  | 
Of course, these proofs are merely examples. As sketched in  | 
493  | 
  \secref{sec:framework-subproof}, there is a fair amount of flexibility in
 | 
|
494  | 
expressing Pure deductions in Isar. Here the user is asked to express  | 
|
495  | 
himself adequately, aiming at proof texts of literary quality.  | 
|
| 58618 | 496  | 
\<close>  | 
| 29734 | 497  | 
|
| 
29730
 
924c1fd5f303
added example "First-Order Logic" -- mostly from Trybulec Festschrift;
 
wenzelm 
parents:  
diff
changeset
 | 
498  | 
end %visible  |