------------------------------------------------------------------------
-- The Agda standard library
--
-- Induction over Subset
------------------------------------------------------------------------

{-# 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 : 

------------------------------------------------------------------------
-- Re-export accessability

open WF public using (Acc; acc)

------------------------------------------------------------------------
-- Complete induction based on _⊂_

⊂-Rec : RecStruct (Subset n)  
⊂-Rec = WfRec _⊂_

⊂-wellFounded : WellFounded {A = Subset n} _⊂_
⊂-wellFounded = Subrelation.wellFounded p⊂q⇒∣p∣<∣q∣
  (On.wellFounded ∣_∣ <-wellFounded)