Skip to content
Snippets Groups Projects
Commit f6a263a8 authored by Ralf Jung's avatar Ralf Jung
Browse files

improve function type notation

parent d5ab6f51
No related branches found
No related tags found
No related merge requests found
......@@ -12,7 +12,7 @@ Section get_x.
Lemma get_x_type :
typed_instruction_ty [] [] [] get_x
fn( α, [α]; &uniq{α} Π[int; int] &shr{α} int).
(fn( α, [α]; &uniq{α} Π[int; int]) &shr{α} int).
(* FIXME: The above is pretty-printed with some explicit scope annotations,
and without using 'typed_instruction_ty'. I think that's related to
the list notation that we added to %TC. *)
......
......@@ -13,7 +13,7 @@ Section init_prod.
Lemma init_prod_type :
typed_instruction_ty [] [] [] init_prod
fn([]; int, int Π[int;int]).
(fn([]; int, int) Π[int;int]).
Proof.
apply type_fn; try apply _. move=> /= [] ret p. inv_vec p=>x y. simpl_subst.
eapply type_deref; [solve_typing..|apply read_own_move|done|]=>x'. simpl_subst.
......
......@@ -18,7 +18,7 @@ Section lazy_lft.
Lemma lazy_lft_type :
typed_instruction_ty [] [] [] lazy_lft
fn([]; unit).
(fn([]) unit).
Proof.
apply type_fn; try apply _. move=> /= [] ret p. inv_vec p. simpl_subst.
eapply (type_newlft []); [solve_typing|]=>α.
......
......@@ -16,7 +16,7 @@ Section option_as_mut.
Lemma option_as_mut_type τ :
typed_instruction_ty [] [] [] option_as_mut
fn( α, [α]; &uniq{α} Σ[unit; τ] Σ[unit; &uniq{α}τ]).
(fn( α, [α]; &uniq{α} Σ[unit; τ]) Σ[unit; &uniq{α}τ]).
Proof.
apply type_fn; try apply _. move=> /= α ret p. inv_vec p=>o. simpl_subst.
eapply (type_cont [_] [] (λ r, [o _; r!!!0 _])%TC) ; [solve_typing..| |].
......
......@@ -17,7 +17,7 @@ Section rebor.
Lemma rebor_type :
typed_instruction_ty [] [] [] rebor
fn([]; Π[int; int], Π[int; int] int).
(fn([]; Π[int; int], Π[int; int]) int).
Proof.
apply type_fn; try apply _. move=> /= [] ret p. inv_vec p=>t1 t2. simpl_subst.
eapply (type_newlft []). apply _. move=> α.
......
......@@ -12,7 +12,7 @@ Section unbox.
Lemma ubox_type :
typed_instruction_ty [] [] [] unbox
fn( α, [α]; &uniq{α}box (Π[int; int]) &uniq{α} int).
(fn( α, [α]; &uniq{α}box (Π[int; int])) &uniq{α} int).
Proof.
apply type_fn; try apply _. move=> /= α ret b. inv_vec b=>b. simpl_subst.
eapply type_deref; [solve_typing..|by apply read_own_move|done|].
......
......@@ -13,7 +13,7 @@ Section unwrap_or.
Lemma unwrap_or_type τ :
typed_instruction_ty [] [] [] (unwrap_or τ)
fn([]; Σ[unit; τ], τ τ).
(fn([]; Σ[unit; τ], τ) τ).
Proof.
apply type_fn; try apply _. move=> /= [] ret p. inv_vec p=>o def. simpl_subst.
eapply type_case_own; [solve_typing..|]. constructor; last constructor; last constructor.
......
......@@ -67,20 +67,22 @@ End fn.
(* TODO: Once we depend on 8.6pl1, extend this to recursive binders so that
patterns like '(a, b) can be used. *)
Notation "'fn(∀' x ',' E ';' T1 ',' .. ',' TN '→' R ')'" :=
Notation "'fn(∀' x ',' E ';' T1 ',' .. ',' TN ')' '→' R" :=
(fn (λ x, E%EL) (λ x, (Vector.cons T1 .. (Vector.cons TN Vector.nil) ..)%T) (λ x, R%T))
(at level 0, E, T1, TN, R at level 50) : lrust_type_scope.
Notation "'fn(∀' x ',' E ';' '→' R ')'" :=
(at level 99, R at level 200,
format "'fn(∀' x ',' E ';' T1 ',' .. ',' TN ')' '→' R") : lrust_type_scope.
Notation "'fn(∀' x ',' E ')' '→' R" :=
(fn (λ x, E%EL) (λ x, Vector.nil) (λ x, R%T))
(at level 0, E, R at level 50) : lrust_type_scope.
Notation "'fn(' E ';' T1 ',' .. ',' TN '→' R ')'" :=
(at level 99, R at level 200,
format "'fn(∀' x ',' E ')' '→' R") : lrust_type_scope.
Notation "'fn(' E ';' T1 ',' .. ',' TN ')' '→' R" :=
(fn (λ _:(), E%EL) (λ _, (Vector.cons T1 .. (Vector.cons TN Vector.nil) ..)%T) (λ _, R%T))
(at level 0, E, T1, TN, R at level 50,
format "'fn(' E ';' T1 ',' .. ',' TN '→' R ')'") : lrust_type_scope.
Notation "'fn(' E ';' '→' R ')'" :=
(at level 99, R at level 200,
format "'fn(' E ';' T1 ',' .. ',' TN ')' '→' R") : lrust_type_scope.
Notation "'fn(' E ')' '→' R" :=
(fn (λ _:(), E%EL) (λ _, Vector.nil) (λ _, R%T))
(at level 0, E, R at level 50,
format "'fn(' E ';' '→' R ')'") : lrust_type_scope.
(at level 99, R at level 200,
format "'fn(' E ')' '→' R") : lrust_type_scope.
Section typing.
Context `{typeG Σ}.
......
......@@ -24,7 +24,7 @@ Section type_soundness.
Definition exit_cont : val := λ: [<>], #().
Definition main_type `{typeG Σ} :=
fn (λ _, []) (λ _, [# ]) (λ _:(), box unit).
(fn([]) unit)%T.
Theorem type_soundness `{typePreG Σ} (main : val) σ t :
( `{typeG Σ}, typed_instruction_ty [] [] [] main main_type)
......
......@@ -81,7 +81,7 @@ Section typing.
Lemma cell_new_type ty :
typed_instruction_ty [] [] [] cell_new
fn([]; ty cell ty).
(fn([]; ty) cell ty).
Proof.
apply type_fn; [apply _..|]. move=>/= _ ret arg. inv_vec arg=>x. simpl_subst.
eapply (type_jump [_]); first solve_typing.
......@@ -93,7 +93,7 @@ Section typing.
Lemma cell_into_inner_type ty :
typed_instruction_ty [] [] [] cell_into_inner
fn([]; cell ty ty).
(fn([]; cell ty) ty).
Proof.
apply type_fn; [apply _..|]. move=>/= _ ret arg. inv_vec arg=>x. simpl_subst.
eapply (type_jump [_]); first solve_typing.
......@@ -105,7 +105,7 @@ Section typing.
Lemma cell_get_mut_type ty :
typed_instruction_ty [] [] [] cell_get_mut
fn( α, [α]; &uniq{α} (cell ty) &uniq{α} ty).
(fn( α, [α]; &uniq{α} (cell ty)) &uniq{α} ty).
Proof.
apply type_fn; [apply _..|]. move=>/= α ret arg. inv_vec arg=>x. simpl_subst.
eapply (type_jump [_]). solve_typing. rewrite /tctx_incl /=.
......@@ -122,7 +122,7 @@ Section typing.
Lemma cell_get_type `(!Copy ty) :
typed_instruction_ty [] [] [] (cell_get ty)
fn( α, [α]; &shr{α} (cell ty) ty).
(fn( α, [α]; &shr{α} (cell ty)) ty).
Proof.
apply type_fn; [apply _..|]. move=>/= α ret arg. inv_vec arg=>x. simpl_subst.
eapply type_deref; [solve_typing..|apply read_own_move|done|]=>x'. simpl_subst.
......@@ -142,7 +142,7 @@ Section typing.
Lemma cell_set_type ty :
typed_instruction_ty [] [] [] (cell_set ty)
fn( α, [α]; &shr{α} cell ty, ty unit).
(fn( α, [α]; &shr{α} cell ty, ty) unit).
Proof.
apply type_fn; try apply _. move=> /= α ret arg. inv_vec arg=>c x. simpl_subst.
eapply type_deref; [solve_typing..|by apply read_own_move|done|].
......
......@@ -122,7 +122,7 @@ Section refmut_functions.
Lemma refmut_drop_type ty :
typed_instruction_ty [] [] [] refmut_drop
fn( α, [α]; refmut α ty unit).
(fn( α, [α]; refmut α ty) unit).
Proof.
apply type_fn; [apply _..|]. move=>/= α ret arg. inv_vec arg=>x. simpl_subst.
iIntros "!# * #LFT Hna Hα HL Hk Hx".
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment