comparison src/MoreSpecif.v @ 78:c49d999fe806

MoreSpecif
author Adam Chlipala <adamc@hcoop.net>
date Fri, 03 Oct 2008 15:59:59 -0400
parents
children 6c2607211cee
comparison
equal deleted inserted replaced
77:d07c77659c20 78:c49d999fe806
1 (* Copyright (c) 2008, Adam Chlipala
2 *
3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License.
6 * The license text is available at:
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
8 *)
9
10 (* Types and notations presented in Chapter 6 *)
11
12 Set Implicit Arguments.
13
14
15 Notation "!" := (False_rec _ _) : specif_scope.
16 Notation "[ e ]" := (exist _ e _) : specif_scope.
17 Notation "x <== e1 ; e2" := (let (x, _) := e1 in e2)
18 (right associativity, at level 60) : specif_scope.
19
20 Open Local Scope specif_scope.
21 Delimit Scope specif_scope with specif.
22
23 Notation "'Yes'" := (left _ _) : specif_scope.
24 Notation "'No'" := (right _ _) : specif_scope.
25 Notation "'Reduce' x" := (if x then Yes else No) (at level 50) : specif_scope.
26
27 Notation "x || y" := (if x then Yes else Reduce y) (at level 50) : specif_scope.
28
29 Inductive maybe (A : Type) (P : A -> Prop) : Set :=
30 | Unknown : maybe P
31 | Found : forall x : A, P x -> maybe P.
32
33 Notation "{{ x | P }}" := (maybe (fun x => P)) : specif_scope.
34 Notation "??" := (Unknown _) : specif_scope.
35 Notation "[[ x ]]" := (Found _ x _) : specif_scope.
36
37 Notation "x <- e1 ; e2" := (match e1 with
38 | Unknown => ??
39 | Found x _ => e2
40 end)
41 (right associativity, at level 60) : specif_scope.
42
43 Notation "x <-- e1 ; e2" := (match e1 with
44 | inright _ => !!
45 | inleft (exist x _) => e2
46 end)
47 (right associativity, at level 60) : specif_scope.
48
49 Notation "e1 ;; e2" := (if e1 then e2 else ??)
50 (right associativity, at level 60) : specif_scope.
51
52 Notation "e1 ;;; e2" := (if e1 then e2 else !!)
53 (right associativity, at level 60) : specif_scope.