File tree Expand file tree Collapse file tree 2 files changed +17
-12
lines changed
Expand file tree Collapse file tree 2 files changed +17
-12
lines changed Original file line number Diff line number Diff line change @@ -250,9 +250,10 @@ Additions to existing modules
250250 if-∧ : ∀ b → (if b ∧ c then x else y) ≡ (if b then (if c then x else y) else y)
251251 if-∨ : ∀ b → (if b ∨ c then x else y) ≡ (if b then x else (if c then x else y))
252252 if-xor : ∀ b → (if b xor c then x else y) ≡ (if b then (if c then y else x) else (if c then x else y))
253- if-congˡ : ∀ b → x ≡ z → (if b then x else y) ≡ (if b then z else y)
254- if-congʳ : ∀ b → y ≡ z → (if b then x else y) ≡ (if b then x else z)
255- if-cong : ∀ b → x ≡ z → y ≡ w → (if b then x else y) ≡ (if b then z else w)
253+ if-cong : b ≡ c → (if b then x else y) ≡ (if c then x else y)
254+ if-cong-then : ∀ b → x ≡ z → (if b then x else y) ≡ (if b then z else y)
255+ if-cong-else : ∀ b → y ≡ z → (if b then x else y) ≡ (if b then x else z)
256+ if-cong₂ : ∀ b → x ≡ z → y ≡ w → (if b then x else y) ≡ (if b then z else w)
256257 ```
257258
258259* In ` Data.Fin.Base ` :
Original file line number Diff line number Diff line change @@ -795,17 +795,21 @@ if-xor false = refl
795795if-xor true {false} = refl
796796if-xor true {true } = refl
797797
798- if-congˡ : ∀ b {x y z : A} → x ≡ z →
799- (if b then x else y) ≡ (if b then z else y)
800- if-congˡ _ refl = refl
798+ if-cong : ∀ {b c} {x y : A} → b ≡ c →
799+ (if b then x else y) ≡ (if c then x else y)
800+ if-cong refl = refl
801801
802- if-congʳ : ∀ b {x y z : A} → y ≡ z →
803- (if b then x else y) ≡ (if b then x else z )
804- if-congʳ _ refl = refl
802+ if-cong-then : ∀ b {x y z : A} → x ≡ z →
803+ (if b then x else y) ≡ (if b then z else y )
804+ if-cong-then _ refl = refl
805805
806- if-cong : ∀ b {x y z w : A} → x ≡ z → y ≡ w →
807- (if b then x else y) ≡ (if b then z else w)
808- if-cong _ refl refl = refl
806+ if-cong-else : ∀ b {x y z : A} → y ≡ z →
807+ (if b then x else y) ≡ (if b then x else z)
808+ if-cong-else _ refl = refl
809+
810+ if-cong₂ : ∀ b {x y z w : A} → x ≡ z → y ≡ w →
811+ (if b then x else y) ≡ (if b then z else w)
812+ if-cong₂ _ refl refl = refl
809813
810814------------------------------------------------------------------------
811815-- Properties of T
You can’t perform that action at this time.
0 commit comments