From iris.program_logic Require Import language. From iris.heap_lang Require Export lang tactics. Set Default Proof Using "Type". Delimit Scope expr_scope with E. Delimit Scope val_scope with V. Coercion LitInt : Z >-> base_lit. Coercion LitBool : bool >-> base_lit. Coercion LitLoc : loc >-> base_lit. Coercion LitProphecy : proph_id >-> base_lit. Coercion App : expr >-> Funclass. Coercion Val : val >-> expr. Coercion Var : string >-> expr. Coercion BNamed : string >-> binder. Notation "<>" := BAnon : binder_scope. (* No scope for the values, does not conflict and scope is often not inferred properly. *) Notation "# l" := (LitV l%Z%V) (at level 8, format "# l"). (** Syntax inspired by Coq/Ocaml. Constructions with higher precedence come first. *) Notation "( e1 , e2 , .. , en )" := (Pair .. (Pair e1 e2) .. en) : expr_scope. Notation "( e1 , e2 , .. , en )" := (PairV .. (PairV e1 e2) .. en) : val_scope. (* Using the '[hv' ']' printing box, we make sure that when the notation for match does not fit on a single line, line breaks will be inserted for *each* breaking point '/'. Note that after each breaking point /, one can put n spaces (for example '/ '). That way, when the breaking point is turned into a line break, indentation of n spaces will appear after the line break. As such, when the match does not fit on one line, it will print it like: match: e0 with InjL x1 => e1 | InjR x2 => e2 end Moreover, if the branches do not fit on a single line, it will be printed as: match: e0 with InjL x1 => lots of stuff bla bla bla bla bla bla bla bla | InjR x2 => even more stuff bla bla bla bla bla bla bla bla end *) Notation "'match:' e0 'with' 'InjL' x1 => e1 | 'InjR' x2 => e2 'end'" := (Match e0 x1%bind e1 x2%bind e2) (e0, x1, e1, x2, e2 at level 200, format "'[hv' 'match:' e0 'with' '/ ' '[' 'InjL' x1 => '/ ' e1 ']' '/' '[' | 'InjR' x2 => '/ ' e2 ']' '/' 'end' ']'") : expr_scope. Notation "'match:' e0 'with' 'InjR' x1 => e1 | 'InjL' x2 => e2 'end'" := (Match e0 x2%bind e2 x1%bind e1) (e0, x1, e1, x2, e2 at level 200, only parsing) : expr_scope. Notation "()" := LitUnit : val_scope. Notation "! e" := (Load e%E) (at level 9, right associativity) : expr_scope. Notation "'ref' e" := (Alloc e%E) (at level 10) : expr_scope. Notation "- e" := (UnOp MinusUnOp e%E) : expr_scope. Notation "e1 + e2" := (BinOp PlusOp e1%E e2%E) : expr_scope. Notation "e1 - e2" := (BinOp MinusOp e1%E e2%E) : expr_scope. Notation "e1 * e2" := (BinOp MultOp e1%E e2%E) : expr_scope. Notation "e1 `quot` e2" := (BinOp QuotOp e1%E e2%E) : expr_scope. Notation "e1 `rem` e2" := (BinOp RemOp e1%E e2%E) : expr_scope. Notation "e1 ≪ e2" := (BinOp ShiftLOp e1%E e2%E) : expr_scope. Notation "e1 ≫ e2" := (BinOp ShiftROp e1%E e2%E) : expr_scope. Notation "e1 ≤ e2" := (BinOp LeOp e1%E e2%E) : expr_scope. Notation "e1 < e2" := (BinOp LtOp e1%E e2%E) : expr_scope. Notation "e1 = e2" := (BinOp EqOp e1%E e2%E) : expr_scope. Notation "e1 ≠ e2" := (UnOp NegOp (BinOp EqOp e1%E e2%E)) : expr_scope. Notation "~ e" := (UnOp NegOp e%E) (at level 75, right associativity) : expr_scope. (* The unicode ← is already part of the notation "_ ← _; _" for bind. *) Notation "e1 <- e2" := (Store e1%E e2%E) (at level 80) : expr_scope. (* The breaking point '/ ' makes sure that the body of the rec is indented by two spaces in case the whole rec does not fit on a single line. *) Notation "'rec:' f x := e" := (Rec f%bind x%bind e%E) (at level 200, f at level 1, x at level 1, e at level 200, format "'[' 'rec:' f x := '/ ' e ']'") : expr_scope. Notation "'rec:' f x := e" := (RecV f%bind x%bind e%E) (at level 200, f at level 1, x at level 1, e at level 200, format "'[' 'rec:' f x := '/ ' e ']'") : val_scope. Notation "'if:' e1 'then' e2 'else' e3" := (If e1%E e2%E e3%E) (at level 200, e1, e2, e3 at level 200) : expr_scope. (** Derived notions, in order of declaration. The notations for let and seq are stated explicitly instead of relying on the Notations Let and Seq as defined above. This is needed because App is now a coercion, and these notations are otherwise not pretty printed back accordingly. *) Notation "'rec:' f x y .. z := e" := (Rec f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..)) (at level 200, f, x, y, z at level 1, e at level 200, format "'[' 'rec:' f x y .. z := '/ ' e ']'") : expr_scope. Notation "'rec:' f x y .. z := e" := (RecV f%bind x%bind (Lam y%bind .. (Lam z%bind e%E) ..)) (at level 200, f, x, y, z at level 1, e at level 200, format "'[' 'rec:' f x y .. z := '/ ' e ']'") : val_scope. (* The breaking point '/ ' makes sure that the body of the λ: is indented by two spaces in case the whole λ: does not fit on a single line. *) Notation "λ: x , e" := (Lam x%bind e%E) (at level 200, x at level 1, e at level 200, format "'[' 'λ:' x , '/ ' e ']'") : expr_scope. Notation "λ: x y .. z , e" := (Lam x%bind (Lam y%bind .. (Lam z%bind e%E) ..)) (at level 200, x, y, z at level 1, e at level 200, format "'[' 'λ:' x y .. z , '/ ' e ']'") : expr_scope. Notation "λ: x , e" := (LamV x%bind e%E) (at level 200, x at level 1, e at level 200, format "'[' 'λ:' x , '/ ' e ']'") : val_scope. Notation "λ: x y .. z , e" := (LamV x%bind (Lam y%bind .. (Lam z%bind e%E) .. )) (at level 200, x, y, z at level 1, e at level 200, format "'[' 'λ:' x y .. z , '/ ' e ']'") : val_scope. Notation "'let:' x := e1 'in' e2" := (Lam x%bind e2%E e1%E) (at level 200, x at level 1, e1, e2 at level 200, format "'[' 'let:' x := '[' e1 ']' 'in' '/' e2 ']'") : expr_scope. Notation "e1 ;; e2" := (Lam BAnon e2%E e1%E) (at level 100, e2 at level 200, format "'[' '[hv' '[' e1 ']' ;; ']' '/' e2 ']'") : expr_scope. (* Shortcircuit Boolean connectives *) Notation "e1 && e2" := (If e1%E e2%E (LitV (LitBool false))) (only parsing) : expr_scope. Notation "e1 || e2" := (If e1%E (LitV (LitBool true)) e2%E) (only parsing) : expr_scope. (** Notations for option *) Notation NONE := (InjL (LitV LitUnit)) (only parsing). Notation NONEV := (InjLV (LitV LitUnit)) (only parsing). Notation SOME x := (InjR x) (only parsing). Notation SOMEV x := (InjRV x) (only parsing). Notation "'match:' e0 'with' 'NONE' => e1 | 'SOME' x => e2 'end'" := (Match e0 BAnon e1 x%bind e2) (e0, e1, x, e2 at level 200, only parsing) : expr_scope. Notation "'match:' e0 'with' 'SOME' x => e2 | 'NONE' => e1 'end'" := (Match e0 BAnon e1 x%bind e2) (e0, e1, x, e2 at level 200, only parsing) : expr_scope. Notation "'resolve_proph:' p 'to:' v" := (ResolveProph p v) (at level 100) : expr_scope.