Add renaming
This commit is contained in:
parent
c28102c0f0
commit
40cedc388f
2 changed files with 113 additions and 9 deletions
|
@ -8,27 +8,31 @@ Definition get_level a :=
|
|||
| Floating ℓ => ℓ
|
||||
end.
|
||||
|
||||
Inductive lookup {A} : nat -> list A -> A -> Prop :=
|
||||
| here A Γ : lookup 0 (cons A Γ) A
|
||||
| there i Γ A B :
|
||||
lookup i Γ A ->
|
||||
lookup (S i) (cons B Γ) A.
|
||||
|
||||
Definition ctx := list (Level * Ty).
|
||||
|
||||
Fixpoint fix_ctx (Γ : ctx) :=
|
||||
match Γ with
|
||||
| nil => nil
|
||||
| (Fixed ℓ, A) :: Γ => (Fixed ℓ, A) :: fix_ctx Γ
|
||||
| (Floating ℓ, A) :: Γ => (Fixed ℓ, A) :: fix_ctx Γ
|
||||
Definition squash a :=
|
||||
match a with
|
||||
| Fixed ℓ => false
|
||||
| Floating ℓ => true
|
||||
end.
|
||||
|
||||
Fixpoint fixed_ctx (Γ : ctx) :=
|
||||
match Γ with
|
||||
| nil => nil
|
||||
| (Fixed ℓ, A) :: Γ => false :: fixed_ctx Γ
|
||||
| (Floating ℓ, A) :: Γ => true :: fixed_ctx Γ
|
||||
| (a, _) :: Γ => squash a :: fixed_ctx Γ
|
||||
end.
|
||||
|
||||
(* false <= true *)
|
||||
(* A term is liftable if it remains well-typed raising the levels of the floating levels to the current level *)
|
||||
Fixpoint liftable Φ (a : Tm) :=
|
||||
match a with
|
||||
| VarTm i => nth_default true Φ i = false
|
||||
| VarTm i => lookup i Φ false
|
||||
| FlApp a b => liftable Φ a /\ liftable Φ b
|
||||
| App a b => liftable Φ a
|
||||
| FlAbs a => liftable (false :: Φ) a
|
||||
|
@ -38,7 +42,7 @@ Fixpoint liftable Φ (a : Tm) :=
|
|||
Inductive Wt (Γ : ctx) : Tm -> bool -> Ty -> Prop :=
|
||||
| T_Var ℓ i ℓ0 A :
|
||||
Bool.le (get_level ℓ0) ℓ ->
|
||||
nth_error Γ i = Some (ℓ0, A) ->
|
||||
lookup i Γ (ℓ0, A) ->
|
||||
Γ ⊢ VarTm i ;; ℓ ∈ A
|
||||
|
||||
| T_FlApp ℓ a b A B :
|
||||
|
@ -53,6 +57,7 @@ Inductive Wt (Γ : ctx) : Tm -> bool -> Ty -> Prop :=
|
|||
| T_App ℓ ℓ0 a b A B :
|
||||
Γ ⊢ b ;; ℓ ∈ Fun ℓ0 A B ->
|
||||
Γ ⊢ a ;; ℓ0 ∈ A ->
|
||||
liftable (fixed_ctx Γ) a ->
|
||||
Γ ⊢ App b a ;; ℓ ∈ A
|
||||
|
||||
| T_Abs ℓ ℓ0 b A B :
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue