{-# OPTIONS --cubical-compatible --safe #-}
module Data.Fin.Subset.Induction where
open import Data.Nat.Base using (ℕ)
open import Data.Nat.Induction using (<-wellFounded)
open import Data.Fin.Subset using (Subset; _⊂_; ∣_∣)
open import Data.Fin.Subset.Properties
open import Induction
open import Induction.WellFounded as WF
open import Level using (Level)
import Relation.Binary.Construct.On as On
private
variable
ℓ : Level
n : ℕ
open WF public using (Acc; acc)
⊂-Rec : RecStruct (Subset n) ℓ ℓ
⊂-Rec = WfRec _⊂_
⊂-wellFounded : WellFounded {A = Subset n} _⊂_
⊂-wellFounded = Subrelation.wellFounded p⊂q⇒∣p∣<∣q∣
(On.wellFounded ∣_∣ <-wellFounded)