You can not select more than 25 topics Topics must start with a letter or number, can include dashes ('-') and can be up to 35 characters long.
semantics-2023/theories/type_systems/systemf/exercises03_sol.v

109 lines
3.2 KiB

This file contains ambiguous Unicode characters!

This file contains ambiguous Unicode characters that may be confused with others in your current locale. If your use case is intentional and legitimate, you can safely ignore this warning. Use the Escape button to highlight these characters.

From stdpp Require Import gmap base relations.
From iris Require Import prelude.
From semantics.ts.systemf Require Import lang notation types tactics.
(** Exercise 3 (LN Exercise 22): Universal Fun *)
Definition fun_comp : val :=
Λ, Λ, Λ, λ: "f" "g" "x", "g" ("f" "x").
Definition fun_comp_type : type :=
: : : (#2 #1) (#1 #0) #2 #0.
Lemma fun_comp_typed :
TY 0; fun_comp : fun_comp_type.
Proof.
(* should be solved by solve_typing. *)
solve_typing.
Qed.
Definition swap_args : val :=
Λ, Λ, Λ, λ: "f" "x" "y", "f" "y" "x".
Definition swap_args_type : type :=
: : : (#2 #1 #0) #1 #2 #0.
Lemma swap_args_typed :
TY 0; swap_args : swap_args_type.
Proof.
(* should be solved by solve_typing. *)
solve_typing.
Qed.
Definition lift_prod : val :=
Λ, Λ, Λ, Λ, λ: "f" "g" "p", ("f" (Fst "p"), "g" (Snd "p")).
Definition lift_prod_type : type :=
(: : : : (#3 #1) (#2 #0) #3 × #2 #1 × #0).
Lemma lift_prod_typed :
TY 0; lift_prod : lift_prod_type.
Proof.
(* should be solved by solve_typing. *)
solve_typing.
Qed.
Definition lift_sum : val :=
Λ, Λ, Λ, Λ, λ: "f" "g" "s",
match: "s" with
InjL "x" => InjL ("f" "x")
| InjR "x" => InjR ("g" "x")
end.
Definition lift_sum_type : type :=
(: : : : (#3 #1) (#2 #0) #3 + #2 #1 + #0).
Lemma lift_sum_typed :
TY 0; lift_sum : lift_sum_type.
Proof.
(* should be solved by solve_typing. *)
solve_typing.
Qed.
(** Exercise 5 (LN Exercise 18): Named to De Bruijn *)
Inductive ptype : Type :=
| PTVar : string ptype
| PInt
| PBool
| PTForall : string ptype ptype
| PTExists : string ptype ptype
| PFun (A B : ptype).
Declare Scope PType_scope.
Delimit Scope PType_scope with pty.
Bind Scope PType_scope with ptype.
Coercion PTVar: string >-> ptype.
Infix "" := PFun : PType_scope.
Notation "∀: x , τ" :=
(PTForall x τ%pty)
(at level 100, τ at level 200) : PType_scope.
Notation "∃: x , τ" :=
(PTExists x τ%pty)
(at level 100, τ at level 200) : PType_scope.
Fixpoint debruijn (m: gmap string nat) (A: ptype) : option type :=
match A with
| PTVar x => match m !! x with None => None | Some n => Some (TVar n) end
| PInt => Some Int
| PBool => Some Bool
| PFun A B => match debruijn m A, debruijn m B with Some A, Some B => Some (A B)%ty | _, _ => None end
| PTForall x A =>
let m' := <[x := 0]> (S <$> m) in
match debruijn m' A with None => None | Some A => Some (TForall A) end
| PTExists x A =>
let m' := <[x := 0]> (S <$> m) in
match debruijn m' A with None => None | Some A => Some (TExists A) end
end.
(* Example *)
Goal debruijn (: "x", : "y", "x" "y")%pty = Some (: : #1 #0)%ty.
Proof.
(* Should be solved by reflexivity. *)
reflexivity.
Qed.
Goal debruijn (: "x", "x" : "y", "y")%pty = Some (: #0 : #0)%ty.
Proof.
(* Should be solved by reflexivity. *)
reflexivity.
Qed.
Goal debruijn (: "x", "x" : "y", "x")%pty = Some (: #0 : #1)%ty.
Proof.
(* Should be solved by reflexivity. *)
reflexivity.
Qed.