2014-12-01 04:34:12 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
2014-12-15 21:43:42 +00:00
|
|
|
|
|
|
|
Module init.relation
|
2014-12-01 04:34:12 +00:00
|
|
|
Authors: Leonardo de Moura
|
|
|
|
-/
|
|
|
|
prelude
|
|
|
|
import init.logic
|
2014-11-19 01:32:22 +00:00
|
|
|
|
|
|
|
-- TODO(Leo): remove duplication between this file and algebra/relation.lean
|
|
|
|
-- We need some of the following definitions asap when "initializing" Lean.
|
|
|
|
|
|
|
|
variables {A B : Type} (R : B → B → Prop)
|
2015-01-26 19:49:08 +00:00
|
|
|
local infix `≺`:50 := R
|
2014-11-19 01:32:22 +00:00
|
|
|
|
|
|
|
definition reflexive := ∀x, x ≺ x
|
|
|
|
|
|
|
|
definition symmetric := ∀⦃x y⦄, x ≺ y → y ≺ x
|
|
|
|
|
|
|
|
definition transitive := ∀⦃x y z⦄, x ≺ y → y ≺ z → x ≺ z
|
|
|
|
|
|
|
|
definition irreflexive := ∀x, ¬ x ≺ x
|
|
|
|
|
|
|
|
definition anti_symmetric := ∀⦃x y⦄, x ≺ y → y ≺ x → x = y
|
|
|
|
|
|
|
|
definition empty_relation := λa₁ a₂ : A, false
|
|
|
|
|
|
|
|
definition subrelation (Q R : B → B → Prop) := ∀⦃x y⦄, Q x y → R x y
|
|
|
|
|
|
|
|
definition inv_image (f : A → B) : A → A → Prop :=
|
|
|
|
λa₁ a₂, f a₁ ≺ f a₂
|
|
|
|
|
|
|
|
theorem inv_image.trans (f : A → B) (H : transitive R) : transitive (inv_image R f) :=
|
|
|
|
λ (a₁ a₂ a₃ : A) (H₁ : inv_image R f a₁ a₂) (H₂ : inv_image R f a₂ a₃), H H₁ H₂
|
|
|
|
|
|
|
|
theorem inv_image.irreflexive (f : A → B) (H : irreflexive R) : irreflexive (inv_image R f) :=
|
|
|
|
λ (a : A) (H₁ : inv_image R f a a), H (f a) H₁
|
|
|
|
|
|
|
|
inductive tc {A : Type} (R : A → A → Prop) : A → A → Prop :=
|
2015-02-26 01:00:10 +00:00
|
|
|
| base : ∀a b, R a b → tc R a b
|
|
|
|
| trans : ∀a b c, tc R a b → tc R b c → tc R a c
|