1
1
------------------------------------------------------------------------
2
2
-- The Agda standard library
3
3
--
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
5
5
------------------------------------------------------------------------
6
6
7
7
{-# OPTIONS --cubical-compatible --safe #-}
@@ -17,6 +17,7 @@ module Relation.Binary.Construct.Add.Infimum.Strict
17
17
open import Level using (_⊔_)
18
18
open import Data.Product.Base using (_,_; map)
19
19
open import Function.Base using (_∘_)
20
+ open import Induction.WellFounded using (WfRec; Acc; acc; WellFounded)
20
21
open import Relation.Binary.PropositionalEquality.Core
21
22
using (_≡_; refl; cong; subst)
22
23
import Relation.Binary.PropositionalEquality.Properties as ≡
@@ -34,6 +35,7 @@ open import Relation.Nullary.Construct.Add.Infimum
34
35
using (⊥₋; [_]; _₋; ≡-dec; []-injective)
35
36
open import Relation.Nullary.Decidable.Core as Dec using (yes; no; map′)
36
37
38
+
37
39
------------------------------------------------------------------------
38
40
-- Definition
39
41
@@ -71,14 +73,28 @@ module _ {r} {_≤_ : Rel A r} where
71
73
open NonStrict _≤_
72
74
73
75
<₋-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 ]
77
79
78
80
<₋-transˡ : Trans _<_ _≤_ _<_ → Trans _<₋_ _≤₋_ _<₋_
79
- <₋-transˡ <-transˡ ⊥₋<[ l ] [ q ] = ⊥₋<[ _ ]
81
+ <₋-transˡ <-transˡ ⊥₋<[ _ ] [ q ] = ⊥₋<[ _ ]
80
82
<₋-transˡ <-transˡ [ p ] [ q ] = [ <-transˡ p q ]
81
83
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
+
82
98
------------------------------------------------------------------------
83
99
-- Relational properties + propositional equality
84
100
0 commit comments