Skip to content

Commit f49e8a0

Browse files
authored
[ add ] wellfoundedness of Relation.Binary.Construct.Add.Infimum.Strict (#2683)
* add: wellfoundedness of lifted relation * tweak * Update CHANGELOG.md fix: name * Update src/Relation/Binary/Construct/Add/Infimum/Strict.agda fix: name
1 parent 4871f37 commit f49e8a0

File tree

2 files changed

+28
-5
lines changed

2 files changed

+28
-5
lines changed

CHANGELOG.md

+7
Original file line numberDiff line numberDiff line change
@@ -200,6 +200,13 @@ Additions to existing modules
200200
filter-↭ : ∀ (P? : Pred.Decidable P) → xs ↭ ys → filter P? xs ↭ filter P? ys
201201
```
202202

203+
* In `Relation.Binary.Construct.Add.Infimum.Strict`:
204+
```agda
205+
<₋-accessible-⊥₋ : Acc _<₋_ ⊥₋
206+
<₋-accessible[_] : Acc _<_ x → Acc _<₋_ [ x ]
207+
<₋-wellFounded : WellFounded _<_ → WellFounded _<₋_
208+
```
209+
203210
* In `Relation.Nullary.Decidable.Core`:
204211
```agda
205212
⊤-dec : Dec {a} ⊤

src/Relation/Binary/Construct/Add/Infimum/Strict.agda

+21-5
Original file line numberDiff line numberDiff line change
@@ -1,7 +1,7 @@
11
------------------------------------------------------------------------
22
-- The Agda standard library
33
--
4-
-- The lifting of a non-strict order to incorporate a new infimum
4+
-- The lifting of a strict order to incorporate a new infimum
55
------------------------------------------------------------------------
66

77
{-# OPTIONS --cubical-compatible --safe #-}
@@ -17,6 +17,7 @@ module Relation.Binary.Construct.Add.Infimum.Strict
1717
open import Level using (_⊔_)
1818
open import Data.Product.Base using (_,_; map)
1919
open import Function.Base using (_∘_)
20+
open import Induction.WellFounded using (WfRec; Acc; acc; WellFounded)
2021
open import Relation.Binary.PropositionalEquality.Core
2122
using (_≡_; refl; cong; subst)
2223
import Relation.Binary.PropositionalEquality.Properties as ≡
@@ -34,6 +35,7 @@ open import Relation.Nullary.Construct.Add.Infimum
3435
using (⊥₋; [_]; _₋; ≡-dec; []-injective)
3536
open import Relation.Nullary.Decidable.Core as Dec using (yes; no; map′)
3637

38+
3739
------------------------------------------------------------------------
3840
-- Definition
3941

@@ -71,14 +73,28 @@ module _ {r} {_≤_ : Rel A r} where
7173
open NonStrict _≤_
7274

7375
<₋-transʳ : Trans _≤_ _<_ _<_ Trans _≤₋_ _<₋_ _<₋_
74-
<₋-transʳ <-transʳ (⊥₋≤ .⊥₋) (⊥₋<[ l ]) = ⊥₋<[ l ]
75-
<₋-transʳ <-transʳ (⊥₋≤ l) [ q ] = ⊥₋<[ _ ]
76-
<₋-transʳ <-transʳ [ p ] [ q ] = [ <-transʳ p q ]
76+
<₋-transʳ <-transʳ (⊥₋≤ ⊥₋) q = q
77+
<₋-transʳ <-transʳ (⊥₋≤ _) [ q ] = ⊥₋<[ _ ]
78+
<₋-transʳ <-transʳ [ p ] [ q ] = [ <-transʳ p q ]
7779

7880
<₋-transˡ : Trans _<_ _≤_ _<_ Trans _<₋_ _≤₋_ _<₋_
79-
<₋-transˡ <-transˡ ⊥₋<[ l ] [ q ] = ⊥₋<[ _ ]
81+
<₋-transˡ <-transˡ ⊥₋<[ _ ] [ q ] = ⊥₋<[ _ ]
8082
<₋-transˡ <-transˡ [ p ] [ q ] = [ <-transˡ p q ]
8183

84+
<₋-accessible-⊥₋ : Acc _<₋_ ⊥₋
85+
<₋-accessible-⊥₋ = acc λ()
86+
87+
<₋-accessible[_] : {x} Acc _<_ x Acc _<₋_ [ x ]
88+
<₋-accessible[_] = acc ∘ wf-acc
89+
where
90+
wf-acc : {x} Acc _<_ x WfRec _<₋_ (Acc _<₋_) [ x ]
91+
wf-acc _ ⊥₋<[ _ ] = <₋-accessible-⊥₋
92+
wf-acc (acc ih) [ y<x ] = <₋-accessible[ ih y<x ]
93+
94+
<₋-wellFounded : WellFounded _<_ WellFounded _<₋_
95+
<₋-wellFounded wf ⊥₋ = <₋-accessible-⊥₋
96+
<₋-wellFounded wf [ x ] = <₋-accessible[ wf x ]
97+
8298
------------------------------------------------------------------------
8399
-- Relational properties + propositional equality
84100

0 commit comments

Comments
 (0)