2014-08-30 03:54:28 +00:00
|
|
|
|
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
|
-- Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
|
-- Author: Jeremy Avigad
|
|
|
|
|
|
2014-10-05 18:11:48 +00:00
|
|
|
|
-- logic.instances
|
2014-08-30 03:54:28 +00:00
|
|
|
|
-- ====================
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-10-05 17:50:13 +00:00
|
|
|
|
import logic.connectives algebra.relation
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
namespace relation
|
|
|
|
|
|
2014-09-03 23:00:38 +00:00
|
|
|
|
open relation
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
|
|
|
|
-- Congruences for logic
|
|
|
|
|
-- ---------------------
|
|
|
|
|
|
2014-08-30 03:54:28 +00:00
|
|
|
|
theorem congruence_not : congruence iff iff not :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
congruence.mk
|
2014-08-05 00:07:59 +00:00
|
|
|
|
(take a b,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
assume H : a ↔ b, iff.intro
|
|
|
|
|
(assume H1 : ¬a, assume H2 : b, H1 (iff.elim_right H H2))
|
|
|
|
|
(assume H1 : ¬b, assume H2 : a, H1 (iff.elim_left H H2)))
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-30 03:54:28 +00:00
|
|
|
|
theorem congruence_and : congruence2 iff iff iff and :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
congruence2.mk
|
2014-08-05 00:07:59 +00:00
|
|
|
|
(take a1 b1 a2 b2,
|
|
|
|
|
assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.intro
|
|
|
|
|
(assume H3 : a1 ∧ a2, and.imp_and H3 (iff.elim_left H1) (iff.elim_left H2))
|
|
|
|
|
(assume H3 : b1 ∧ b2, and.imp_and H3 (iff.elim_right H1) (iff.elim_right H2)))
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-30 03:54:28 +00:00
|
|
|
|
theorem congruence_or : congruence2 iff iff iff or :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
congruence2.mk
|
2014-08-05 00:07:59 +00:00
|
|
|
|
(take a1 b1 a2 b2,
|
|
|
|
|
assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.intro
|
|
|
|
|
(assume H3 : a1 ∨ a2, or.imp_or H3 (iff.elim_left H1) (iff.elim_left H2))
|
|
|
|
|
(assume H3 : b1 ∨ b2, or.imp_or H3 (iff.elim_right H1) (iff.elim_right H2)))
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-30 03:54:28 +00:00
|
|
|
|
theorem congruence_imp : congruence2 iff iff iff imp :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
congruence2.mk
|
2014-08-05 00:07:59 +00:00
|
|
|
|
(take a1 b1 a2 b2,
|
|
|
|
|
assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.intro
|
|
|
|
|
(assume H3 : a1 → a2, assume Hb1 : b1, iff.elim_left H2 (H3 ((iff.elim_right H1) Hb1)))
|
|
|
|
|
(assume H3 : b1 → b2, assume Ha1 : a1, iff.elim_right H2 (H3 ((iff.elim_left H1) Ha1))))
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-30 03:54:28 +00:00
|
|
|
|
theorem congruence_iff : congruence2 iff iff iff iff :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
congruence2.mk
|
2014-08-05 00:07:59 +00:00
|
|
|
|
(take a1 b1 a2 b2,
|
|
|
|
|
assume H1 : a1 ↔ b1, assume H2 : a2 ↔ b2,
|
2014-09-05 04:25:21 +00:00
|
|
|
|
iff.intro
|
|
|
|
|
(assume H3 : a1 ↔ a2, iff.trans (iff.symm H1) (iff.trans H3 H2))
|
|
|
|
|
(assume H3 : b1 ↔ b2, iff.trans H1 (iff.trans H3 (iff.symm H2))))
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-09-05 04:25:21 +00:00
|
|
|
|
-- theorem congruence_const_iff [instance] := congruence.const iff iff.refl
|
2014-08-30 03:54:28 +00:00
|
|
|
|
definition congruence_not_compose [instance] := congruence.compose congruence_not
|
|
|
|
|
definition congruence_and_compose [instance] := congruence.compose21 congruence_and
|
|
|
|
|
definition congruence_or_compose [instance] := congruence.compose21 congruence_or
|
|
|
|
|
definition congruence_implies_compose [instance] := congruence.compose21 congruence_imp
|
|
|
|
|
definition congruence_iff_compose [instance] := congruence.compose21 congruence_iff
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
|
|
|
|
-- Generalized substitution
|
|
|
|
|
-- ------------------------
|
|
|
|
|
|
|
|
|
|
-- TODO: note that the target has to be "iff". Otherwise, there is not enough
|
|
|
|
|
-- information to infer an mp-like relation.
|
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
namespace general_operations
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-10-12 20:06:00 +00:00
|
|
|
|
theorem subst {T : Type} (R : T → T → Prop) ⦃P : T → Prop⦄ [C : congruence R iff P]
|
2014-09-05 04:25:21 +00:00
|
|
|
|
{a b : T} (H : R a b) (H1 : P a) : P b := iff.elim_left (congruence.app C H) H1
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
end general_operations
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
|
|
|
|
-- = is an equivalence relation
|
|
|
|
|
-- ----------------------------
|
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
theorem is_reflexive_eq [instance] (T : Type) : relation.is_reflexive (@eq T) :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
relation.is_reflexive.mk (@eq.refl T)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
|
|
|
|
theorem is_symmetric_eq [instance] (T : Type) : relation.is_symmetric (@eq T) :=
|
2014-09-05 01:41:06 +00:00
|
|
|
|
relation.is_symmetric.mk (@eq.symm T)
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
theorem is_transitive_eq [instance] (T : Type) : relation.is_transitive (@eq T) :=
|
2014-09-05 01:41:06 +00:00
|
|
|
|
relation.is_transitive.mk (@eq.trans T)
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
-- TODO: this is only temporary, needed to inform Lean that is_equivalence is a class
|
|
|
|
|
theorem is_equivalence_eq [instance] (T : Type) : relation.is_equivalence (@eq T) :=
|
2014-09-04 23:36:06 +00:00
|
|
|
|
relation.is_equivalence.mk _ _ _
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- iff is an equivalence relation
|
|
|
|
|
-- ------------------------------
|
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
theorem is_reflexive_iff [instance] : relation.is_reflexive iff :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
relation.is_reflexive.mk (@iff.refl)
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
theorem is_symmetric_iff [instance] : relation.is_symmetric iff :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
relation.is_symmetric.mk (@iff.symm)
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
theorem is_transitive_iff [instance] : relation.is_transitive iff :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
relation.is_transitive.mk (@iff.trans)
|
2014-08-05 00:07:59 +00:00
|
|
|
|
|
|
|
|
|
|
|
|
|
|
-- Mp-like for iff
|
|
|
|
|
-- ---------------
|
|
|
|
|
|
2014-08-26 04:26:17 +00:00
|
|
|
|
theorem mp_like_iff [instance] (a b : Prop) (H : a ↔ b) : @relation.mp_like iff a b H :=
|
2014-09-05 04:25:21 +00:00
|
|
|
|
relation.mp_like.mk (iff.elim_left H)
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
2014-08-20 22:49:44 +00:00
|
|
|
|
|
2014-08-20 02:32:44 +00:00
|
|
|
|
-- Substition for iff
|
|
|
|
|
-- ------------------
|
2014-09-05 04:25:21 +00:00
|
|
|
|
namespace iff
|
2014-10-12 20:06:00 +00:00
|
|
|
|
theorem subst {P : Prop → Prop} [C : congruence iff iff P] {a b : Prop} (H : a ↔ b) (H1 : P a) :
|
2014-08-20 02:32:44 +00:00
|
|
|
|
P b :=
|
|
|
|
|
@general_operations.subst Prop iff P C a b H H1
|
2014-09-05 04:25:21 +00:00
|
|
|
|
end iff
|
2014-08-20 22:49:44 +00:00
|
|
|
|
|
|
|
|
|
-- Support for calculations with iff
|
2014-08-20 02:32:44 +00:00
|
|
|
|
-- ----------------
|
|
|
|
|
|
2014-09-05 04:25:21 +00:00
|
|
|
|
calc_subst iff.subst
|
2014-08-20 02:32:44 +00:00
|
|
|
|
|
|
|
|
|
namespace iff_ops
|
2014-09-05 04:25:21 +00:00
|
|
|
|
postfix `⁻¹`:100 := iff.symm
|
|
|
|
|
infixr `⬝`:75 := iff.trans
|
|
|
|
|
infixr `▸`:75 := iff.subst
|
2014-09-17 21:39:05 +00:00
|
|
|
|
definition refl := iff.refl
|
|
|
|
|
definition symm := @iff.symm
|
|
|
|
|
definition trans := @iff.trans
|
|
|
|
|
definition subst := @iff.subst
|
|
|
|
|
definition mp := @iff.mp
|
2014-08-20 02:32:44 +00:00
|
|
|
|
end iff_ops
|
|
|
|
|
end relation
|