{-
This second-order term syntax was created from the following second-order syntax description:

syntax STLC | Λ

type
  N   : 0-ary
  _↣_ : 2-ary | r30

term
  app : α ↣ β  α  ->  β | _$_ l20
  lam : α.β  ->  α ↣ β | ƛ_ r10

theory
  (ƛβ) b : α.β  a : α |> app (lam(x.b[x]), a) = b[a]
  (ƛη) f : α ↣ β      |> lam (x. app(f, x))   = f
-}


module STLC.Syntax where

open import SOAS.Common
open import SOAS.Context
open import SOAS.Variable
open import SOAS.Families.Core
open import SOAS.Construction.Structure
open import SOAS.ContextMaps.Inductive

open import SOAS.Metatheory.Syntax

open import STLC.Signature

private
  variable
    Γ Δ Π : Ctx
    α β : ΛT
    𝔛 : Familyₛ

-- Inductive term declaration
module Λ:Terms (𝔛 : Familyₛ) where

  data Λ : Familyₛ where
    var  :  ⇾̣ Λ
    mvar : 𝔛 α Π  Sub Λ Π Γ  Λ α Γ

    _$_ : Λ (α  β) Γ  Λ α Γ  Λ β Γ
    ƛ_  : Λ β (α  Γ)  Λ (α  β) Γ

  infixl 20 _$_
  infixr 10 ƛ_

  open import SOAS.Metatheory.MetaAlgebra ⅀F 𝔛

  Λᵃ : MetaAlg Λ
  Λᵃ = record
    { 𝑎𝑙𝑔 = λ where
      (appₒ  a , b)  _$_ a b
      (lamₒ  a)      ƛ_  a
    ; 𝑣𝑎𝑟 = var ; 𝑚𝑣𝑎𝑟 = λ 𝔪   mvar 𝔪 (tabulate ) }

  module Λᵃ = MetaAlg Λᵃ

  module _ {𝒜 : Familyₛ}(𝒜ᵃ : MetaAlg 𝒜) where

    open MetaAlg 𝒜ᵃ

    𝕤𝕖𝕞 : Λ ⇾̣ 𝒜
    𝕊 : Sub Λ Π Γ  Π ~[ 𝒜 ]↝ Γ
    𝕊 (t  σ) new = 𝕤𝕖𝕞 t
    𝕊 (t  σ) (old v) = 𝕊 σ v
    𝕤𝕖𝕞 (mvar 𝔪 ) = 𝑚𝑣𝑎𝑟 𝔪 (𝕊 )
    𝕤𝕖𝕞 (var v) = 𝑣𝑎𝑟 v

    𝕤𝕖𝕞 (_$_ a b) = 𝑎𝑙𝑔 (appₒ  𝕤𝕖𝕞 a , 𝕤𝕖𝕞 b)
    𝕤𝕖𝕞 (ƛ_  a)   = 𝑎𝑙𝑔 (lamₒ  𝕤𝕖𝕞 a)

    𝕤𝕖𝕞ᵃ⇒ : MetaAlg⇒ Λᵃ 𝒜ᵃ 𝕤𝕖𝕞
    𝕤𝕖𝕞ᵃ⇒ = record
      { ⟨𝑎𝑙𝑔⟩ = λ{ {t = t}  ⟨𝑎𝑙𝑔⟩ t }
      ; ⟨𝑣𝑎𝑟⟩ = refl
      ; ⟨𝑚𝑣𝑎𝑟⟩ = λ{ {𝔪 = 𝔪}{}  cong (𝑚𝑣𝑎𝑟 𝔪) (dext (𝕊-tab )) }  }
      where
      open ≡-Reasoning
      ⟨𝑎𝑙𝑔⟩ : (t :  Λ α Γ)  𝕤𝕖𝕞 (Λᵃ.𝑎𝑙𝑔 t)  𝑎𝑙𝑔 (⅀₁ 𝕤𝕖𝕞 t)
      ⟨𝑎𝑙𝑔⟩ (appₒ  _) = refl
      ⟨𝑎𝑙𝑔⟩ (lamₒ  _) = refl

      𝕊-tab : ( : Π ~[ Λ ]↝ Γ)(v :  α Π)  𝕊 (tabulate ) v  𝕤𝕖𝕞 ( v)
      𝕊-tab  new = refl
      𝕊-tab  (old v) = 𝕊-tab (  old) v

    module _ (g : Λ ⇾̣ 𝒜)(gᵃ⇒ : MetaAlg⇒ Λᵃ 𝒜ᵃ g) where

      open MetaAlg⇒ gᵃ⇒

      𝕤𝕖𝕞! : (t : Λ α Γ)  𝕤𝕖𝕞 t  g t
      𝕊-ix : ( : Sub Λ Π Γ)(v :  α Π)  𝕊  v  g (index  v)
      𝕊-ix (x  ) new = 𝕤𝕖𝕞! x
      𝕊-ix (x  ) (old v) = 𝕊-ix  v
      𝕤𝕖𝕞! (mvar 𝔪 ) rewrite cong (𝑚𝑣𝑎𝑟 𝔪) (dext (𝕊-ix ))
        = trans (sym ⟨𝑚𝑣𝑎𝑟⟩) (cong (g  mvar 𝔪) (tab∘ix≈id ))
      𝕤𝕖𝕞! (var v) = sym ⟨𝑣𝑎𝑟⟩

      𝕤𝕖𝕞! (_$_ a b) rewrite 𝕤𝕖𝕞! a | 𝕤𝕖𝕞! b = sym ⟨𝑎𝑙𝑔⟩
      𝕤𝕖𝕞! (ƛ_ a) rewrite 𝕤𝕖𝕞! a = sym ⟨𝑎𝑙𝑔⟩


-- Syntax instance for the signature
Λ:Syn : Syntax
Λ:Syn = record
  { ⅀F = ⅀F
  ; ⅀:CS = ⅀:CompatStr
  ; mvarᵢ = Λ:Terms.mvar
  ; 𝕋:Init = λ 𝔛  let open Λ:Terms 𝔛 in record
    {  = Λ  Λᵃ
    ; ⊥-is-initial = record { ! = λ{ {𝒜  𝒜ᵃ}  𝕤𝕖𝕞 𝒜ᵃ  𝕤𝕖𝕞ᵃ⇒ 𝒜ᵃ }
    ; !-unique = λ{ {𝒜  𝒜ᵃ} (f  fᵃ⇒) {x = t}  𝕤𝕖𝕞! 𝒜ᵃ f fᵃ⇒ t } } } }

-- Instantiation of the syntax and metatheory
open Syntax Λ:Syn public
open Λ:Terms public
open import SOAS.Families.Build public
open import SOAS.Syntax.Shorthands Λᵃ public
open import SOAS.Metatheory Λ:Syn public