2014-12-15 21:43:42 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
Author: Leonardo de Moura, Jeremy Avigad, Floris van Doorn
|
|
|
|
-/
|
2014-12-01 04:34:12 +00:00
|
|
|
prelude
|
2015-02-01 19:14:01 +00:00
|
|
|
import init.datatypes init.num init.wf init.logic init.tactic
|
2014-12-01 04:34:12 +00:00
|
|
|
|
2015-02-01 16:39:47 +00:00
|
|
|
definition dpair := @sigma.mk
|
2015-09-30 15:06:31 +00:00
|
|
|
notation `Σ` binders `, ` r:(scoped P, sigma P) := r
|
2015-02-01 19:14:01 +00:00
|
|
|
-- notation for n-ary tuples; input ⟨ ⟩ as \< \>
|
2015-09-30 15:06:31 +00:00
|
|
|
notation `⟨`:max t:(foldr `, ` (e r, sigma.mk e r)) `⟩`:0 := t
|
2014-11-17 04:08:52 +00:00
|
|
|
|
2015-07-15 19:49:47 +00:00
|
|
|
lemma ex_of_sig {A : Type} {P : A → Prop} : (Σ x, P x) → ∃ x, P x :=
|
|
|
|
assume h, obtain x hx, from h, exists.intro x hx
|
|
|
|
|
2014-11-17 04:08:52 +00:00
|
|
|
namespace sigma
|
2014-12-20 02:23:08 +00:00
|
|
|
notation `pr₁` := pr1
|
|
|
|
notation `pr₂` := pr2
|
2014-12-01 04:34:12 +00:00
|
|
|
|
|
|
|
namespace ops
|
2014-12-20 02:23:08 +00:00
|
|
|
postfix `.1`:(max+1) := pr1
|
|
|
|
postfix `.2`:(max+1) := pr2
|
2014-12-01 04:34:12 +00:00
|
|
|
end ops
|
|
|
|
|
2014-12-11 23:48:09 +00:00
|
|
|
open ops well_founded
|
2014-12-01 04:34:12 +00:00
|
|
|
|
2014-11-17 04:08:52 +00:00
|
|
|
section
|
|
|
|
variables {A : Type} {B : A → Type}
|
|
|
|
variable (Ra : A → A → Prop)
|
|
|
|
variable (Rb : ∀ a, B a → B a → Prop)
|
|
|
|
|
2015-04-01 02:56:05 +00:00
|
|
|
theorem dpair_eq : ∀ {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} (H₁ : a₁ = a₂), eq.rec_on H₁ b₁ = b₂ → ⟨a₁, b₁⟩ = ⟨a₂, b₂⟩
|
2015-06-16 19:52:08 +00:00
|
|
|
| a₁ a₁ b₁ b₁ (eq.refl a₁) (eq.refl b₁) := rfl
|
2015-04-01 02:56:05 +00:00
|
|
|
|
2015-05-25 12:13:23 +00:00
|
|
|
protected theorem eq {p₁ p₂ : Σa : A, B a} :
|
2015-04-01 02:56:05 +00:00
|
|
|
∀(H₁ : p₁.1 = p₂.1) (H₂ : eq.rec_on H₁ p₁.2 = p₂.2), p₁ = p₂ :=
|
|
|
|
destruct p₁ (take a₁ b₁, destruct p₂ (take a₂ b₂ H₁ H₂, dpair_eq H₁ H₂))
|
|
|
|
|
2014-11-17 04:08:52 +00:00
|
|
|
-- Lexicographical order based on Ra and Rb
|
|
|
|
inductive lex : sigma B → sigma B → Prop :=
|
2015-02-26 01:00:10 +00:00
|
|
|
| left : ∀{a₁ b₁} a₂ b₂, Ra a₁ a₂ → lex ⟨a₁, b₁⟩ ⟨a₂, b₂⟩
|
|
|
|
| right : ∀a {b₁ b₂}, Rb a b₁ b₂ → lex ⟨a, b₁⟩ ⟨a, b₂⟩
|
2014-11-17 04:08:52 +00:00
|
|
|
end
|
|
|
|
|
2015-04-22 02:13:19 +00:00
|
|
|
section
|
2014-11-17 04:08:52 +00:00
|
|
|
parameters {A : Type} {B : A → Type}
|
|
|
|
parameters {Ra : A → A → Prop} {Rb : Π a : A, B a → B a → Prop}
|
2015-04-22 02:13:19 +00:00
|
|
|
local infix `≺`:50 := lex Ra Rb
|
2014-11-17 04:08:52 +00:00
|
|
|
|
2014-12-11 23:48:09 +00:00
|
|
|
definition lex.accessible {a} (aca : acc Ra a) (acb : ∀a, well_founded (Rb a)) : ∀ (b : B a), acc (lex Ra Rb) ⟨a, b⟩ :=
|
2014-11-17 04:08:52 +00:00
|
|
|
acc.rec_on aca
|
2014-12-11 23:48:09 +00:00
|
|
|
(λxa aca (iHa : ∀y, Ra y xa → ∀b : B y, acc (lex Ra Rb) ⟨y, b⟩),
|
2014-11-17 04:08:52 +00:00
|
|
|
λb : B xa, acc.rec_on (acb xa b)
|
|
|
|
(λxb acb
|
2014-12-12 20:36:51 +00:00
|
|
|
(iHb : ∀ (y : B xa), Rb xa y xb → acc (lex Ra Rb) ⟨xa, y⟩),
|
2014-12-11 23:48:09 +00:00
|
|
|
acc.intro ⟨xa, xb⟩ (λp (lt : p ≺ ⟨xa, xb⟩),
|
2014-11-17 04:08:52 +00:00
|
|
|
have aux : xa = xa → xb == xb → acc (lex Ra Rb) p, from
|
2015-02-11 20:49:27 +00:00
|
|
|
@sigma.lex.rec_on A B Ra Rb (λp₁ p₂, p₂.1 = xa → p₂.2 == xb → acc (lex Ra Rb) p₁)
|
|
|
|
p ⟨xa, xb⟩ lt
|
2014-12-12 20:36:51 +00:00
|
|
|
(λ (a₁ : A) (b₁ : B a₁) (a₂ : A) (b₂ : B a₂) (H : Ra a₁ a₂) (eq₂ : a₂ = xa) (eq₃ : b₂ == xb),
|
|
|
|
begin cases eq₂, exact (iHa a₁ H b₁) end)
|
|
|
|
(λ (a : A) (b₁ b₂ : B a) (H : Rb a b₁ b₂) (eq₂ : a = xa) (eq₃ : b₂ == xb),
|
|
|
|
begin cases eq₂, cases eq₃, exact (iHb b₁ H) end),
|
2014-11-17 04:08:52 +00:00
|
|
|
aux rfl !heq.refl)))
|
|
|
|
|
|
|
|
-- The lexicographical order of well founded relations is well-founded
|
|
|
|
definition lex.wf (Ha : well_founded Ra) (Hb : ∀ x, well_founded (Rb x)) : well_founded (lex Ra Rb) :=
|
|
|
|
well_founded.intro (λp, destruct p (λa b, lex.accessible (Ha a) Hb b))
|
|
|
|
end
|
|
|
|
end sigma
|