{-# OPTIONS --safe #-}
module Cubical.Functions.Fixpoint where
open import Cubical.Foundations.Prelude
open import Cubical.Foundations.Function
open import Cubical.Foundations.GroupoidLaws
private
  variable
    ℓ : Level
    A : Type ℓ
Fixpoint : (A → A) → Type _
Fixpoint {A = A} f = Σ A (λ x → f x ≡ x)
fixpoint : {f : A → A} → Fixpoint f → A
fixpoint = fst
fixpointPath : {f : A → A} → (p : Fixpoint f) → f (fixpoint p) ≡ fixpoint p
fixpointPath = snd
2-Constant→isPropFixpoint : (f : A → A) → 2-Constant f → isProp (Fixpoint f)
2-Constant→isPropFixpoint f fconst (x , p) (y , q) i = s i , t i where
  noose : ∀ x y → f x ≡ f y
  noose x y = sym (fconst x x) ∙ fconst x y
  
  
  KrausInsight : ∀ {x y} → (p : x ≡ y) → noose x y ≡ cong f p
  KrausInsight {x} = J (λ y p → noose x y ≡ cong f p) (lCancel (fconst x x))
  
  
  s : x ≡ y
  s = sym p ∙∙ noose x y ∙∙ q
  t' : PathP (λ i → noose x y i ≡ s i) p q
  t' i j = doubleCompPath-filler (sym p) (noose x y) q j i
  t : PathP (λ i → cong f s i ≡ s i) p q
  t = subst (λ kraus → PathP (λ i → kraus i ≡ s i) p q) (KrausInsight s) t'