Skip to content

Commit

Permalink
constructor sup of W can be well directed
Browse files Browse the repository at this point in the history
  • Loading branch information
kyoDralliam committed Oct 25, 2023
1 parent f7c8a6a commit 7030a3a
Showing 1 changed file with 34 additions and 2 deletions.
36 changes: 34 additions & 2 deletions theories/DirectedTyping.v
Original file line number Diff line number Diff line change
Expand Up @@ -284,8 +284,10 @@ Module Examples.

Definition ctx := εᵈ ,, Fun \ U @ Discr ,, Cofun \ tProd (tRel 0) U @ Cofun.

Definition sup_ctx (tW : term -> term -> term) :=
ctx ,, Discr \ tRel 1 @ Fun ,, Discr \ tProd (tApp (tRel 1) (tRel 0)) (tW (tRel 3) (tRel 2)) @ Cofun.
Definition sup_ctx
(* tW should bind one variable in its second argument *)
(tW : term -> term -> term) :=
ctx ,, Discr \ tRel 1 @ Fun ,, Discr \ tProd (tApp (tRel 1) (tRel 0)) (tW (tRel 3) (tApp (tRel 3) (tRel 0))) @ Fun.

Section WTyping.
Context `{GenericConsequences}.
Expand All @@ -299,6 +301,36 @@ Module Examples.
+ intros; now eapply wfTypeU.
Qed.

Lemma supWfCtx (tW : term -> term -> term)
(ih : forall Θ A B,
[Θ |-(Fun) A] ->
[Θ ,, Discr \ A @ Fun |-(Cofun) B ] ->
[Θ |-(Fun) tW A B]) :
[|-() sup_ctx tW ].
Proof.
eapply wfCtxCons'; [eapply wfCtxCons'|]; try apply ctxWfCtx; intros.
1: eapply wfTypeVar; tea; [do 2 econstructor|reflexivity].
eapply wfTypeProd'; intros.
+ eapply (wfTypeUniv (dU:=Cofun)).
change U with (U[(tRel 0)..]).
eapply (wfTermApp (A:=(tRel 2))).
1,2: eapply wfTermVar; tea; try reflexivity.
2: econstructor.
change (?d \ tProd _ _ @ ?dT) with (shift_n (d \ tProd (tRel 0) U @ dT) 2).
do 2 econstructor.
+ set (ctx' := _,, _ ,, _).
assert [ctx' |-(Fun) tRel 3].
1:{ eapply wfTypeVar; tea; try reflexivity; repeat econstructor. }
assert [|-() ctx' ,, Discr \ tRel 3 @ Fun] by now eapply wfCtxCons.
apply ih; tea.
eapply (wfTypeUniv (dU:=Cofun)).
change U with (U[(tRel 0)..]); eapply (wfTermApp (A:=(tRel 4))).
1,2: eapply wfTermVar; tea; try reflexivity.
2: econstructor.
change (?d \ tProd _ _ @ ?dT) with (shift_n (d \ tProd (tRel 1) U @ dT) 3).
repeat econstructor.
Qed.

End WTyping.
End W.

Expand Down

0 comments on commit 7030a3a

Please sign in to comment.