35 lines
1.4 KiB
Text
35 lines
1.4 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 decidable tactic
|
||
|
using bit decidable tactic
|
||
|
|
||
|
definition ite (c : Bool) {H : decidable c} {A : Type} (t e : A) : A
|
||
|
:= rec H (assume Hc, t) (assume Hnc, e)
|
||
|
|
||
|
notation `if` c `then` t `else` e:45 := ite c t e
|
||
|
|
||
|
theorem if_pos {c : Bool} {H : decidable c} (Hc : c) {A : Type} (t e : A) : (if c then t else e) = t
|
||
|
:= decidable_rec
|
||
|
(assume Hc : c, calc (@ite c (inl Hc) A t e) = t : refl _)
|
||
|
(assume Hnc : ¬c, absurd_elim (@ite c (inr Hnc) A t e = t) Hc Hnc)
|
||
|
H
|
||
|
|
||
|
theorem if_neg {c : Bool} {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, calc (@ite c (inr Hnc) A t e) = e : refl _)
|
||
|
H
|
||
|
|
||
|
theorem if_t_t (c : Bool) {H : decidable c} {A : Type} (t : A) : (@ite c H A t t) = t -- check why '(if c then t else t) = t' fails
|
||
|
:= decidable_rec
|
||
|
(assume Hc : c, calc (@ite c (inl Hc) A t t) = t : refl _)
|
||
|
(assume Hnc : ¬c, calc (@ite c (inr Hnc) A t t) = t : refl _)
|
||
|
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
|