Skip to content

Commit 4084915

Browse files
committed
Conjunction-related operations over Data.So were added.
1 parent 21434e0 commit 4084915

File tree

1 file changed

+33
-0
lines changed

1 file changed

+33
-0
lines changed

libs/base/Data/So.idr

Lines changed: 33 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -23,14 +23,47 @@ choose : (b : Bool) -> Either (So b) (So (not b))
2323
choose True = Left Oh
2424
choose False = Right Oh
2525

26+
--------------------------------------------------------------------------------
27+
-- Absurd- and negation-related properties
28+
--------------------------------------------------------------------------------
29+
2630
||| Absurd when you have proof that both `b` and `not b` is true.
2731
export
2832
soAbsurd : So b -> So (not b) -> Void
2933
soAbsurd sb snb with (sb)
3034
| Oh = uninhabited snb
3135

36+
||| Absurd when you have a proof of both `b` and `not b` (with something else).
37+
export
38+
soConjAbsurd : So b -> So (not b && c) -> Void
39+
soConjAbsurd sb sbc with (sb)
40+
| Oh = uninhabited sbc
41+
3242
||| Transmission between usage of value-level `not` and type-level `Not`.
3343
export
3444
soNotToNotSo : So (not b) -> Not (So b)
3545
soNotToNotSo = flip soAbsurd
3646

47+
--------------------------------------------------------------------------------
48+
--- Operations for `So` of conjunction
49+
--------------------------------------------------------------------------------
50+
51+
||| Given proofs of two properties you have a proof of a conjunction.
52+
export
53+
joinSoConj : So b -> So c -> So (b && c)
54+
joinSoConj sob soc with (sob)
55+
joinSoConj _ soc | Oh with (soc)
56+
joinSoConj _ _ | Oh | Oh = Oh
57+
58+
||| A proof of the right side of a conjunction given a proof of the left side.
59+
export
60+
takeSoConjPart : So b -> So (b && c) -> So c
61+
takeSoConjPart sb sbc with (sb)
62+
takeSoConjPart _ sbc | Oh with (sbc)
63+
takeSoConjPart _ _ | Oh | Oh = Oh
64+
65+
||| Splits the proof of a conjunction to a pair of proofs for each part.
66+
export
67+
splitSoConj : So (b && c) -> (So b, So c)
68+
splitSoConj {b=True} sbc with (sbc)
69+
| Oh = (Oh, Oh)

0 commit comments

Comments
 (0)