2024-08-21 20:03:46 +02:00
|
|
|
|
From Coq Require Import Strings.String.
|
2024-07-24 11:18:57 +02:00
|
|
|
|
Require Import terms.
|
|
|
|
|
|
|
|
|
|
Include Terms.
|
|
|
|
|
|
|
|
|
|
Module Subst.
|
2024-08-21 20:03:46 +02:00
|
|
|
|
|
|
|
|
|
(* Type Variable "x" is a free variable in type *)
|
|
|
|
|
Inductive type_var_free (x:string) : type_term -> Prop :=
|
|
|
|
|
| TFree_Var :
|
|
|
|
|
(type_var_free x (type_var x))
|
|
|
|
|
|
|
|
|
|
| TFree_Ladder : forall τ1 τ2,
|
|
|
|
|
(type_var_free x τ1) ->
|
|
|
|
|
(type_var_free x τ2) ->
|
|
|
|
|
(type_var_free x (type_ladder τ1 τ2))
|
|
|
|
|
|
|
|
|
|
| TFree_Fun : forall τ1 τ2,
|
|
|
|
|
(type_var_free x τ1) ->
|
|
|
|
|
(type_var_free x τ2) ->
|
|
|
|
|
(type_var_free x (type_fun τ1 τ2))
|
|
|
|
|
|
|
|
|
|
| TFree_Morph : forall τ1 τ2,
|
|
|
|
|
(type_var_free x τ1) ->
|
|
|
|
|
(type_var_free x τ2) ->
|
|
|
|
|
(type_var_free x (type_morph τ1 τ2))
|
|
|
|
|
|
|
|
|
|
| TFree_Spec : forall τ1 τ2,
|
|
|
|
|
(type_var_free x τ1) ->
|
|
|
|
|
(type_var_free x τ2) ->
|
|
|
|
|
(type_var_free x (type_spec τ1 τ2))
|
|
|
|
|
|
|
|
|
|
| TFree_Univ : forall y τ,
|
|
|
|
|
~(y = x) ->
|
|
|
|
|
(type_var_free x τ) ->
|
|
|
|
|
(type_var_free x (type_univ y τ))
|
|
|
|
|
.
|
|
|
|
|
|
|
|
|
|
|
|
|
|
|
Open Scope ladder_type_scope.
|
|
|
|
|
Example ex_type_free_var1 :
|
|
|
|
|
(type_var_free "T" (type_univ "U" (type_var "T")))
|
|
|
|
|
.
|
|
|
|
|
Proof.
|
|
|
|
|
apply TFree_Univ.
|
|
|
|
|
easy.
|
|
|
|
|
apply TFree_Var.
|
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
|
|
Open Scope ladder_type_scope.
|
|
|
|
|
Example ex_type_free_var2 :
|
|
|
|
|
~(type_var_free "T" (type_univ "T" (type_var "T")))
|
|
|
|
|
.
|
|
|
|
|
Proof.
|
|
|
|
|
intro H.
|
|
|
|
|
inversion H.
|
|
|
|
|
contradiction.
|
|
|
|
|
Qed.
|
|
|
|
|
|
2024-08-07 15:59:03 +02:00
|
|
|
|
(* scoped variable substitution in type terms $\label{coq:subst-type}$ *)
|
2024-07-25 11:04:56 +02:00
|
|
|
|
Fixpoint type_subst (v:string) (n:type_term) (t0:type_term) :=
|
2024-07-24 11:18:57 +02:00
|
|
|
|
match t0 with
|
|
|
|
|
| type_var name => if (eqb v name) then n else t0
|
|
|
|
|
| type_fun t1 t2 => (type_fun (type_subst v n t1) (type_subst v n t2))
|
2024-07-25 11:04:56 +02:00
|
|
|
|
| type_univ x t => if (eqb v x) then t0 else type_univ x (type_subst v n t)
|
|
|
|
|
| type_spec t1 t2 => (type_spec (type_subst v n t1) (type_subst v n t2))
|
|
|
|
|
| type_ladder t1 t2 => (type_ladder (type_subst v n t1) (type_subst v n t2))
|
2024-07-24 11:18:57 +02:00
|
|
|
|
| t => t
|
|
|
|
|
end.
|
|
|
|
|
|
2024-08-21 20:03:46 +02:00
|
|
|
|
Inductive type_subst1 (x:string) (σ:type_term) : type_term -> type_term -> Prop :=
|
|
|
|
|
| TSubst_VarReplace :
|
|
|
|
|
(type_subst1 x σ (type_var x) σ)
|
|
|
|
|
|
|
|
|
|
| TSubst_VarKeep : forall y,
|
|
|
|
|
~(x = y) ->
|
|
|
|
|
(type_subst1 x σ (type_var y) (type_var y))
|
|
|
|
|
|
|
|
|
|
| TSubst_UnivReplace : forall y τ τ',
|
|
|
|
|
~(x = y) ->
|
|
|
|
|
~(type_var_free y σ) ->
|
|
|
|
|
(type_subst1 x σ τ τ') ->
|
|
|
|
|
(type_subst1 x σ (type_univ y τ) (type_univ y τ'))
|
|
|
|
|
|
|
|
|
|
| TSubst_Id : forall n,
|
|
|
|
|
(type_subst1 x σ (type_id n) (type_id n))
|
|
|
|
|
|
|
|
|
|
| TSubst_Spec : forall τ1 τ2 τ1' τ2',
|
|
|
|
|
(type_subst1 x σ τ1 τ1') ->
|
|
|
|
|
(type_subst1 x σ τ2 τ2') ->
|
|
|
|
|
(type_subst1 x σ (type_spec τ1 τ2) (type_spec τ1' τ2'))
|
|
|
|
|
|
|
|
|
|
| TSubst_Fun : forall τ1 τ1' τ2 τ2',
|
|
|
|
|
(type_subst1 x σ τ1 τ1') ->
|
|
|
|
|
(type_subst1 x σ τ2 τ2') ->
|
|
|
|
|
(type_subst1 x σ (type_fun τ1 τ2) (type_fun τ1' τ2'))
|
|
|
|
|
|
|
|
|
|
| TSubst_Morph : forall τ1 τ1' τ2 τ2',
|
|
|
|
|
(type_subst1 x σ τ1 τ1') ->
|
|
|
|
|
(type_subst1 x σ τ2 τ2') ->
|
|
|
|
|
(type_subst1 x σ (type_morph τ1 τ2) (type_morph τ1' τ2'))
|
|
|
|
|
|
|
|
|
|
| TSubst_Ladder : forall τ1 τ1' τ2 τ2',
|
|
|
|
|
(type_subst1 x σ τ1 τ1') ->
|
|
|
|
|
(type_subst1 x σ τ2 τ2') ->
|
|
|
|
|
(type_subst1 x σ (type_ladder τ1 τ2) (type_ladder τ1' τ2'))
|
|
|
|
|
.
|
|
|
|
|
|
2024-07-24 11:18:57 +02:00
|
|
|
|
(* scoped variable substitution, replaces free occurences of v with n in expression e *)
|
2024-07-25 11:04:56 +02:00
|
|
|
|
Fixpoint expr_subst (v:string) (n:expr_term) (e0:expr_term) :=
|
2024-07-24 11:18:57 +02:00
|
|
|
|
match e0 with
|
|
|
|
|
| expr_var name => if (eqb v name) then n else e0
|
|
|
|
|
| expr_ty_abs x e => if (eqb v x) then e0 else expr_ty_abs x (expr_subst v n e)
|
|
|
|
|
| expr_ty_app e t => expr_ty_app (expr_subst v n e) t
|
2024-08-22 01:04:40 +02:00
|
|
|
|
| expr_abs x t e => if (eqb v x) then e0 else expr_abs x t (expr_subst v n e)
|
|
|
|
|
| expr_morph x t e => if (eqb v x) then e0 else expr_morph x t (expr_subst v n e)
|
|
|
|
|
| expr_app e a => expr_app (expr_subst v n e) (expr_subst v n a)
|
2024-07-24 11:18:57 +02:00
|
|
|
|
| expr_let x t a e => expr_let x t (expr_subst v n a) (expr_subst v n e)
|
|
|
|
|
| expr_ascend t e => expr_ascend t (expr_subst v n e)
|
|
|
|
|
| expr_descend t e => expr_descend t (expr_subst v n e)
|
|
|
|
|
end.
|
|
|
|
|
|
|
|
|
|
(* replace only type variables in expression *)
|
2024-07-25 11:04:56 +02:00
|
|
|
|
Fixpoint expr_specialize (v:string) (n:type_term) (e0:expr_term) :=
|
2024-07-24 11:18:57 +02:00
|
|
|
|
match e0 with
|
|
|
|
|
| expr_ty_app e t => expr_ty_app (expr_specialize v n e) (type_subst v n t)
|
|
|
|
|
| expr_ascend t e => expr_ascend (type_subst v n t) (expr_specialize v n e)
|
|
|
|
|
| expr_descend t e => expr_descend (type_subst v n t) (expr_specialize v n e)
|
|
|
|
|
| e => e
|
|
|
|
|
end.
|
2024-07-25 11:04:56 +02:00
|
|
|
|
|
2024-07-24 11:18:57 +02:00
|
|
|
|
End Subst.
|