lean2/library/standard/logic/connectives/if.lean
Leonardo de Moura c3f57cdb1c feat(library/standard/logic/classes): add 'by_contradiction' theorem for decidable propositions
Signed-off-by: Leonardo de Moura <leonardo@microsoft.com>
2014-08-03 22:58:12 -07:00

73 lines
3.5 KiB
Text

----------------------------------------------------------------------------------------------------
-- Copyright (c) 2014 Microsoft Corporation. All rights reserved.
-- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Leonardo de Moura
----------------------------------------------------------------------------------------------------
import logic.classes.decidable tools.tactic
using decidable tactic eq_proofs
definition ite (c : Prop) {H : decidable c} {A : Type} (t e : A) : A :=
rec_on H (assume Hc, t) (assume Hnc, e)
notation `if` c `then` t `else` e:45 := ite c t e
theorem if_pos {c : Prop} {H : decidable c} (Hc : c) {A : Type} (t e : A) : (if c then t else e) = t :=
decidable_rec
(assume Hc : c, refl (@ite c (inl Hc) A t e))
(assume Hnc : ¬c, absurd_elim (@ite c (inr Hnc) A t e = t) Hc Hnc)
H
theorem if_neg {c : Prop} {H : decidable c} (Hnc : ¬c) {A : Type} (t e : A) : (if c then t else e) = e :=
decidable_rec
(assume Hc : c, absurd_elim (@ite c (inl Hc) A t e = e) Hc Hnc)
(assume Hnc : ¬c, refl (@ite c (inr Hnc) A t e))
H
theorem if_t_t (c : Prop) {H : decidable c} {A : Type} (t : A) : (if c then t else t) = t :=
decidable_rec
(assume Hc : c, refl (@ite c (inl Hc) A t t))
(assume Hnc : ¬c, refl (@ite c (inr Hnc) A t t))
H
theorem if_true {A : Type} (t e : A) : (if true then t else e) = t :=
if_pos trivial t e
theorem if_false {A : Type} (t e : A) : (if false then t else e) = e :=
if_neg not_false_trivial t e
theorem if_congr_aux {c₁ c₂ : Prop} {H₁ : decidable c₁} {H₂ : decidable c₂} {A : Type} {t₁ t₂ e₁ e₂ : A}
(Hc : c₁ = c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) :
(if c₁ then t₁ else e₁) = (if c₂ then t₂ else e₂) :=
have H1 : ∀ (H : decidable c₁), (@ite c₁ H₁ A t₁ e₁) = (@ite c₁ H A t₁ e₁), from
take H : decidable c₁,
have Hd : H₁ = H, from irrelevant H₁ H,
Hd ▸ refl (@ite c₁ H₁ A t₁ e₁),
have H2 : (@ite c₁ H₁ A t₁ e₁) = (@ite c₂ H₂ A t₁ e₁), from
(Hc ▸ H1) H₂,
Ht ▸ He ▸ H2
theorem if_congr {c₁ c₂ : Prop} {H₁ : decidable c₁} {A : Type} {t₁ t₂ e₁ e₂ : A} (Hc : c₁ = c₂) (Ht : t₁ = t₂) (He : e₁ = e₂) :
(if c₁ then t₁ else e₁) = (@ite c₂ (decidable_eq_equiv H₁ Hc) A t₂ e₂) :=
have H2 [fact] : decidable c₂, from (decidable_eq_equiv H₁ Hc),
if_congr_aux Hc Ht He
exit
theorem app_if_distribute {A B : Type} (c : Prop) {H : decidable c} (f : A → B) (a b : A) : f (if c then a else b) = if c then f a else f b :=
or_elim (decidable.em H)
(assume Hc : c, calc
f (if c then a else b) = f (if true then a else b) : { eqt_intro Hc }
(assume Hnc : ¬c, sorry)
exit
:= or_elim (em c)
(λ Hc : c , calc
f (if c then a else b) = f (if true then a else b) : { eqt_intro Hc }
... = f a : { if_true a b }
... = if true then f a else f b : symm (if_true (f a) (f b))
... = if c then f a else f b : { symm (eqt_intro Hc) })
(λ Hnc : ¬ c, calc
f (if c then a else b) = f (if false then a else b) : { eqf_intro Hnc }
... = f b : { if_false a b }
... = if false then f a else f b : symm (if_false (f a) (f b))
... = if c then f a else f b : { symm (eqf_intro Hnc) })