Isabelle/ZF builds on the first order logic. Almost everything one would like to have in this area is covered in the standard Isabelle libraries. The material in this theory provides some lemmas that are missing or allow for a more readable proof style.
This section contains mostly shortcuts and workarounds that allow to use more readable coding style.
The next lemma serves as a workaround to problems with applying the definition of transitivity (of a relation) in our coding style (any attempt to do something like using trans_def puts Isabelle in an infinite loop).
lemma Fol1_L2:
assumes A1: \( \forall x y z.\ \langle x, y\rangle \in r \wedge \langle y, z\rangle \in r \longrightarrow \langle x, z\rangle \in r \)
shows \( \text{trans}(r) \)proofAnother workaround for the problem of Isabelle simplifier looping when the transitivity definition is used.
lemma Fol1_L3:
assumes A1: \( \text{trans}(r) \) and A2: \( \langle a,b\rangle \in r \wedge \langle b,c\rangle \in r \)
shows \( \langle a,c\rangle \in r \)proofThere is a problem with application of the definition of asymetry for relations. The next lemma is a workaround.
lemma Fol1_L4:
assumes A1: \( \text{antisym}(r) \) and A2: \( \langle a,b\rangle \in r \), \( \langle b,a\rangle \in r \)
shows \( a=b \)proofThe definition below implements a common idiom that states that (perhaps under some assumptions) exactly one of given three statements is true.
definition
\( \text{Exactly_1_of_3_holds} (p,q,r) \equiv \) \( (p\vee q\vee r) \wedge (p \longrightarrow \neg q \wedge \neg r) \wedge (q \longrightarrow \neg p \wedge \neg r) \wedge (r \longrightarrow \neg p \wedge \neg q) \)
The next lemma allows to prove statements of the form \( \text{Exactly_1_of_3_holds} (p,q,r) \).
lemma Fol1_L5:
assumes \( p\vee q\vee r \) and \( p \longrightarrow \neg q \wedge \neg r \) and \( q \longrightarrow \neg p \wedge \neg r \) and \( r \longrightarrow \neg p \wedge \neg q \)
shows \( \text{Exactly_1_of_3_holds} (p,q,r) \)proofIf exactly one of \(p,q,r\) holds and \(p\) is not true, then \(q\) or \(r\).
lemma Fol1_L6:
assumes A1: \( \neg p \) and A2: \( \text{Exactly_1_of_3_holds} (p,q,r) \)
shows \( q\vee r \)proofIf exactly one of \(p,q,r\) holds and \(q\) is true, then \(r\) can not be true.
lemma Fol1_L7:
assumes A1: \( q \) and A2: \( \text{Exactly_1_of_3_holds} (p,q,r) \)
shows \( \neg r \)proofThe next lemma demonstrates an elegant form of the \( \text{Exactly_1_of_3_holds} (p,q,r) \) predicate.
lemma Fol1_L8:
shows \( \text{Exactly_1_of_3_holds} (p,q,r) \longleftrightarrow (p\longleftrightarrow q\longleftrightarrow r) \wedge \neg (p\wedge q\wedge r) \)proofA property of the Exactly_1_of_3_holds predicate.
lemma Fol1_L8A:
assumes A1: \( \text{Exactly_1_of_3_holds} (p,q,r) \)
shows \( p \longleftrightarrow \neg (q \vee r) \)proofExclusive or definition. There is one also defined in the standard Isabelle, denoted xor, but it relates to boolean values, which are sets. Here we define a logical functor.
definition
\( p \text{ Xor } q \equiv (p\vee q) \wedge \neg (p \wedge q) \)
The "exclusive or" is the same as negation of equivalence.
lemma Fol1_L9:
shows \( p \text{ Xor } q \longleftrightarrow \neg (p\longleftrightarrow q) \) using Xor_defConstructions from the same sets are the same. It is suprising but we do have to use this as a rule in rare cases.
lemma same_constr:
assumes \( x=y \)
shows \( P(x) = P(y) \) using assmsEquivalence relations are symmetric.
lemma equiv_is_sym:
assumes A1: \( \text{equiv}(X,r) \) and A2: \( \langle x,y\rangle \in r \)
shows \( \langle y,x\rangle \in r \)proof