2014-08-02 23:59:01 +00:00
|
|
|
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
2014-06-30 18:44:47 +00:00
|
|
|
-- Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
-- Author: Leonardo de Moura
|
2014-08-20 22:49:44 +00:00
|
|
|
|
2014-10-05 17:50:13 +00:00
|
|
|
import logic.eq logic.inhabited logic.decidable
|
2014-10-02 00:51:17 +00:00
|
|
|
open eq.ops decidable
|
2014-08-01 01:40:09 +00:00
|
|
|
|
2014-06-30 02:30:38 +00:00
|
|
|
inductive option (A : Type) : Type :=
|
2014-09-05 05:31:52 +00:00
|
|
|
none {} : option A,
|
|
|
|
some : A → option A
|
2014-06-30 02:30:38 +00:00
|
|
|
|
2014-09-04 23:36:06 +00:00
|
|
|
namespace option
|
2014-09-05 05:31:52 +00:00
|
|
|
definition is_none {A : Type} (o : option A) : Prop :=
|
|
|
|
rec true (λ a, false) o
|
|
|
|
|
|
|
|
theorem is_none_none {A : Type} : is_none (@none A) :=
|
|
|
|
trivial
|
|
|
|
|
|
|
|
theorem not_is_none_some {A : Type} (a : A) : ¬ is_none (some a) :=
|
|
|
|
not_false_trivial
|
|
|
|
|
|
|
|
theorem none_ne_some {A : Type} (a : A) : none ≠ some a :=
|
|
|
|
assume H : none = some a, absurd
|
|
|
|
(H ▸ is_none_none)
|
|
|
|
(not_is_none_some a)
|
|
|
|
|
2014-09-19 22:04:52 +00:00
|
|
|
protected theorem equal {A : Type} {a₁ a₂ : A} (H : some a₁ = some a₂) : a₁ = a₂ :=
|
2014-09-05 05:31:52 +00:00
|
|
|
congr_arg (option.rec a₁ (λ a, a)) H
|
|
|
|
|
2014-10-02 16:00:34 +00:00
|
|
|
protected definition is_inhabited [instance] (A : Type) : inhabited (option A) :=
|
2014-09-05 05:31:52 +00:00
|
|
|
inhabited.mk none
|
|
|
|
|
2014-09-30 16:02:37 +00:00
|
|
|
protected definition has_decidable_eq [instance] {A : Type} (H : decidable_eq A) : decidable_eq (option A) :=
|
2014-09-09 23:07:07 +00:00
|
|
|
take o₁ o₂ : option A,
|
2014-09-08 05:22:04 +00:00
|
|
|
rec_on o₁
|
|
|
|
(rec_on o₂ (inl rfl) (take a₂, (inr (none_ne_some a₂))))
|
|
|
|
(take a₁ : A, rec_on o₂
|
|
|
|
(inr (ne.symm (none_ne_some a₁)))
|
|
|
|
(take a₂ : A, decidable.rec_on (H a₁ a₂)
|
|
|
|
(assume Heq : a₁ = a₂, inl (Heq ▸ rfl))
|
2014-09-09 23:07:07 +00:00
|
|
|
(assume Hne : a₁ ≠ a₂, inr (assume Hn : some a₁ = some a₂, absurd (equal Hn) Hne))))
|
2014-08-20 02:32:44 +00:00
|
|
|
end option
|