Substitution versus (OPE-based) renaming

做~自己de王妃 提交于 2020-01-06 19:55:13

问题


Based on this suggestion I am trying to use order-preserving embeddings to represent renamings in a project where I am going to need two levels of contexts (types live in a kinding context of type variable kinds, and terms live in a typing context of term variable types). In my full code, I have substitution for types fully implemented; but a single hole remains in the function that does type-renaming of terms (so this is for renaming type variables). Below is my attempt at recreating the hole with minimal context.

It all starts with a simple List representation of contexts, with a simple strengthening operation that enables easy substitution of variables; and an implementation of OPEs based on the link above:

open import Data.List
open import Data.List.Any
open Data.List.Any.Membership-≡ hiding (Kind)
open import Data.Sum using (_⊎_; inj₁; inj₂; [_,_]′) renaming (map to mapSum)
open import Function using (id)
open import Relation.Binary.PropositionalEquality

module Con where
  Con : Set → Set
  Con = List

  Var : {A : Set} → A → Con A → Set
  Var t Γ = t ∈ Γ

  _∖_ : ∀ {A t} → (Γ : Con A) → Var t Γ → Con A
  (._ ∷ Γ) ∖ here _ = Γ
  (_ ∷ []) ∖ there ()
  (a ∷ b ∷ Γ) ∖ there x = a ∷ (b ∷ Γ) ∖ x



  ∖-there : ∀ {A a b} {Γ : Con A} (x : Var b Γ) → (a ∷ Γ) ∖ there x ≡ a ∷ (Γ ∖ x)
  ∖-there {Γ = []} ()
  ∖-there {Γ = _ ∷ Γ} _ = refl

  strengthen : ∀ {A a b} {Γ : Con A} → (x : Var a Γ) → (y : Var b Γ) → (a ≡ b) ⊎ (Var b (Γ ∖ x))
  strengthen (here refl) (here refl) = inj₁ refl
  strengthen (here _) (there y) = inj₂ y
  strengthen {Γ = _ ∷ []} (there ())
  strengthen {Γ = _ ∷ _ ∷ _} (there y) (here refl) = inj₂ (here refl)
  strengthen {Γ = _ ∷ _ ∷ _} (there x) (there y) = mapSum id (lemma x) (strengthen x y)
    where
      lemma : ∀ {A a b c} {Γ : Con A} (x : Var a Γ) (y : Var b (Γ ∖ x)) → Var b ((c ∷ Γ) ∖ there x)
      lemma {Γ = Γ} x y = subst (Var _) (sym (∖-there x)) (there y)

  substVar : ∀ {A} {_⊢_ : Con A → A → Set}
           → (var : ∀ {a} {Γ : Con A} → Var a Γ → Γ ⊢ a)
           → ∀ {Γ : Con A} {a b}
           → (x : Var a Γ)
           → (e : (Γ ∖ x) ⊢ a)
           → (y : Var b Γ)
           → (Γ ∖ x) ⊢ b
  substVar var x e y = [ (λ eq → subst _ eq e) , var ]′ (strengthen x y)

module OPE where
  open Con

  data _≼_ {A : Set} : Con A -> Con A -> Set where
    done : [] ≼ []
    keep : ∀ {Γ Δ a} -> Γ ≼ Δ -> (a ∷ Γ) ≼ (a ∷ Δ)
    skip : ∀ {Γ Δ a} -> Γ ≼ Δ -> (a ∷ Γ) ≼ Δ

  ≼-refl : ∀ {A} {Γ : Con A} → Γ ≼ Γ
  ≼-refl {Γ = []}     = done
  ≼-refl {Γ = x ∷ Γ} = keep ≼-refl

  weak : ∀ {A} {a : A} {Γ} → (a ∷ Γ) ≼ Γ
  weak = skip ≼-refl

  ren : ∀ {A Γ Γ′} {a : A} → Γ′ ≼ Γ → Var a Γ → Var a Γ′
  ren done ()
  ren (keep σ) (here refl) = here refl
  ren (keep σ) (there x) = there (ren σ x)
  ren (skip σ) x = there (ren σ x)

  ope-∖ : ∀ {A} {Γ Γ′} {t : A} → (x : Var t Γ) → (σ : Γ′ ≼ Γ) → (Γ′ ∖ ren σ x) ≼ (Γ ∖ x)
  ope-∖ (here refl) (keep σ) = σ
  ope-∖ {Γ = _ ∷ Γ} {x′ ∷ Γ′} (here refl) (skip σ) rewrite ∖-there {a = x′} (ren σ (here refl)) = skip (ope-∖ (here refl) σ)
  ope-∖ {Γ = t ∷ Γ} {.t ∷ Γ′} (there x) (keep σ) rewrite ∖-there {a = t} (ren σ x) | ∖-there {a = t} x = keep (ope-∖ x σ)
  ope-∖ {Γ = t ∷ Γ} {t′ ∷ Γ′} (there x) (skip σ) rewrite ∖-there {a = t′} (ren σ (there x)) = skip (ope-∖ (there x) σ)

I then define a type representing well-scoped, well-kinded System F types, and implement substitution of type variables:

module _ where
  open Con
  open OPE

  data Kind : Set where
    ⋆ : Kind

  data _⊢_ (Γ : Con Kind) : Kind → Set where
    tyvar : ∀ {κ} → (α : Var κ Γ) → Γ ⊢ κ
    _⟶_ : (t₁ t₂ : Γ ⊢ ⋆) → Γ ⊢ ⋆
    Pi : ∀ {κ′} → (κ : Kind) → (t : (κ ∷ Γ) ⊢ κ′) → Γ ⊢ κ′

  renType : ∀ {κ} {Γ Γ′ : Con Kind} → Γ′ ≼ Γ → Γ ⊢ κ → Γ′ ⊢ κ
  renType σ (tyvar α) = tyvar (ren σ α)
  renType σ (t₁ ⟶ t₂) = renType σ t₁ ⟶ renType σ t₂
  renType σ (Pi κ t) = Pi κ (renType (keep σ) t)

  substType :  ∀ {κ κ′} {Γ : Con Kind}
            → (β : Var κ′ Γ)
            → (t′ : (Γ ∖ β) ⊢ κ′)
            → (t : Γ ⊢ κ)
            → (Γ ∖ β) ⊢ κ
  substType β t′ (tyvar α) = substVar tyvar β t′ α
  substType β t′ (t₁ ⟶ t₂) = substType β t′ t₁ ⟶ substType β t′ t₂
  substType {Γ = []} () t′ (Pi κ t)
  substType {Γ = _ ∷ Γ} β t′ (Pi κ t) = Pi κ (substType (there β) (renType weak t′) t)

And now I finally have everything in place to formulate my question: how do I prove the following relationship between substType and ren?

  substType-ren : ∀ {κ κ′ Γ Γ′}
                → (σ : Γ′ ≼ Γ)
                → (β : Var κ′ Γ)
                → (t′ : (Γ ∖ β) ⊢ κ′)
                → (t : Γ ⊢ κ)
                → substType (ren σ β) (renType (ope-∖ β σ) t′) (renType σ t) ≡ renType (ope-∖ β σ) (substType β t′ t)
  substType-ren σ β t′ t = {!!}

Appendix

Just to show my work so far, here are the cases I was able to cover:

  strengthen-self : ∀ {A a} {Γ : Con A} (x : Var a Γ) → strengthen x x ≡ inj₁ refl
  strengthen-self (here refl) = refl
  strengthen-self {Γ = x ∷ []} (there ())
  strengthen-self {Γ = _ ∷ _ ∷ Γ} (there x) rewrite strengthen-self x = refl

  substVar-self : ∀ {A} {_⊢_ : Con A → A → Set}
                → (var : ∀ {a} {Γ : Con A} → Var a Γ → Γ ⊢ a)
                → ∀ {Γ : Con A} {a}
                → (x : Var a Γ)
                → (e : (Γ ∖ x) ⊢ a)
                → substVar var x e x ≡ e
  substVar-self var x e rewrite strengthen-self x = refl

  substType-ren : ∀ {κ κ′ Γ Γ′}
                → (σ : Γ′ ≼ Γ)
                → (β : Var κ′ Γ)
                → (t′ : (Γ ∖ β) ⊢ κ′)
                → (t : Γ ⊢ κ)
                → substType (ren σ β) (renType (ope-∖ β σ) t′) (renType σ t) ≡ renType (ope-∖ β σ) (substType β t′ t)
  substType-ren σ (here refl) t′ (tyvar (here refl))
    rewrite substVar-self tyvar (ren σ (here refl)) (renType (ope-∖ (here refl) σ) t′) = refl
  substType-ren (keep σ) (here refl) t′ (tyvar (there α)) = refl
  substType-ren (skip σ) (here refl) t′ (tyvar (there α)) = {!!}
  substType-ren σ (there β) t′ (tyvar (here refl)) = {!!}
  substType-ren σ (there β) t′ (tyvar (there α)) = {!!}
  substType-ren σ β t′ (t₁ ⟶ t₂) = cong₂ _⟶_ (substType-ren σ β t′ t₁) (substType-ren σ β t′ t₂)
  substType-ren σ β t′ (Pi κ t) = {!!}

Also, I have a ton of lemmas by now about how renType lifts equally rening OPEs into equal functions, or how renTypes compose, or that ren σ is injective; so, the usual suspects, but none has been helpful so far in proving substType-ren.

来源:https://stackoverflow.com/questions/34045074/substitution-versus-ope-based-renaming

易学教程内所有资源均来自网络或用户发布的内容,如有违反法律规定的内容欢迎反馈
该文章没有解决你所遇到的问题?点击提问,说说你的问题,让更多的人一起探讨吧!