├── .gitignore ├── CSL.v ├── CoqMakefile ├── CoqMakefile.conf ├── Delay.v ├── Hoare.v ├── Makefile ├── Monads.v ├── README.md ├── Separation.v ├── Seplog.v ├── Sequences.v ├── _CoqProject └── docs ├── .nojekyll ├── CDF.CSL.html ├── CDF.Delay.html ├── CDF.Hoare.html ├── CDF.Monads.html ├── CDF.Separation.html ├── CDF.Seplog.html ├── CDF.Sequences.html ├── coq2html.css └── coq2html.js /.gitignore: -------------------------------------------------------------------------------- 1 | .*.aux 2 | .*.d 3 | *.a 4 | *.cma 5 | *.cmi 6 | *.cmo 7 | *.cmx 8 | *.cmxa 9 | *.cmxs 10 | *.glob 11 | *.ml.d 12 | *.ml4.d 13 | *.mli.d 14 | *.mllib.d 15 | *.mlpack.d 16 | *.native 17 | *.o 18 | *.v.d 19 | *.vio 20 | *.vo 21 | *.vok 22 | *.vos 23 | .coq-native/ 24 | .csdp.cache 25 | .lia.cache 26 | .nia.cache 27 | .nlia.cache 28 | .nra.cache 29 | csdp.cache 30 | lia.cache 31 | nia.cache 32 | nlia.cache 33 | nra.cache 34 | -------------------------------------------------------------------------------- /CSL.v: -------------------------------------------------------------------------------- 1 | (** Concurrent separation logic. *) 2 | 3 | From Coq Require Import ZArith Lia Bool String List. 4 | From Coq Require Import FunctionalExtensionality PropExtensionality. 5 | From CDF Require Import Sequences Separation. 6 | 7 | Local Open Scope Z_scope. 8 | 9 | (** * 1. A language with pointers and concurrency *) 10 | 11 | (** Here is a variant of the PTR language (from the course on separation logic) 12 | with concurrency (the PAR and ATOMIC constructs). 13 | 14 | Like PTR, it's an ML-like language with immutable variables and 15 | references to mutable memory locations, represented using higher-order 16 | abstract syntax. *) 17 | 18 | Inductive com: Type := 19 | | PURE (x: Z) (**r command without effects *) 20 | | LET (c: com) (f: Z -> com) (**r sequencing of commands *) 21 | | IFTHENELSE (b: Z) (c1 c2: com (**r conditional *)) 22 | | REPEAT (c: com) (**r iterate [c] until it returns not 0 *) 23 | | PAR (c1 c2: com) (**r run [c1] and [c2] in parallel *) 24 | | ATOMIC (c: com) (**r run [c] as one atomic step *) 25 | | ALLOC (sz: nat) (**r allocate [sz] words of storage *) 26 | | GET (l: Z) (**r dereference a pointer *) 27 | | SET (l: Z) (v: Z) (**r assign through a pointer *) 28 | | FREE (l: Z). (**r free one word of storage *) 29 | 30 | (** Some derived forms. *) 31 | 32 | Definition SKIP: com := PURE 0. 33 | 34 | Definition SEQ (c1 c2: com) := LET c1 (fun _ => c2). 35 | 36 | (** Locations that can be read / written right now. *) 37 | 38 | Fixpoint immacc (l: addr) (c: com) : Prop := 39 | match c with 40 | | LET c _ => immacc l c 41 | | PAR c1 c2 => immacc l c1 \/ immacc l c2 42 | | GET l' => l = l' 43 | | SET l' _ => l = l' 44 | | FREE l' => l = l' 45 | | _ => False 46 | end. 47 | 48 | (** Reduction semantics. *) 49 | 50 | Inductive red: com * heap -> com * heap -> Prop := 51 | | red_let_done: forall x f h, 52 | red (LET (PURE x) f, h) (f x, h) 53 | | red_let_step: forall c f h c' h', 54 | red (c, h) (c', h') -> 55 | red (LET c f, h) (LET c' f, h') 56 | | red_ifthenelse: forall b c1 c2 h, 57 | red (IFTHENELSE b c1 c2, h) ((if Z.eqb b 0 then c2 else c1), h) 58 | | red_repeat: forall c h, 59 | red (REPEAT c, h) (LET c (fun b => IFTHENELSE b (PURE b) (REPEAT c)), h) 60 | | red_par_done: forall v1 v2 h, 61 | red (PAR (PURE v1) (PURE v2), h) (SKIP, h) 62 | | red_par_left: forall c1 c2 h c1' h', 63 | red (c1, h) (c1', h') -> 64 | red (PAR c1 c2, h) (PAR c1' c2, h') 65 | | red_par_right: forall c1 c2 h c2' h', 66 | red (c2, h) (c2', h') -> 67 | red (PAR c1 c2, h) (PAR c1 c2', h') 68 | | red_atomic: forall c h v h', 69 | star red (c, h) (PURE v, h') -> 70 | red (ATOMIC c, h) (PURE v, h') 71 | | red_alloc: forall sz (h: heap) l, 72 | (forall i, l <= i < l + Z.of_nat sz -> h i = None) -> 73 | l <> 0 -> 74 | red (ALLOC sz, h) (PURE l, hinit l sz h) 75 | | red_get: forall l (h: heap) v, 76 | h l = Some v -> 77 | red (GET l, h) (PURE v, h) 78 | | red_set: forall l v (h: heap), 79 | h l <> None -> 80 | red (SET l v, h) (SKIP, hupdate l v h) 81 | | red_free: forall l (h: heap), 82 | h l <> None -> 83 | red (FREE l, h) (SKIP, hfree l h). 84 | 85 | (** Run-time errors. This includes race conditions, where a location is 86 | immediately accessed by two commands running in parallel. *) 87 | 88 | Inductive erroneous: com * heap -> Prop := 89 | | erroneous_let: forall c f h, 90 | erroneous (c, h) -> erroneous (LET c f, h) 91 | | erroneous_par_race: forall c1 c2 h l, 92 | immacc l c1 /\ immacc l c2 -> 93 | erroneous (PAR c1 c2, h) 94 | | erroneous_par_l: forall c1 c2 h, 95 | erroneous (c1, h) -> erroneous (PAR c1 c2, h) 96 | | erroneous_par_r: forall c1 c2 h, 97 | erroneous (c2, h) -> erroneous (PAR c1 c2, h) 98 | | erroneous_atomic: forall c h c' h', 99 | star red (c, h) (c', h') -> 100 | erroneous (c', h') -> 101 | erroneous (ATOMIC c, h) 102 | | erroneous_get: forall l (h: heap), 103 | h l = None -> erroneous (GET l, h) 104 | | erroneous_set: forall l v (h: heap), 105 | h l = None -> erroneous (SET l v, h) 106 | | erroneous_free: forall l (h: heap), 107 | h l = None -> erroneous (FREE l, h). 108 | 109 | (** * 2. The rules of concurrent separation logic *) 110 | 111 | Definition invariant := assertion. 112 | Definition precond := assertion. 113 | Definition postcond := Z -> assertion. 114 | 115 | (** ** 2.1. Semantic definition of weak triples *) 116 | 117 | (** We now define "triples" (actually, quadruples) [ J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ ], 118 | where [c] is a command, [P] a precondition (on the initial memory heap), 119 | [Q] a postcondition (on the return value and the final memory heap), 120 | and [J] an invariant about the shared heap that is accessed by atomic 121 | commands. This is a weak triple: termination is not guaranteed. 122 | 123 | As in the [Seplog] module, we define the "triple" semantically 124 | in terms of a [safe n c h1 Q J] predicate over the possible reductions 125 | of command [c] in heap [h1]. 126 | 127 | The definition of [safe] follows Vafeiadis (2011) and uses quantification 128 | over all possible shared heaps [hj] and framing heaps [hf]. 129 | 130 | Step-indexing (the [n] parameter) provides an induction principle 131 | over the [safe] predicate. *) 132 | 133 | Inductive safe: nat -> com -> heap -> postcond -> invariant -> Prop := 134 | | safe_zero: forall c h Q J, 135 | safe O c h Q J 136 | | safe_done: forall n v h (Q: postcond) (J: invariant), 137 | Q v h -> 138 | safe (S n) (PURE v) h Q J 139 | | safe_step: forall n c (h1: heap) (Q: postcond) (J: invariant) 140 | (NOTDONE: match c with PURE _ => False | _ => True end) 141 | (ACC: forall l, immacc l c -> h1 l <> None) 142 | (IMM: forall hf hj h, 143 | hdisj3 h1 hj hf -> 144 | h = hunion h1 (hunion hj hf) -> 145 | J hj -> 146 | ~ erroneous (c, h)) 147 | (STEP: forall hf hj h c' h', 148 | hdisj3 h1 hj hf -> 149 | h = hunion h1 (hunion hj hf) -> 150 | J hj -> 151 | red (c, h) (c', h') -> 152 | exists h1' hj', 153 | hdisj3 h1' hj' hf /\ 154 | h' = hunion h1' (hunion hj' hf) /\ 155 | J hj' /\ safe n c' h1' Q J), 156 | safe (S n) c h1 Q J. 157 | 158 | Definition triple (J: invariant) (P: precond) (c: com) (Q: postcond) := 159 | forall n h, P h -> safe n c h Q J. 160 | 161 | Notation "J '⊢' ⦃ P ⦄ c ⦃ Q ⦄" := (triple J P c Q) (at level 90, c at next level). 162 | 163 | (** ** 2.2. Properties about [safe] *) 164 | 165 | Ltac inv H := inversion H; clear H; subst. 166 | 167 | Lemma safe_pure: forall n v h (Q: postcond) J, 168 | Q v h -> safe n (PURE v) h Q J. 169 | Proof. 170 | intros. destruct n; constructor; auto. 171 | Qed. 172 | 173 | Lemma safe_pure_inv: forall n v h Q J, 174 | safe (S n) (PURE v) h Q J -> Q v h. 175 | Proof. 176 | intros. inv H. auto. contradiction. 177 | Qed. 178 | 179 | Lemma safe_red: forall n c h1 Q J hj hf c' h', 180 | red (c, hunion h1 (hunion hj hf)) (c', h') -> 181 | safe (S n) c h1 Q J -> 182 | J hj -> 183 | hdisj3 h1 hj hf -> 184 | exists h1' hj', 185 | hdisj3 h1' hj' hf /\ 186 | h' = hunion h1' (hunion hj' hf) /\ 187 | J hj' /\ safe n c' h1' Q J. 188 | Proof. 189 | intros. inv H0. 190 | - inv H. 191 | - eauto. 192 | Qed. 193 | 194 | Lemma safe_redN: forall n c h1 Q J hj hf c' h', 195 | starN red n (c, hunion h1 (hunion hj hf)) (c', h') -> 196 | safe (S n) c h1 Q J -> 197 | J hj -> 198 | hdisj3 h1 hj hf -> 199 | exists h1' hj', 200 | hdisj3 h1' hj' hf /\ 201 | h' = hunion h1' (hunion hj' hf) /\ 202 | J hj' /\ safe 1%nat c' h1' Q J. 203 | Proof. 204 | induction n; intros. 205 | - inv H. exists h1, hj; auto. 206 | - inv H. rename c' into c''. rename h' into h''. destruct b as [c' h']. 207 | edestruct safe_red as (h1' & hj' & A & B & C & D). 208 | eassumption. eassumption. assumption. assumption. 209 | subst h'. eapply IHn; eauto. 210 | Qed. 211 | 212 | Lemma safe_not_erroneous: forall n c h1 Q J hj hf, 213 | safe (S n) c h1 Q J -> 214 | hdisj3 h1 hj hf -> 215 | J hj -> 216 | ~ erroneous (c, hunion h1 (hunion hj hf)). 217 | Proof. 218 | intros. inv H. 219 | - intros ST; inv ST. 220 | - eauto. 221 | Qed. 222 | 223 | Lemma safe_immacc: forall n c h1 Q J l, 224 | safe (S n) c h1 Q J -> 225 | immacc l c -> 226 | h1 l <> None. 227 | Proof. 228 | intros. inv H. 229 | - contradiction. 230 | - eapply ACC; auto. 231 | Qed. 232 | 233 | Lemma safe_mono: forall n c h Q J, 234 | safe n c h Q J -> forall n', (n' <= n)%nat -> safe n' c h Q J. 235 | Proof. 236 | induction n; intros. 237 | - replace n' with O by lia. constructor. 238 | - destruct n' as [ | n']. constructor. inv H. 239 | + constructor; auto. 240 | + constructor; auto; intros. 241 | edestruct STEP as (h1' & hj' & A & B & C & D); eauto. 242 | exists h1', hj'; intuition auto. 243 | apply IHn; auto. lia. 244 | Qed. 245 | 246 | (** ** 2.3. The rules of concurrent separation logic *) 247 | 248 | (** *** The frame rule *) 249 | 250 | Lemma safe_frame: 251 | forall (R: assertion) (Q: postcond) J n c h h', 252 | safe n c h Q J -> hdisjoint h h' -> R h' -> 253 | safe n c (hunion h h') (fun v => Q v ** R) J. 254 | Proof. 255 | induction n; intros. 256 | - constructor. 257 | - inv H. 258 | + constructor. exists h, h'; auto. 259 | + constructor; auto. 260 | * intros. apply ACC in H. cbn. destruct (h l); congruence. 261 | * intros. subst h0. 262 | apply (IMM (hunion h' hf) hj). 263 | HDISJ. 264 | rewrite hunion_assoc. f_equal. 265 | rewrite hunion_comm by HDISJ. rewrite hunion_assoc. f_equal. 266 | apply hunion_comm. HDISJ. 267 | auto. 268 | * intros. subst h0. 269 | edestruct (STEP (hunion h' hf) hj) as (h1' & hj' & A & B & C & D). 270 | 4: eauto. HDISJ. 271 | rewrite hunion_assoc. f_equal. 272 | rewrite hunion_comm by HDISJ. rewrite hunion_assoc. f_equal. 273 | apply hunion_comm. HDISJ. 274 | auto. 275 | subst h'0. 276 | exists (hunion h1' h'), hj'. 277 | split. HDISJ. 278 | split. rewrite hunion_assoc. f_equal. 279 | rewrite hunion_comm by HDISJ. rewrite hunion_assoc. f_equal. 280 | apply hunion_comm. HDISJ. 281 | split. assumption. 282 | apply IHn; auto. HDISJ. 283 | Qed. 284 | 285 | Lemma triple_frame: forall J P c Q R, 286 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> 287 | J ⊢ ⦃ P ** R ⦄ c ⦃ fun v => Q v ** R ⦄. 288 | Proof. 289 | intros; red; intros. destruct H0 as (h1 & h2 & P1 & R2 & D & U). 290 | subst h. apply safe_frame; auto. 291 | Qed. 292 | 293 | (** *** The frame rule for invariants *) 294 | 295 | Lemma safe_frame_invariant: 296 | forall Q (J J': invariant) n c h, 297 | safe n c h Q J -> 298 | safe n c h Q (J ** J'). 299 | Proof. 300 | induction n; intros. 301 | - constructor. 302 | - inv H. 303 | + constructor; auto. 304 | + constructor; auto. 305 | * intros. destruct H1 as (hj1 & hj2 & A & B & C & D). subst hj h0. 306 | apply (IMM (hunion hj2 hf) hj1). 307 | HDISJ. 308 | f_equal. rewrite hunion_assoc. auto. 309 | auto. 310 | * intros. destruct H1 as (hj1 & hj2 & A & B & C & D). subst hj h0. 311 | edestruct (STEP (hunion hj2 hf) hj1) as (h1' & hj1' & U & V & X & Y). 312 | 4: eauto. HDISJ. 313 | f_equal. rewrite hunion_assoc. auto. 314 | auto. 315 | subst h'. 316 | exists h1', (hunion hj1' hj2). 317 | split. HDISJ. 318 | split. f_equal. rewrite hunion_assoc. auto. 319 | split. exists hj1', hj2; intuition auto. HDISJ. 320 | apply IHn; auto. 321 | Qed. 322 | 323 | Lemma triple_frame_invariant: forall J J' P c Q, 324 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> 325 | J ** J' ⊢ ⦃ P ⦄ c ⦃ Q ⦄. 326 | Proof. 327 | intros; red; intros. apply safe_frame_invariant; auto. 328 | Qed. 329 | 330 | (** *** Atomic commands *) 331 | 332 | Lemma triple_atomic: forall J P c (Q: postcond), 333 | emp ⊢ ⦃ P ** J ⦄ c ⦃ fun v => Q v ** J ⦄ -> 334 | J ⊢ ⦃ P ⦄ ATOMIC c ⦃ Q ⦄. 335 | Proof. 336 | intros until Q; intros TR n h Ph. destruct n; constructor; auto. 337 | - intros. intro ST; inv ST. 338 | apply star_starN in H4. destruct H4 as (N & STEPS). 339 | rewrite <- hunion_assoc in STEPS. rewrite <- (hunion_empty hf) in STEPS. 340 | edestruct safe_redN as (h1' & hj' & A & B & C & D). 341 | eexact STEPS. 342 | apply TR. exists h, hj. intuition auto. HDISJ. 343 | reflexivity. 344 | HDISJ. 345 | eelim safe_not_erroneous. eexact D. eexact A. eauto. subst h'; auto. 346 | - intros. inv H2. 347 | apply star_starN in H4. destruct H4 as (N & STEPS). 348 | rewrite <- hunion_assoc in STEPS. rewrite <- (hunion_empty hf) in STEPS. 349 | edestruct safe_redN as (h1' & hj' & A & B & C & D). 350 | eexact STEPS. 351 | apply TR. exists h, hj. intuition auto. HDISJ. 352 | reflexivity. 353 | HDISJ. 354 | subst h'. 355 | apply safe_pure_inv in D. destruct D as (h1'' & hj'' & U & V & X & Y). 356 | subst h1'. 357 | exists h1'', hj''. 358 | split. HDISJ. 359 | split. rewrite hunion_assoc. rewrite C. rewrite hunion_empty. auto. 360 | split. auto. 361 | apply safe_pure. auto. 362 | Qed. 363 | 364 | (** *** Sharing some state in the invariant *) 365 | 366 | Lemma safe_share: 367 | forall Q (J J': invariant) n c h h', 368 | safe n c h Q (J ** J') -> 369 | hdisjoint h h' -> J' h' -> 370 | safe n c (hunion h h') (fun v => Q v ** J') J. 371 | Proof. 372 | induction n; intros. 373 | - constructor. 374 | - inv H. 375 | + constructor. exists h, h'; auto. 376 | + constructor; auto. 377 | * intros. apply ACC in H. cbn. destruct (h l); congruence. 378 | * intros. apply (IMM hf (hunion h' hj)). HDISJ. 379 | subst h0. rewrite ! hunion_assoc. auto. 380 | rewrite hunion_comm by HDISJ. exists hj, h'; intuition auto. HDISJ. 381 | * intros. 382 | edestruct (STEP hf (hunion h' hj)) as (h1' & hj' & U & V & X & Y). 383 | 4: eauto. 384 | HDISJ. 385 | subst h0. rewrite ! hunion_assoc. auto. 386 | rewrite hunion_comm by HDISJ. exists hj, h'; intuition auto. HDISJ. 387 | destruct X as (hj1' & hj2' & A & B & C & D). 388 | subst hj' h'0 h0. 389 | exists (hunion h1' hj2'), hj1'. 390 | split. HDISJ. 391 | split. rewrite (hunion_comm hj2') by HDISJ. rewrite ! hunion_assoc. auto. 392 | split. auto. 393 | apply IHn; auto. HDISJ. 394 | Qed. 395 | 396 | Lemma triple_share: forall J J' P c Q, 397 | J ** J' ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> 398 | J ⊢ ⦃ P ** J' ⦄ c ⦃ fun v => Q v ** J' ⦄. 399 | Proof. 400 | intros; intros n h (h1 & h2 & Ph1 & J'h2 & D & U). subst h. 401 | apply safe_share; auto. 402 | Qed. 403 | 404 | (** *** Sequential commands *) 405 | 406 | Lemma triple_pure: forall J P Q v, 407 | aimp P (Q v) -> 408 | J ⊢ ⦃ P ⦄ PURE v ⦃ Q ⦄. 409 | Proof. 410 | intros J P Q v IMP n h Ph. apply safe_pure; auto. 411 | Qed. 412 | 413 | Lemma safe_let: 414 | forall (Q R: postcond) (J: invariant) f n c h, 415 | safe n c h Q J -> 416 | (forall v n' h', Q v h' -> (n' < n)%nat -> safe n' (f v) h' R J) -> 417 | safe n (LET c f) h R J. 418 | Proof. 419 | induction n; intros until h; intros S1 S2. 420 | - constructor. 421 | - constructor; auto; intros. 422 | + eapply safe_immacc; eauto. 423 | + red; intros. inv H2. eelim safe_not_erroneous; eauto. 424 | + subst h0. inv H2. 425 | * exists h, hj; intuition auto. apply S2. eapply safe_pure_inv; eauto. lia. 426 | * edestruct safe_red as (h1' & hj' & A & B & C & D); eauto. 427 | exists h1', hj'; intuition auto. 428 | Qed. 429 | 430 | Lemma triple_let: 431 | forall c f (J: invariant) (P: precond) (Q R: postcond), 432 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> 433 | (forall v, J ⊢ ⦃ Q v ⦄ f v ⦃ R ⦄) -> 434 | J ⊢ ⦃ P ⦄ LET c f ⦃ R ⦄. 435 | Proof. 436 | intros; red; intros. apply safe_let with Q. 437 | apply H; auto. 438 | intros; apply H0; auto. 439 | Qed. 440 | 441 | Corollary triple_seq: 442 | forall c1 c2 (J: invariant) (P Q: precond) (R: postcond), 443 | J ⊢ ⦃ P ⦄ c1 ⦃ fun _ => Q ⦄ -> 444 | J ⊢ ⦃ Q ⦄ c2 ⦃ R ⦄ -> 445 | J ⊢ ⦃ P ⦄ SEQ c1 c2 ⦃ R ⦄. 446 | Proof. 447 | intros. apply triple_let with (fun _ => Q); auto. 448 | Qed. 449 | 450 | (** *** Conditionals and loops *) 451 | 452 | Lemma safe_ifthenelse: 453 | forall n b c1 c2 h Q J, 454 | (b <> 0 -> safe n c1 h Q J) -> 455 | (b = 0 -> safe n c2 h Q J) -> 456 | safe (S n) (IFTHENELSE b c1 c2) h Q J. 457 | Proof. 458 | intros. constructor; auto; intros. 459 | - intro ST; inv ST. 460 | - inv H4. exists h, hj; intuition auto. destruct (Z.eqb_spec b 0); auto. 461 | Qed. 462 | 463 | Lemma triple_ifthenelse: forall J b c1 c2 P Q, 464 | J ⊢ ⦃ (b <> 0) //\\ P ⦄ c1 ⦃ Q ⦄ -> 465 | J ⊢ ⦃ (b = 0) //\\ P ⦄ c2 ⦃ Q ⦄ -> 466 | J ⊢ ⦃ P ⦄ IFTHENELSE b c1 c2 ⦃ Q ⦄. 467 | Proof. 468 | intros; red; intros. 469 | destruct n. constructor. apply safe_ifthenelse. 470 | - intros. apply H. split; auto. 471 | - intros. apply H0. split; auto. 472 | Qed. 473 | 474 | Lemma triple_repeat: forall J P c Q, 475 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> 476 | aimp (Q 0) P -> 477 | J ⊢ ⦃ P ⦄ REPEAT c ⦃ fun v => (v <> 0) //\\ Q v ⦄. 478 | Proof. 479 | intros J P c Q TR IMP. red; induction n; intros h Ph. 480 | - constructor. 481 | - constructor; auto. 482 | + intros; intro ST. inv ST. 483 | + intros. inv H2. exists h, hj; intuition auto. 484 | apply safe_let with Q. 485 | * apply TR; auto. 486 | * intros. destruct n'. constructor. apply safe_ifthenelse. 487 | ** intros. destruct n'; constructor. split; auto. 488 | ** intros. apply safe_mono with n. apply IHn. apply IMP. congruence. lia. 489 | Qed. 490 | 491 | (** *** Parallel composition *) 492 | 493 | Lemma safe_par: 494 | forall (J: invariant) (Q1 Q2: assertion) n c1 h1 c2 h2, 495 | safe n c1 h1 (fun _ => Q1) J -> 496 | safe n c2 h2 (fun _ => Q2) J -> 497 | hdisjoint h1 h2 -> 498 | safe n (PAR c1 c2) (hunion h1 h2) (fun _ => Q1 ** Q2) J. 499 | Proof. 500 | induction n; intros until h2; intros S1 S2 DISJ; constructor; auto. 501 | - cbn; intros. destruct H as [H | H]; eapply safe_immacc in H; eauto. 502 | destruct (h1 l); congruence. 503 | destruct (h1 l); congruence. 504 | - intros. intros ST; inv ST. 505 | + destruct H3 as [IM1 IM2]. 506 | eapply safe_immacc in IM1; eauto. 507 | eapply safe_immacc in IM2; eauto. 508 | specialize (DISJ l). tauto. 509 | + elim (safe_not_erroneous _ _ _ _ _ hj (hunion h2 hf) S1). 510 | HDISJ. 511 | auto. 512 | rewrite hunion_assoc in H3. rewrite <- (hunion_comm hj) by HDISJ. 513 | rewrite hunion_assoc. rewrite (hunion_comm hj) by HDISJ. assumption. 514 | + elim (safe_not_erroneous _ _ _ _ _ hj (hunion h1 hf) S2). 515 | HDISJ. 516 | auto. 517 | rewrite <- (hunion_comm h1) in H3 by HDISJ. 518 | rewrite hunion_assoc in H3. rewrite <- (hunion_comm hj) by HDISJ. 519 | rewrite hunion_assoc. rewrite (hunion_comm hj) by HDISJ. assumption. 520 | - intros; subst h. inv H2. 521 | + (* c1 and c2 are PURE *) 522 | apply safe_pure_inv in S1. apply safe_pure_inv in S2. 523 | exists (hunion h1 h2), hj; intuition auto. 524 | apply safe_pure. exists h1, h2; intuition auto. 525 | + (* c1 reduces *) 526 | rewrite hunion_assoc in H3. rewrite <- (hunion_comm h2) in H3 by HDISJ. 527 | rewrite hunion_assoc in H3. 528 | destruct (safe_red _ _ _ _ _ _ _ _ _ H3 S1) as (h1' & hj' & A & B & C & D). 529 | auto. HDISJ. 530 | subst h'. 531 | exists (hunion h1' h2), hj'. 532 | split. HDISJ. 533 | split. rewrite hunion_assoc. f_equal. rewrite <- (hunion_comm h2) by HDISJ. 534 | rewrite hunion_assoc. auto. 535 | split. assumption. 536 | apply IHn. auto. apply safe_mono with (S n); auto. HDISJ. 537 | + (* c2 reduces *) 538 | rewrite <- (hunion_comm h1) in H3 by HDISJ. 539 | rewrite hunion_assoc in H3. rewrite <- (hunion_comm h1) in H3 by HDISJ. 540 | rewrite hunion_assoc in H3. 541 | destruct (safe_red _ _ _ _ _ _ _ _ _ H3 S2) as (h2' & hj' & A & B & C & D). 542 | auto. HDISJ. 543 | subst h'. 544 | exists (hunion h2' h1), hj'. 545 | split. HDISJ. 546 | split. rewrite hunion_assoc. f_equal. rewrite <- (hunion_comm h1) by HDISJ. 547 | rewrite hunion_assoc. auto. 548 | split. assumption. 549 | rewrite hunion_comm by HDISJ. 550 | apply IHn. apply safe_mono with (S n); auto. auto. HDISJ. 551 | Qed. 552 | 553 | Lemma triple_par: forall J P1 c1 Q1 P2 c2 Q2, 554 | J ⊢ ⦃ P1 ⦄ c1 ⦃ fun _ => Q1 ⦄ -> 555 | J ⊢ ⦃ P2 ⦄ c2 ⦃ fun _ => Q2 ⦄ -> 556 | J ⊢ ⦃ P1 ** P2 ⦄ PAR c1 c2 ⦃ fun _ => Q1 ** Q2 ⦄. 557 | Proof. 558 | intros until Q2; intros TR1 TR2 n h Ph. 559 | destruct Ph as (h1 & h2 & Ph1 & Ph2 & D & U). 560 | subst h. apply safe_par; auto. 561 | Qed. 562 | 563 | (** *** The "small rules" for heap operations *) 564 | 565 | Lemma triple_get: forall J l v, 566 | J ⊢ ⦃ contains l v ⦄ GET l ⦃ fun v' => (v' = v) //\\ contains l v ⦄. 567 | Proof. 568 | intros J l v n h Ph. 569 | assert (L: h l = Some v) by (rewrite Ph; apply hupdate_same). 570 | destruct n; constructor; auto. 571 | - cbn; intros. congruence. 572 | - intros. subst h0. intro ST; inv ST. cbn in H2. rewrite L in H2. congruence. 573 | - intros. subst h0. inv H2. 574 | assert (v0 = v). 575 | { cbn in H3. rewrite L in H3. congruence. } 576 | subst v0. 577 | exists h, hj; intuition auto. 578 | apply safe_pure. split; auto. 579 | Qed. 580 | 581 | Lemma triple_set: forall J l v, 582 | J ⊢ ⦃ valid l ⦄ SET l v ⦃ fun _ => contains l v ⦄. 583 | Proof. 584 | intros J l v n h Ph. 585 | destruct Ph as (v0 & Ph). 586 | assert (L: h l = Some v0) by (rewrite Ph; apply hupdate_same). 587 | destruct n; constructor; auto. 588 | - cbn; intros. congruence. 589 | - intros; intro ST; inv ST. cbn in H3. rewrite L in H3; congruence. 590 | - intros. subst h0. rewrite Ph in H. inv H2. 591 | exists (hupdate l v hempty), hj; intuition auto. 592 | + HDISJ. 593 | red; intros l0; generalize (H l0); cbn. 594 | destruct (Z.eq_dec l l0); intuition congruence. 595 | red; intros l0; generalize (H0 l0); cbn. 596 | destruct (Z.eq_dec l l0); intuition congruence. 597 | + rewrite Ph. apply heap_extensionality; intros l0; cbn. 598 | destruct (Z.eq_dec l l0); auto. 599 | + apply safe_pure. reflexivity. 600 | Qed. 601 | 602 | Fixpoint valid_N (l: addr) (sz: nat) : assertion := 603 | match sz with O => emp | S sz => valid l ** valid_N (l + 1) sz end. 604 | 605 | Remark valid_N_init: 606 | forall sz l, 607 | (valid_N l sz) (hinit l sz hempty). 608 | Proof. 609 | induction sz as [ | sz]; intros l; cbn. 610 | - red; auto. 611 | - exists (hupdate l 0 hempty), (hinit (l + 1) sz hempty). 612 | split. exists 0. red; auto. 613 | split. apply IHsz. 614 | split. intros x. unfold hupdate, hempty at 1; cbn. 615 | destruct (Z.eq_dec l x); auto. 616 | right. rewrite hinit_outside by lia. auto. 617 | apply heap_extensionality; intros x. cbn. destruct (Z.eq_dec l x); auto. 618 | Qed. 619 | 620 | Lemma triple_alloc: forall J sz, 621 | J ⊢ ⦃ emp ⦄ ALLOC sz ⦃ fun l => (l <> 0) //\\ valid_N l sz ⦄. 622 | Proof. 623 | intros J sz n h Ph. red in Ph; subst h. 624 | destruct n; constructor; auto. 625 | - intros; intro ST. inv ST. 626 | - intros. rewrite hunion_empty in H0. subst h. inv H2. 627 | exists (hinit l sz hempty), hj; intuition auto. 628 | + assert (D: hdisjoint (hinit l sz hempty) (hunion hj hf)). 629 | { red; intros l0. 630 | assert (EITHER: l <= l0 < l + Z.of_nat sz \/ l0 < l \/ l + Z.of_nat sz <= l0) by lia. 631 | destruct EITHER. right; auto. left; apply hinit_outside; auto. 632 | } 633 | HDISJ. 634 | + apply heap_extensionality; intros l0; cbn. 635 | assert (EITHER: l <= l0 < l + Z.of_nat sz \/ l0 < l \/ l + Z.of_nat sz <= l0) by lia. 636 | destruct EITHER. 637 | rewrite ! hinit_inside by auto. auto. 638 | rewrite ! hinit_outside by auto. auto. 639 | + apply safe_pure. split. auto. apply valid_N_init. 640 | Qed. 641 | 642 | Lemma triple_free: forall J l, 643 | J ⊢ ⦃ valid l ⦄ FREE l ⦃ fun _ => emp ⦄. 644 | Proof. 645 | intros J l n h Ph. 646 | destruct Ph as (v0 & Ph). 647 | assert (L: h l = Some v0) by (rewrite Ph; apply hupdate_same). 648 | destruct n; constructor; auto. 649 | - cbn; intros. congruence. 650 | - intros; intro ST; inv ST. cbn in H3. rewrite L in H3; congruence. 651 | - intros. subst h0. rewrite Ph in H. inv H2. 652 | exists hempty, hj; intuition auto. 653 | + HDISJ. 654 | + assert (D: hdisjoint (hupdate l v0 hempty) (hunion hj hf)) by HDISJ. 655 | rewrite Ph. apply heap_extensionality; intros l0; cbn. 656 | destruct (Z.eq_dec l l0); auto. 657 | subst l0. destruct (D l); auto. rewrite hupdate_same in H0; congruence. 658 | + apply safe_pure. reflexivity. 659 | Qed. 660 | 661 | (** *** Structural rules *) 662 | 663 | Lemma triple_consequence_pre: forall P' J P c Q, 664 | J ⊢ ⦃ P' ⦄ c ⦃ Q ⦄ -> 665 | aimp P P' -> 666 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄. 667 | Proof. 668 | intros. intros n h Ph. apply H. auto. 669 | Qed. 670 | 671 | Lemma safe_consequence: 672 | forall (Q Q': postcond) (J: invariant), 673 | (forall v, aimp (Q' v) (Q v)) -> 674 | forall n c h, 675 | safe n c h Q' J -> 676 | safe n c h Q J. 677 | Proof. 678 | induction n; intros. 679 | - constructor. 680 | - inv H0. 681 | + constructor. apply H; auto. 682 | + constructor; auto. 683 | intros. 684 | edestruct STEP as (h1' & hj' & A & B & C & D); eauto. 685 | exists h1', hj'; intuition auto. 686 | Qed. 687 | 688 | Lemma triple_consequence_post: 689 | forall Q' J P c Q, 690 | J ⊢ ⦃ P ⦄ c ⦃ Q' ⦄ -> 691 | (forall v, aimp (Q' v) (Q v)) -> 692 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄. 693 | Proof. 694 | intros. intros n h Ph. apply safe_consequence with Q'; auto. 695 | Qed. 696 | 697 | Lemma triple_exists_pre: forall {X: Type} J (P: X -> assertion) c Q, 698 | (forall v, J ⊢ ⦃ P v ⦄ c ⦃ Q ⦄) -> 699 | J ⊢ ⦃ aexists P ⦄ c ⦃ Q ⦄. 700 | Proof. 701 | intros. intros n h Ph. destruct Ph as (v & Ph). apply (H v). auto. 702 | Qed. 703 | 704 | Lemma triple_simple_conj_pre: forall J (P1: Prop) P2 c Q, 705 | (P1 -> J ⊢ ⦃ P2 ⦄ c ⦃ Q ⦄) -> 706 | J ⊢ ⦃ P1 //\\ P2 ⦄ c ⦃ Q ⦄. 707 | Proof. 708 | intros. intros n h Ph. destruct Ph. apply H; auto. 709 | Qed. 710 | 711 | Lemma triple_or: forall J P c Q P' Q', 712 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> J ⊢ ⦃ P' ⦄ c ⦃ Q' ⦄ -> 713 | J ⊢ ⦃ aor P P' ⦄ c ⦃ fun v => aor (Q v) (Q' v) ⦄. 714 | Proof. 715 | intros until Q'; intros TR1 TR2 n h [Ph | P'h]. 716 | - apply safe_consequence with Q. intros v1 h1; red; auto. apply TR1; auto. 717 | - apply safe_consequence with Q'. intros v1 h1; red; auto. apply TR2; auto. 718 | Qed. 719 | 720 | Lemma safe_and: forall J Q Q', 721 | precise J -> 722 | forall n c h, 723 | safe n c h Q J -> safe n c h Q' J -> safe n c h (fun v => aand (Q v) (Q' v)) J. 724 | Proof. 725 | induction n; intros c h S S'. 726 | - constructor. 727 | - inv S; inv S'. 728 | + constructor; split; auto. 729 | + contradiction. 730 | + contradiction. 731 | + constructor; auto. 732 | * intros. 733 | edestruct STEP as (h1' & hj' & D' & E' & J' & SAFE'); eauto. 734 | edestruct STEP0 as (h1'' & hj'' & D'' & E'' & J'' & SAFE''); eauto. 735 | assert (hj' = hj''). 736 | { apply H with (hunion h1' hf) (hunion h1'' hf); auto. 737 | HDISJ. HDISJ. 738 | rewrite ! (hunion_comm hf) by HDISJ. 739 | rewrite <- ! hunion_assoc. 740 | rewrite (hunion_comm h1'), (hunion_comm h1'') by HDISJ. 741 | congruence. 742 | } 743 | subst hj''. 744 | assert (h1' = h1''). 745 | { apply hunion_invert_l with (hunion hj' hf). congruence. HDISJ. HDISJ. } 746 | subst h1''. 747 | exists h1', hj'; auto. 748 | Qed. 749 | 750 | Lemma triple_and: forall J P c Q P' Q', 751 | precise J -> 752 | J ⊢ ⦃ P ⦄ c ⦃ Q ⦄ -> J ⊢ ⦃ P' ⦄ c ⦃ Q' ⦄ -> 753 | J ⊢ ⦃ aand P P' ⦄ c ⦃ fun v => aand (Q v) (Q' v) ⦄. 754 | Proof. 755 | intros until Q'; intros PR TR1 TR2 n h (Ph & P'h). 756 | apply safe_and; auto. 757 | Qed. 758 | 759 | (** * 3. Mutual exclusion *) 760 | 761 | (** ** 3.1. Binary semaphores *) 762 | 763 | (** A binary semaphore is a memory location that contains 0 if it is empty 764 | and 1 if it is busy. *) 765 | 766 | Definition sem_invariant (lck: addr) (R: assertion) : assertion := 767 | aexists (fun v => contains lck v ** (if Z.eqb v 0 then emp else R)). 768 | 769 | (** Acquiring a semaphore (the P operation) is achieved by atomically 770 | setting it to 0 until its previous value was not 0. *) 771 | 772 | Definition SWAP (l: addr) (new_v: Z) : com := 773 | ATOMIC (LET (GET l) (fun old_v => SEQ (SET l new_v) (PURE old_v))). 774 | 775 | Definition ACQUIRE (lck: addr) : com := 776 | REPEAT (SWAP lck 0). 777 | 778 | (** Releasing a semaphore (the V operation) is achieved by atomically 779 | setting it to 1. *) 780 | 781 | Definition RELEASE (lck: addr) : com := 782 | ATOMIC (SET lck 1). 783 | 784 | Lemma triple_swap: 785 | forall lck R, 786 | sem_invariant lck R ⊢ ⦃ emp ⦄ SWAP lck 0 ⦃ fun v => if Z.eqb v 0 then emp else R ⦄. 787 | Proof. 788 | intros. apply triple_atomic. 789 | rewrite sepconj_emp. unfold sem_invariant at 1. 790 | apply triple_exists_pre; intros v. 791 | eapply triple_let with 792 | (Q := fun v' => ((v' = v) //\\ contains lck v) ** (if v =? 0 then emp else R)). 793 | apply triple_frame. apply triple_get. 794 | cbn. intros v'. rewrite lift_pureconj. apply triple_simple_conj_pre. 795 | intros EQ; subst v'. 796 | apply triple_seq with (Q := contains lck 0 ** (if v =? 0 then emp else R)). 797 | apply triple_frame. 798 | apply triple_consequence_pre with (valid lck). 799 | apply triple_set. 800 | red; intros. exists v; auto. 801 | apply triple_pure. 802 | unfold sem_invariant. red; intros. rewrite sepconj_comm, lift_aexists. exists 0. 803 | rewrite Z.eqb_refl. rewrite <- (sepconj_comm emp), sepconj_emp. assumption. 804 | Qed. 805 | 806 | Lemma triple_acquire: 807 | forall lck R, 808 | sem_invariant lck R ⊢ ⦃ emp ⦄ ACQUIRE lck ⦃ fun _ => R ⦄. 809 | Proof. 810 | intros. 811 | apply triple_consequence_post with (Q' := fun v => (v <> 0) //\\ (if Z.eqb v 0 then emp else R)). 812 | apply triple_repeat. apply triple_swap. 813 | rewrite Z.eqb_refl. red; auto. 814 | intros v h [H1 H2]. apply Z.eqb_neq in H1. rewrite H1 in H2. auto. 815 | Qed. 816 | 817 | Lemma triple_release: 818 | forall lck R, 819 | precise R -> 820 | sem_invariant lck R ⊢ ⦃ R ⦄ RELEASE lck ⦃ fun _ => emp ⦄. 821 | Proof. 822 | intros. apply triple_atomic. 823 | rewrite sepconj_comm. unfold sem_invariant at 1. rewrite lift_aexists. 824 | apply triple_exists_pre. intros v. rewrite sepconj_assoc. 825 | apply triple_consequence_post with (Q' := fun _ => contains lck 1 ** (if v =? 0 then emp else R) ** R). 826 | apply triple_frame. 827 | apply triple_consequence_pre with (valid lck). apply triple_set. 828 | red; intros. exists v; auto. 829 | intros _. intros h P. 830 | assert ((contains lck 1 ** R) h). 831 | { eapply sepconj_imp_r; eauto. 832 | destruct (v =? 0). 833 | rewrite sepconj_emp. red; auto. 834 | apply sepconj_self; auto. } 835 | rewrite sepconj_emp. exists 1. simpl. auto. 836 | Qed. 837 | 838 | (** ** 3.2. Critical regions *) 839 | 840 | (** A critical region is a command that is run in mutual exclusion, 841 | while holding the associated lock. *) 842 | 843 | Definition CRITREGION (lck: addr) (c: com) := 844 | SEQ (ACQUIRE lck) (LET c (fun v => SEQ (RELEASE lck) (PURE v))). 845 | 846 | Lemma triple_critregion: 847 | forall lck R c P Q, 848 | precise R -> 849 | emp ⊢ ⦃ P ** R ⦄ c ⦃ fun v => Q v ** R ⦄ -> 850 | sem_invariant lck R ⊢ ⦃ P ⦄ CRITREGION lck c ⦃ Q ⦄. 851 | Proof. 852 | intros. 853 | apply triple_seq with (Q := R ** P). 854 | rewrite <- (sepconj_emp P) at 1. apply triple_frame. apply triple_acquire. 855 | eapply triple_let. 856 | rewrite sepconj_comm. rewrite <- sepconj_emp at 1. apply triple_frame_invariant. apply H0. 857 | intros. simpl. apply triple_seq with (Q := emp ** Q v). 858 | rewrite sepconj_comm. apply triple_frame. apply triple_release; auto. 859 | rewrite sepconj_emp. apply triple_pure. red; auto. 860 | Qed. 861 | 862 | (** ** 3.3. Conditional critical regions *) 863 | 864 | (** A conditional critical region (CCR), as introduced by 865 | Brinch-Hansen and Hoare, is a command [c] that is run in mutual 866 | exclusion but only when a condition [b] is true. *) 867 | 868 | Definition CCR (lck: addr) (b: com) (c: com) := 869 | REPEAT (SEQ (ACQUIRE lck) 870 | (LET b (fun v => IFTHENELSE v (SEQ c (SEQ (RELEASE lck) (PURE 1))) 871 | (SEQ (RELEASE lck) (PURE 0))))). 872 | 873 | Lemma triple_ccr: 874 | forall lck R b c B P Q, 875 | precise R -> 876 | emp ⊢ ⦃ P ** R ⦄ b ⦃ fun v => if v =? 0 then P ** R else B ⦄ -> 877 | emp ⊢ ⦃ B ⦄ c ⦃ fun _ => Q ** R ⦄ -> 878 | sem_invariant lck R ⊢ ⦃ P ⦄ CCR lck b c ⦃ fun _ => Q ⦄. 879 | Proof. 880 | intros. 881 | set (Qloop := fun v => if v =? 0 then P else Q). 882 | apply triple_consequence_post with (fun v => (v <> 0) //\\ Qloop v). 883 | 2: { intros v h (U & V). red in V. apply Z.eqb_neq in U. rewrite U in V. auto. } 884 | apply triple_repeat. 885 | 2: { unfold Qloop. intros v U. simpl in U. auto. } 886 | apply triple_seq with (Q := R ** P). 887 | { rewrite <- (sepconj_emp P) at 1. apply triple_frame. apply triple_acquire. } 888 | rewrite sepconj_comm at 1. 889 | eapply triple_let. rewrite <- sepconj_emp at 1. apply triple_frame_invariant. eexact H0. 890 | intros v. apply triple_ifthenelse. 891 | - (* B succeeded *) 892 | eapply triple_seq. 893 | { eapply triple_consequence_pre. 894 | rewrite <- sepconj_emp at 1. apply triple_frame_invariant. eexact H1. 895 | intros h (X & Y). apply Z.eqb_neq in X. rewrite X in Y. auto. } 896 | apply triple_seq with (Q := emp ** Q). 897 | { rewrite sepconj_comm at 1. apply triple_frame. apply triple_release; auto. } 898 | apply triple_pure. rewrite sepconj_emp. unfold Qloop. cbn. red; auto. 899 | - (* B failed *) 900 | apply triple_consequence_pre with (P ** R). 901 | 2: { intros h (X & Y). subst v. auto. } 902 | eapply triple_seq with (Q := emp ** P). 903 | { rewrite sepconj_comm at 1. apply triple_frame. apply triple_release; auto. } 904 | apply triple_pure. 905 | rewrite sepconj_emp. unfold Qloop. cbn. red; auto. 906 | Qed. 907 | 908 | (** * 4. The producer/consumer problem *) 909 | 910 | (** 4.1. With a one-place buffer and binary semaphores *) 911 | 912 | Module ProdCons1. 913 | 914 | Definition PRODUCE (buff free busy: addr) (data: Z) : com := 915 | SEQ (ACQUIRE free) 916 | (SEQ (SET buff data) 917 | (RELEASE busy)). 918 | 919 | Definition CONSUME (buff free busy: addr) : com := 920 | SEQ (ACQUIRE busy) 921 | (LET (GET buff) (fun data => 922 | (SEQ (RELEASE free) (PURE data)))). 923 | 924 | Definition buffer_invariant (R: Z -> assertion) (buff free busy: addr) := 925 | sem_invariant free (valid buff) 926 | ** sem_invariant busy (aexists (fun v => contains buff v ** R v)). 927 | 928 | Remark precise_buffer_invariant: forall (R: Z -> assertion) buff, 929 | (forall v, precise (R v)) -> 930 | precise (aexists (fun v => contains buff v ** R v)). 931 | Proof. 932 | intros. apply aexists_precise. apply sepconj_param_precise; auto. apply contains_param_precise. 933 | Qed. 934 | 935 | Lemma triple_consume: forall R buff free busy, 936 | buffer_invariant R buff free busy ⊢ 937 | ⦃ emp ⦄ CONSUME buff free busy ⦃ fun v => R v ⦄. 938 | Proof. 939 | intros. 940 | eapply triple_seq. 941 | unfold buffer_invariant. rewrite sepconj_comm. 942 | apply triple_frame_invariant. 943 | apply triple_acquire. 944 | apply triple_exists_pre. intros v. 945 | eapply triple_let. 946 | apply triple_frame. apply triple_get. 947 | intros v'. cbn. rewrite lift_pureconj. apply triple_simple_conj_pre. intros EQ; subst v'. 948 | apply triple_seq with (emp ** R v). 949 | unfold buffer_invariant. apply triple_frame_invariant. apply triple_frame. 950 | eapply triple_consequence_pre. apply triple_release. apply valid_precise. 951 | red; intros; exists v; auto. 952 | apply triple_pure. rewrite sepconj_emp. red; auto. 953 | Qed. 954 | 955 | Lemma triple_produce: forall (R: Z -> assertion) buff free busy data, 956 | (forall v, precise (R v)) -> 957 | buffer_invariant R buff free busy ⊢ 958 | ⦃ R data ⦄ PRODUCE buff free busy data ⦃ fun _ => emp ⦄. 959 | Proof. 960 | intros. 961 | apply triple_seq with (valid buff ** R data). 962 | unfold buffer_invariant. apply triple_frame_invariant. 963 | rewrite <- (sepconj_emp (R data)) at 1. 964 | apply triple_frame. apply triple_acquire. 965 | apply triple_seq with (contains buff data ** R data). 966 | apply triple_frame. apply triple_set. 967 | unfold buffer_invariant. rewrite sepconj_comm. apply triple_frame_invariant. 968 | eapply triple_consequence_pre. 969 | apply triple_release. apply precise_buffer_invariant. assumption. 970 | red; intros. exists data; auto. 971 | Qed. 972 | 973 | End ProdCons1. 974 | 975 | (** ** 4.2. With an unbounded buffer implemented as a list *) 976 | 977 | Module ProdCons2. 978 | 979 | Definition PRODUCE (buff: addr) (data: Z) : com := 980 | LET (ALLOC 2) (fun a => 981 | SEQ (SET a data) 982 | (ATOMIC (LET (GET buff) (fun prev => 983 | SEQ (SET (a + 1) prev) (SET buff a))))). 984 | 985 | Definition POP (buff: addr) : com := 986 | REPEAT (ATOMIC ( 987 | LET (GET buff) (fun b => 988 | IFTHENELSE b 989 | (LET (GET (b + 1)) (fun next => SEQ (SET buff next) (PURE b))) 990 | (PURE 0)))). 991 | 992 | Definition CONSUME (buff: addr) : com := 993 | LET (POP buff) (fun b => 994 | LET (GET b) (fun data => 995 | SEQ (FREE b) (SEQ (FREE (b + 1)) (PURE data)))). 996 | 997 | Fixpoint list_invariant (R: Z -> assertion) (l: list Z) (p: addr) : assertion := 998 | match l with 999 | | nil => (p = 0) //\\ emp 1000 | | x :: l => (p <> 0) //\\ aexists (fun q => contains p x ** contains (p + 1) q ** R x ** list_invariant R l q) 1001 | end. 1002 | 1003 | Definition buffer_invariant (R: Z -> assertion) (buff: addr) : assertion := 1004 | aexists (fun l => aexists (fun p => contains buff p ** list_invariant R l p)). 1005 | 1006 | Lemma triple_produce: forall R buff data, 1007 | buffer_invariant R buff ⊢ 1008 | ⦃ R data ⦄ PRODUCE buff data ⦃ fun _ => emp ⦄. 1009 | Proof. 1010 | intros. eapply triple_let. 1011 | { rewrite <- (sepconj_emp (R data)). apply triple_frame. apply triple_alloc. } 1012 | intros a; cbn. 1013 | rewrite lift_pureconj. apply triple_simple_conj_pre; intros NOT0. 1014 | rewrite ! sepconj_assoc, sepconj_emp. 1015 | apply triple_seq with (contains a data ** valid (a + 1) ** R data). 1016 | { apply triple_frame. apply triple_set. } 1017 | apply triple_atomic. 1018 | rewrite sepconj_comm. unfold buffer_invariant. 1019 | rewrite lift_aexists; apply triple_exists_pre; intros l. 1020 | rewrite lift_aexists; apply triple_exists_pre; intros p. 1021 | rewrite sepconj_assoc. 1022 | eapply triple_let. 1023 | { apply triple_frame. apply triple_get. } 1024 | intros p'; cbn. rewrite lift_pureconj. apply triple_simple_conj_pre; intros EQ; subst p'. 1025 | eapply triple_seq. 1026 | { rewrite (sepconj_pick3 (valid (a + 1))). rewrite sepconj_pick2. 1027 | apply triple_frame with (Q := fun _ => contains (a + 1) p). 1028 | apply triple_set. } 1029 | rewrite sepconj_pick2. eapply triple_consequence_post. 1030 | { apply triple_frame. eapply triple_consequence_pre. apply triple_set. 1031 | intros h A; exists p; auto. } 1032 | cbn. intros _. rewrite sepconj_emp. 1033 | rewrite <- (sepconj_swap3 (list_invariant R l p)). 1034 | rewrite (sepconj_pick2 (contains a data)). 1035 | intros h A. exists (data :: l), a. 1036 | revert h A. apply sepconj_imp_r. 1037 | intros h A. cbn. split; auto. exists p; exact A. 1038 | Qed. 1039 | 1040 | Lemma triple_pop: forall R buff, 1041 | buffer_invariant R buff ⊢ 1042 | ⦃ emp ⦄ POP buff ⦃ fun p => aexists (fun x => contains p x ** valid (p + 1) ** R x) ⦄. 1043 | Proof. 1044 | intros. 1045 | set (Qloop := fun p => if p =? 0 then emp else aexists (fun x => contains p x ** valid (p + 1) ** R x)). 1046 | apply triple_consequence_post with (fun p => (p <> 0) //\\ Qloop p). 1047 | apply triple_repeat. 1048 | - apply triple_atomic. 1049 | rewrite sepconj_emp. 1050 | apply triple_exists_pre; intros l. 1051 | apply triple_exists_pre; intros p. 1052 | eapply triple_let. 1053 | { apply triple_frame. apply triple_get. } 1054 | cbn. intros p'. rewrite lift_pureconj; apply triple_simple_conj_pre; intros E; subst p'. 1055 | apply triple_ifthenelse. 1056 | + apply triple_simple_conj_pre; intros NOTZERO. 1057 | rewrite sepconj_comm. destruct l as [ | x l]; cbn; rewrite lift_pureconj; apply triple_simple_conj_pre; intro; try lia. 1058 | rewrite lift_aexists; apply triple_exists_pre; intros t. 1059 | eapply triple_let. 1060 | { rewrite ! sepconj_assoc, sepconj_pick2. apply triple_frame. apply triple_get. } 1061 | intros t'; cbn. rewrite lift_pureconj; apply triple_simple_conj_pre; intros E; subst t'. 1062 | rewrite <- ! sepconj_assoc, sepconj_comm, ! sepconj_assoc. 1063 | eapply triple_seq. 1064 | { apply triple_frame with (Q := fun _ => contains buff t). 1065 | eapply triple_consequence_pre. apply triple_set. 1066 | intros h A; exists p; auto. } 1067 | apply triple_pure. 1068 | unfold Qloop. apply Z.eqb_neq in NOTZERO; rewrite NOTZERO. 1069 | rewrite (sepconj_pick2 (contains p x)). 1070 | rewrite <- (sepconj_pick3 (contains buff t)). 1071 | rewrite <- (sepconj_pick2 (contains buff t)). 1072 | intros h A. rewrite lift_aexists. exists x. rewrite ! sepconj_assoc. 1073 | eapply sepconj_imp_r; eauto. 1074 | intros h' B. apply sepconj_imp_l with (P := contains (p + 1) t). 1075 | intros h'' C. exists t; auto. 1076 | revert h' B. apply sepconj_imp_r. apply sepconj_imp_r. 1077 | intros h''' D. red. exists l; exists t; auto. 1078 | + apply triple_simple_conj_pre; intros ZERO. 1079 | apply triple_pure. unfold Qloop; cbn. rewrite sepconj_emp. intros h A; exists l, p; auto. 1080 | - unfold Qloop; cbn. red; auto. 1081 | - unfold Qloop. intros v h [A B]. apply Z.eqb_neq in A. rewrite A in B. auto. 1082 | Qed. 1083 | 1084 | Lemma triple_consume: forall R buff, 1085 | buffer_invariant R buff ⊢ 1086 | ⦃ emp ⦄ CONSUME buff ⦃ fun data => R data ⦄. 1087 | Proof. 1088 | intros. eapply triple_let. apply triple_pop. 1089 | intros b. cbn. apply triple_exists_pre; intros p. 1090 | eapply triple_let. 1091 | { apply triple_frame. apply triple_get. } 1092 | intros p'; cbn; rewrite lift_pureconj; apply triple_simple_conj_pre; intros E; subst p'. 1093 | eapply triple_seq. 1094 | { apply triple_frame with (Q := fun _ => emp). 1095 | eapply triple_consequence_pre. apply triple_free. intros h A; exists p; auto. } 1096 | rewrite sepconj_emp. 1097 | eapply triple_seq. 1098 | { apply triple_frame with (Q := fun _ => emp). apply triple_free. } 1099 | apply triple_pure. rewrite sepconj_emp. red; auto. 1100 | Qed. 1101 | 1102 | End ProdCons2. 1103 | 1104 | -------------------------------------------------------------------------------- /CoqMakefile: -------------------------------------------------------------------------------- 1 | ########################################################################## 2 | ## # The Coq Proof Assistant / The Coq Development Team ## 3 | ## v # Copyright INRIA, CNRS and contributors ## 4 | ## /dev/null 2>/dev/null; echo $$?)) 72 | STDTIME?=command time -f $(TIMEFMT) 73 | else 74 | ifeq (0,$(shell gtime -f "" true >/dev/null 2>/dev/null; echo $$?)) 75 | STDTIME?=gtime -f $(TIMEFMT) 76 | else 77 | STDTIME?=command time 78 | endif 79 | endif 80 | else 81 | STDTIME?=command time -f $(TIMEFMT) 82 | endif 83 | 84 | ifneq (,$(COQBIN)) 85 | # add an ending / 86 | COQBIN:=$(COQBIN)/ 87 | endif 88 | 89 | # Coq binaries 90 | COQC ?= "$(COQBIN)coqc" 91 | COQTOP ?= "$(COQBIN)coqtop" 92 | COQCHK ?= "$(COQBIN)coqchk" 93 | COQDEP ?= "$(COQBIN)coqdep" 94 | COQDOC ?= "$(COQBIN)coqdoc" 95 | COQPP ?= "$(COQBIN)coqpp" 96 | COQMKFILE ?= "$(COQBIN)coq_makefile" 97 | OCAMLLIBDEP ?= "$(COQBIN)ocamllibdep" 98 | 99 | # Timing scripts 100 | COQMAKE_ONE_TIME_FILE ?= "$(COQLIB)/tools/make-one-time-file.py" 101 | COQMAKE_BOTH_TIME_FILES ?= "$(COQLIB)/tools/make-both-time-files.py" 102 | COQMAKE_BOTH_SINGLE_TIMING_FILES ?= "$(COQLIB)/tools/make-both-single-timing-files.py" 103 | BEFORE ?= 104 | AFTER ?= 105 | 106 | # FIXME this should be generated by Coq (modules already linked by Coq) 107 | CAMLDONTLINK=num,str,unix,dynlink,threads 108 | 109 | # OCaml binaries 110 | CAMLC ?= "$(OCAMLFIND)" ocamlc -c 111 | CAMLOPTC ?= "$(OCAMLFIND)" opt -c 112 | CAMLLINK ?= "$(OCAMLFIND)" ocamlc -linkpkg -dontlink $(CAMLDONTLINK) 113 | CAMLOPTLINK ?= "$(OCAMLFIND)" opt -linkpkg -dontlink $(CAMLDONTLINK) 114 | CAMLDOC ?= "$(OCAMLFIND)" ocamldoc 115 | CAMLDEP ?= "$(OCAMLFIND)" ocamldep -slash -ml-synonym .mlpack 116 | 117 | # DESTDIR is prepended to all installation paths 118 | DESTDIR ?= 119 | 120 | # Debug builds, typically -g to OCaml, -debug to Coq. 121 | CAMLDEBUG ?= 122 | COQDEBUG ?= 123 | 124 | # Extra packages to be linked in (as in findlib -package) 125 | CAMLPKGS ?= 126 | 127 | # Option for making timing files 128 | TIMING?= 129 | # Option for changing sorting of timing output file 130 | TIMING_SORT_BY ?= auto 131 | # Option for changing the fuzz parameter on the output file 132 | TIMING_FUZZ ?= 0 133 | # Option for changing whether to use real or user time for timing tables 134 | TIMING_REAL?= 135 | # Option for including the memory column(s) 136 | TIMING_INCLUDE_MEM?= 137 | # Option for sorting by the memory column 138 | TIMING_SORT_BY_MEM?= 139 | # Output file names for timed builds 140 | TIME_OF_BUILD_FILE ?= time-of-build.log 141 | TIME_OF_BUILD_BEFORE_FILE ?= time-of-build-before.log 142 | TIME_OF_BUILD_AFTER_FILE ?= time-of-build-after.log 143 | TIME_OF_PRETTY_BUILD_FILE ?= time-of-build-pretty.log 144 | TIME_OF_PRETTY_BOTH_BUILD_FILE ?= time-of-build-both.log 145 | TIME_OF_PRETTY_BUILD_EXTRA_FILES ?= - # also output to the command line 146 | 147 | TGTS ?= 148 | 149 | ########## End of parameters ################################################## 150 | # What follows may be relevant to you only if you need to 151 | # extend this Makefile. If so, look for 'Extension point' here and 152 | # put in CoqMakefile.local double colon rules accordingly. 153 | # E.g. to perform some work after the all target completes you can write 154 | # 155 | # post-all:: 156 | # echo "All done!" 157 | # 158 | # in CoqMakefile.local 159 | # 160 | ############################################################################### 161 | 162 | 163 | 164 | 165 | # Flags ####################################################################### 166 | # 167 | # We define a bunch of variables combining the parameters. 168 | # To add additional flags to coq, coqchk or coqdoc, set the 169 | # {COQ,COQCHK,COQDOC}EXTRAFLAGS variable to whatever you want to add. 170 | # To overwrite the default choice and set your own flags entirely, set the 171 | # {COQ,COQCHK,COQDOC}FLAGS variable. 172 | 173 | SHOW := $(if $(VERBOSE),@true "",@echo "") 174 | HIDE := $(if $(VERBOSE),,@) 175 | 176 | TIMER=$(if $(TIMED), $(STDTIME), $(TIMECMD)) 177 | 178 | OPT?= 179 | 180 | # The DYNOBJ and DYNLIB variables are used by "coqdep -dyndep var" in .v.d 181 | ifeq '$(OPT)' '-byte' 182 | USEBYTE:=true 183 | DYNOBJ:=.cma 184 | DYNLIB:=.cma 185 | else 186 | USEBYTE:= 187 | DYNOBJ:=.cmxs 188 | DYNLIB:=.cmxs 189 | endif 190 | 191 | # these variables are meant to be overridden if you want to add *extra* flags 192 | COQEXTRAFLAGS?= 193 | COQCHKEXTRAFLAGS?= 194 | COQDOCEXTRAFLAGS?= 195 | 196 | # these flags do NOT contain the libraries, to make them easier to overwrite 197 | COQFLAGS?=-q $(OTHERFLAGS) $(COQEXTRAFLAGS) 198 | COQCHKFLAGS?=-silent -o $(COQCHKEXTRAFLAGS) 199 | COQDOCFLAGS?=-interpolate -utf8 $(COQDOCEXTRAFLAGS) 200 | 201 | COQDOCLIBS?=$(COQLIBS_NOML) 202 | 203 | # The version of Coq being run and the version of coq_makefile that 204 | # generated this makefile 205 | COQ_VERSION:=$(shell $(COQC) --print-version | cut -d " " -f 1) 206 | COQMAKEFILE_VERSION:=8.12.2 207 | 208 | COQSRCLIBS?= $(foreach d,$(COQ_SRC_SUBDIRS), -I "$(COQLIB)/$(d)") 209 | 210 | CAMLFLAGS+=$(OCAMLLIBS) $(COQSRCLIBS) 211 | # ocamldoc fails with unknown argument otherwise 212 | CAMLDOCFLAGS:=$(filter-out -annot, $(filter-out -bin-annot, $(CAMLFLAGS))) 213 | CAMLFLAGS+=$(OCAMLWARN) 214 | 215 | ifneq (,$(TIMING)) 216 | TIMING_ARG=-time 217 | ifeq (after,$(TIMING)) 218 | TIMING_EXT=after-timing 219 | else 220 | ifeq (before,$(TIMING)) 221 | TIMING_EXT=before-timing 222 | else 223 | TIMING_EXT=timing 224 | endif 225 | endif 226 | else 227 | TIMING_ARG= 228 | endif 229 | 230 | # Retro compatibility (DESTDIR is standard on Unix, DSTROOT is not) 231 | ifdef DSTROOT 232 | DESTDIR := $(DSTROOT) 233 | endif 234 | 235 | concat_path = $(if $(1),$(1)/$(if $(COQMF_WINDRIVE),$(subst $(COQMF_WINDRIVE),/,$(2)),$(2)),$(2)) 236 | 237 | COQLIBINSTALL = $(call concat_path,$(DESTDIR),$(COQLIB)/user-contrib) 238 | COQDOCINSTALL = $(call concat_path,$(DESTDIR),$(DOCDIR)/user-contrib) 239 | COQTOPINSTALL = $(call concat_path,$(DESTDIR),$(COQLIB)/toploop) 240 | 241 | # Files ####################################################################### 242 | # 243 | # We here define a bunch of variables about the files being part of the 244 | # Coq project in order to ease the writing of build target and build rules 245 | 246 | VDFILE := .CoqMakefile.d 247 | 248 | ALLSRCFILES := \ 249 | $(MLGFILES) \ 250 | $(MLFILES) \ 251 | $(MLPACKFILES) \ 252 | $(MLLIBFILES) \ 253 | $(MLIFILES) 254 | 255 | # helpers 256 | vo_to_obj = $(addsuffix .o,\ 257 | $(filter-out Warning: Error:,\ 258 | $(shell $(COQTOP) -q -noinit -batch -quiet -print-mod-uid $(1)))) 259 | strip_dotslash = $(patsubst ./%,%,$(1)) 260 | 261 | # without this we get undefined variables in the expansion for the 262 | # targets of the [deprecated,use-mllib-or-mlpack] rule 263 | with_undef = $(if $(filter-out undefined, $(origin $(1))),$($(1))) 264 | 265 | VO = vo 266 | VOS = vos 267 | 268 | VOFILES = $(VFILES:.v=.$(VO)) 269 | GLOBFILES = $(VFILES:.v=.glob) 270 | HTMLFILES = $(VFILES:.v=.html) 271 | GHTMLFILES = $(VFILES:.v=.g.html) 272 | BEAUTYFILES = $(addsuffix .beautified,$(VFILES)) 273 | TEXFILES = $(VFILES:.v=.tex) 274 | GTEXFILES = $(VFILES:.v=.g.tex) 275 | CMOFILES = \ 276 | $(MLGFILES:.mlg=.cmo) \ 277 | $(MLFILES:.ml=.cmo) \ 278 | $(MLPACKFILES:.mlpack=.cmo) 279 | CMXFILES = $(CMOFILES:.cmo=.cmx) 280 | OFILES = $(CMXFILES:.cmx=.o) 281 | CMAFILES = $(MLLIBFILES:.mllib=.cma) $(MLPACKFILES:.mlpack=.cma) 282 | CMXAFILES = $(CMAFILES:.cma=.cmxa) 283 | CMIFILES = \ 284 | $(CMOFILES:.cmo=.cmi) \ 285 | $(MLIFILES:.mli=.cmi) 286 | # the /if/ is because old _CoqProject did not list a .ml(pack|lib) but just 287 | # a .mlg file 288 | CMXSFILES = \ 289 | $(MLPACKFILES:.mlpack=.cmxs) \ 290 | $(CMXAFILES:.cmxa=.cmxs) \ 291 | $(if $(MLPACKFILES)$(CMXAFILES),,\ 292 | $(MLGFILES:.mlg=.cmxs) $(MLFILES:.ml=.cmxs)) 293 | 294 | # files that are packed into a plugin (no extension) 295 | PACKEDFILES = \ 296 | $(call strip_dotslash, \ 297 | $(foreach lib, \ 298 | $(call strip_dotslash, \ 299 | $(MLPACKFILES:.mlpack=_MLPACK_DEPENDENCIES)),$(call with_undef,$(lib)))) 300 | # files that are archived into a .cma (mllib) 301 | LIBEDFILES = \ 302 | $(call strip_dotslash, \ 303 | $(foreach lib, \ 304 | $(call strip_dotslash, \ 305 | $(MLLIBFILES:.mllib=_MLLIB_DEPENDENCIES)),$(call with_undef,$(lib)))) 306 | CMIFILESTOINSTALL = $(filter-out $(addsuffix .cmi,$(PACKEDFILES)),$(CMIFILES)) 307 | CMOFILESTOINSTALL = $(filter-out $(addsuffix .cmo,$(PACKEDFILES)),$(CMOFILES)) 308 | OBJFILES = $(call vo_to_obj,$(VOFILES)) 309 | ALLNATIVEFILES = \ 310 | $(OBJFILES:.o=.cmi) \ 311 | $(OBJFILES:.o=.cmx) \ 312 | $(OBJFILES:.o=.cmxs) 313 | # trick: wildcard filters out non-existing files, so that `install` doesn't show 314 | # warnings and `clean` doesn't pass to rm a list of files that is too long for 315 | # the shell. 316 | NATIVEFILES = $(wildcard $(ALLNATIVEFILES)) 317 | FILESTOINSTALL = \ 318 | $(VOFILES) \ 319 | $(VFILES) \ 320 | $(GLOBFILES) \ 321 | $(NATIVEFILES) \ 322 | $(CMIFILESTOINSTALL) 323 | BYTEFILESTOINSTALL = \ 324 | $(CMOFILESTOINSTALL) \ 325 | $(CMAFILES) 326 | ifeq '$(HASNATDYNLINK)' 'true' 327 | DO_NATDYNLINK = yes 328 | FILESTOINSTALL += $(CMXSFILES) $(CMXAFILES) $(CMOFILESTOINSTALL:.cmo=.cmx) 329 | else 330 | DO_NATDYNLINK = 331 | endif 332 | 333 | ALLDFILES = $(addsuffix .d,$(ALLSRCFILES)) $(VDFILE) 334 | 335 | # Compilation targets ######################################################### 336 | 337 | all: 338 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all 339 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all 340 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all 341 | .PHONY: all 342 | 343 | all.timing.diff: 344 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" pre-all 345 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" real-all.timing.diff TIME_OF_PRETTY_BUILD_EXTRA_FILES="" 346 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" post-all 347 | .PHONY: all.timing.diff 348 | 349 | ifeq (0,$(TIMING_REAL)) 350 | TIMING_REAL_ARG := 351 | TIMING_USER_ARG := --user 352 | else 353 | ifeq (1,$(TIMING_REAL)) 354 | TIMING_REAL_ARG := --real 355 | TIMING_USER_ARG := 356 | else 357 | TIMING_REAL_ARG := 358 | TIMING_USER_ARG := 359 | endif 360 | endif 361 | 362 | ifeq (0,$(TIMING_INCLUDE_MEM)) 363 | TIMING_INCLUDE_MEM_ARG := --no-include-mem 364 | else 365 | TIMING_INCLUDE_MEM_ARG := 366 | endif 367 | 368 | ifeq (1,$(TIMING_SORT_BY_MEM)) 369 | TIMING_SORT_BY_MEM_ARG := --sort-by-mem 370 | else 371 | TIMING_SORT_BY_MEM_ARG := 372 | endif 373 | 374 | make-pretty-timed-before:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_BEFORE_FILE) 375 | make-pretty-timed-after:: TIME_OF_BUILD_FILE=$(TIME_OF_BUILD_AFTER_FILE) 376 | make-pretty-timed make-pretty-timed-before make-pretty-timed-after:: 377 | $(HIDE)rm -f pretty-timed-success.ok 378 | $(HIDE)($(MAKE) --no-print-directory -f "$(PARENT)" $(TGTS) TIMED=1 2>&1 && touch pretty-timed-success.ok) | tee -a $(TIME_OF_BUILD_FILE) 379 | $(HIDE)rm pretty-timed-success.ok # must not be -f; must fail if the touch failed 380 | print-pretty-timed:: 381 | $(HIDE)$(COQMAKE_ONE_TIME_FILE) $(TIMING_INCLUDE_MEM_ARG) $(TIMING_SORT_BY_MEM_ARG) $(TIMING_REAL_ARG) $(TIME_OF_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 382 | print-pretty-timed-diff:: 383 | $(HIDE)$(COQMAKE_BOTH_TIME_FILES) --sort-by=$(TIMING_SORT_BY) $(TIMING_INCLUDE_MEM_ARG) $(TIMING_SORT_BY_MEM_ARG) $(TIMING_REAL_ARG) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 384 | ifeq (,$(BEFORE)) 385 | print-pretty-single-time-diff:: 386 | @echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing' 387 | $(HIDE)false 388 | else 389 | ifeq (,$(AFTER)) 390 | print-pretty-single-time-diff:: 391 | @echo 'Error: Usage: $(MAKE) print-pretty-single-time-diff AFTER=path/to/file.v.after-timing BEFORE=path/to/file.v.before-timing' 392 | $(HIDE)false 393 | else 394 | print-pretty-single-time-diff:: 395 | $(HIDE)$(COQMAKE_BOTH_SINGLE_TIMING_FILES) --fuzz=$(TIMING_FUZZ) --sort-by=$(TIMING_SORT_BY) $(TIMING_USER_ARG) $(AFTER) $(BEFORE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BUILD_EXTRA_FILES) 396 | endif 397 | endif 398 | pretty-timed: 399 | $(HIDE)$(MAKE) --no-print-directory -f "$(PARENT)" make-pretty-timed 400 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-timed 401 | .PHONY: pretty-timed make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff 402 | 403 | # Extension points for actions to be performed before/after the all target 404 | pre-all:: 405 | @# Extension point 406 | $(HIDE)if [ "$(COQMAKEFILE_VERSION)" != "$(COQ_VERSION)" ]; then\ 407 | echo "W: This Makefile was generated by Coq $(COQMAKEFILE_VERSION)";\ 408 | echo "W: while the current Coq version is $(COQ_VERSION)";\ 409 | fi 410 | .PHONY: pre-all 411 | 412 | post-all:: 413 | @# Extension point 414 | .PHONY: post-all 415 | 416 | real-all: $(VOFILES) $(if $(USEBYTE),bytefiles,optfiles) 417 | .PHONY: real-all 418 | 419 | real-all.timing.diff: $(VOFILES:.vo=.v.timing.diff) 420 | .PHONY: real-all.timing.diff 421 | 422 | bytefiles: $(CMOFILES) $(CMAFILES) 423 | .PHONY: bytefiles 424 | 425 | optfiles: $(if $(DO_NATDYNLINK),$(CMXSFILES)) 426 | .PHONY: optfiles 427 | 428 | # FIXME, see Ralf's bugreport 429 | # quick is deprecated, now renamed vio 430 | vio: $(VOFILES:.vo=.vio) 431 | .PHONY: vio 432 | quick: vio 433 | $(warning "'make quick' is deprecated, use 'make vio' or consider using 'vos' files") 434 | .PHONY: quick 435 | 436 | vio2vo: 437 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) \ 438 | -schedule-vio2vo $(J) $(VOFILES:%.vo=%.vio) 439 | .PHONY: vio2vo 440 | 441 | # quick2vo is undocumented 442 | quick2vo: 443 | $(HIDE)make -j $(J) vio 444 | $(HIDE)VIOFILES=$$(for vofile in $(VOFILES); do \ 445 | viofile="$$(echo "$$vofile" | sed "s/\.vo$$/.vio/")"; \ 446 | if [ "$$vofile" -ot "$$viofile" -o ! -e "$$vofile" ]; then printf "$$viofile "; fi; \ 447 | done); \ 448 | echo "VIO2VO: $$VIOFILES"; \ 449 | if [ -n "$$VIOFILES" ]; then \ 450 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -schedule-vio2vo $(J) $$VIOFILES; \ 451 | fi 452 | .PHONY: quick2vo 453 | 454 | checkproofs: 455 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) \ 456 | -schedule-vio-checking $(J) $(VOFILES:%.vo=%.vio) 457 | .PHONY: checkproofs 458 | 459 | vos: $(VOFILES:%.vo=%.vos) 460 | .PHONY: vos 461 | 462 | vok: $(VOFILES:%.vo=%.vok) 463 | .PHONY: vok 464 | 465 | validate: $(VOFILES) 466 | $(TIMER) $(COQCHK) $(COQCHKFLAGS) $(COQLIBS) $^ 467 | .PHONY: validate 468 | 469 | only: $(TGTS) 470 | .PHONY: only 471 | 472 | # Documentation targets ####################################################### 473 | 474 | html: $(GLOBFILES) $(VFILES) 475 | $(SHOW)'COQDOC -d html $(GAL)' 476 | $(HIDE)mkdir -p html 477 | $(HIDE)$(COQDOC) \ 478 | -toc $(COQDOCFLAGS) -html $(GAL) $(COQDOCLIBS) -d html $(VFILES) 479 | 480 | mlihtml: $(MLIFILES:.mli=.cmi) 481 | $(SHOW)'CAMLDOC -d $@' 482 | $(HIDE)mkdir $@ || rm -rf $@/* 483 | $(HIDE)$(CAMLDOC) -html \ 484 | -d $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES) 485 | 486 | all-mli.tex: $(MLIFILES:.mli=.cmi) 487 | $(SHOW)'CAMLDOC -latex $@' 488 | $(HIDE)$(CAMLDOC) -latex \ 489 | -o $@ -m A $(CAMLDEBUG) $(CAMLDOCFLAGS) $(MLIFILES) 490 | 491 | all.ps: $(VFILES) 492 | $(SHOW)'COQDOC -ps $(GAL)' 493 | $(HIDE)$(COQDOC) \ 494 | -toc $(COQDOCFLAGS) -ps $(GAL) $(COQDOCLIBS) \ 495 | -o $@ `$(COQDEP) -sort $(VFILES)` 496 | 497 | all.pdf: $(VFILES) 498 | $(SHOW)'COQDOC -pdf $(GAL)' 499 | $(HIDE)$(COQDOC) \ 500 | -toc $(COQDOCFLAGS) -pdf $(GAL) $(COQDOCLIBS) \ 501 | -o $@ `$(COQDEP) -sort $(VFILES)` 502 | 503 | # FIXME: not quite right, since the output name is different 504 | gallinahtml: GAL=-g 505 | gallinahtml: html 506 | 507 | all-gal.ps: GAL=-g 508 | all-gal.ps: all.ps 509 | 510 | all-gal.pdf: GAL=-g 511 | all-gal.pdf: all.pdf 512 | 513 | # ? 514 | beautify: $(BEAUTYFILES) 515 | for file in $^; do mv $${file%.beautified} $${file%beautified}old && mv $${file} $${file%.beautified}; done 516 | @echo 'Do not do "make clean" until you are sure that everything went well!' 517 | @echo 'If there were a problem, execute "for file in $$(find . -name \*.v.old -print); do mv $${file} $${file%.old}; done" in your shell/' 518 | .PHONY: beautify 519 | 520 | # Installation targets ######################################################## 521 | # 522 | # There rules can be extended in CoqMakefile.local 523 | # Extensions can't assume when they run. 524 | 525 | install: 526 | $(HIDE)code=0; for f in $(FILESTOINSTALL); do\ 527 | if ! [ -f "$$f" ]; then >&2 echo $$f does not exist; code=1; fi \ 528 | done; exit $$code 529 | $(HIDE)for f in $(FILESTOINSTALL); do\ 530 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\ 531 | if [ "$$?" != "0" -o -z "$$df" ]; then\ 532 | echo SKIP "$$f" since it has no logical path;\ 533 | else\ 534 | install -d "$(COQLIBINSTALL)/$$df" &&\ 535 | install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\ 536 | echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\ 537 | fi;\ 538 | done 539 | $(HIDE)$(MAKE) install-extra -f "$(SELF)" 540 | install-extra:: 541 | @# Extension point 542 | .PHONY: install install-extra 543 | 544 | install-byte: 545 | $(HIDE)for f in $(BYTEFILESTOINSTALL); do\ 546 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`";\ 547 | if [ "$$?" != "0" -o -z "$$df" ]; then\ 548 | echo SKIP "$$f" since it has no logical path;\ 549 | else\ 550 | install -d "$(COQLIBINSTALL)/$$df" &&\ 551 | install -m 0644 "$$f" "$(COQLIBINSTALL)/$$df" &&\ 552 | echo INSTALL "$$f" "$(COQLIBINSTALL)/$$df";\ 553 | fi;\ 554 | done 555 | 556 | install-doc:: html mlihtml 557 | @# Extension point 558 | $(HIDE)install -d "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" 559 | $(HIDE)for i in html/*; do \ 560 | dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\ 561 | install -m 0644 "$$i" "$$dest";\ 562 | echo INSTALL "$$i" "$$dest";\ 563 | done 564 | $(HIDE)install -d \ 565 | "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml" 566 | $(HIDE)for i in mlihtml/*; do \ 567 | dest="$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/$$i";\ 568 | install -m 0644 "$$i" "$$dest";\ 569 | echo INSTALL "$$i" "$$dest";\ 570 | done 571 | .PHONY: install-doc 572 | 573 | uninstall:: 574 | @# Extension point 575 | $(HIDE)for f in $(FILESTOINSTALL); do \ 576 | df="`$(COQMKFILE) -destination-of "$$f" $(COQLIBS)`" &&\ 577 | instf="$(COQLIBINSTALL)/$$df/`basename $$f`" &&\ 578 | rm -f "$$instf" &&\ 579 | echo RM "$$instf" &&\ 580 | (rmdir "$(call concat_path,,$(COQLIBINSTALL)/$$df/)" 2>/dev/null || true); \ 581 | done 582 | .PHONY: uninstall 583 | 584 | uninstall-doc:: 585 | @# Extension point 586 | $(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html' 587 | $(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/html" 588 | $(SHOW)'RM $(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml' 589 | $(HIDE)rm -rf "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/mlihtml" 590 | $(HIDE) rmdir "$(COQDOCINSTALL)/$(INSTALLCOQDOCROOT)/" || true 591 | .PHONY: uninstall-doc 592 | 593 | # Cleaning #################################################################### 594 | # 595 | # There rules can be extended in CoqMakefile.local 596 | # Extensions can't assume when they run. 597 | 598 | clean:: 599 | @# Extension point 600 | $(SHOW)'CLEAN' 601 | $(HIDE)rm -f $(CMOFILES) 602 | $(HIDE)rm -f $(CMIFILES) 603 | $(HIDE)rm -f $(CMAFILES) 604 | $(HIDE)rm -f $(CMOFILES:.cmo=.cmx) 605 | $(HIDE)rm -f $(CMXAFILES) 606 | $(HIDE)rm -f $(CMXSFILES) 607 | $(HIDE)rm -f $(CMOFILES:.cmo=.o) 608 | $(HIDE)rm -f $(CMXAFILES:.cmxa=.a) 609 | $(HIDE)rm -f $(MLGFILES:.mlg=.ml) 610 | $(HIDE)rm -f $(ALLDFILES) 611 | $(HIDE)rm -f $(NATIVEFILES) 612 | $(HIDE)find . -name .coq-native -type d -empty -delete 613 | $(HIDE)rm -f $(VOFILES) 614 | $(HIDE)rm -f $(VOFILES:.vo=.vio) 615 | $(HIDE)rm -f $(VOFILES:.vo=.vos) 616 | $(HIDE)rm -f $(VOFILES:.vo=.vok) 617 | $(HIDE)rm -f $(BEAUTYFILES) $(VFILES:=.old) 618 | $(HIDE)rm -f all.ps all-gal.ps all.pdf all-gal.pdf all.glob all-mli.tex 619 | $(HIDE)rm -f $(VFILES:.v=.glob) 620 | $(HIDE)rm -f $(VFILES:.v=.tex) 621 | $(HIDE)rm -f $(VFILES:.v=.g.tex) 622 | $(HIDE)rm -f pretty-timed-success.ok 623 | $(HIDE)rm -rf html mlihtml 624 | .PHONY: clean 625 | 626 | cleanall:: clean 627 | @# Extension point 628 | $(SHOW)'CLEAN *.aux *.timing' 629 | $(HIDE)rm -f $(foreach f,$(VFILES:.v=),$(dir $(f)).$(notdir $(f)).aux) 630 | $(HIDE)rm -f $(TIME_OF_BUILD_FILE) $(TIME_OF_BUILD_BEFORE_FILE) $(TIME_OF_BUILD_AFTER_FILE) $(TIME_OF_PRETTY_BUILD_FILE) $(TIME_OF_PRETTY_BOTH_BUILD_FILE) 631 | $(HIDE)rm -f $(VOFILES:.vo=.v.timing) 632 | $(HIDE)rm -f $(VOFILES:.vo=.v.before-timing) 633 | $(HIDE)rm -f $(VOFILES:.vo=.v.after-timing) 634 | $(HIDE)rm -f $(VOFILES:.vo=.v.timing.diff) 635 | $(HIDE)rm -f .lia.cache .nia.cache 636 | .PHONY: cleanall 637 | 638 | archclean:: 639 | @# Extension point 640 | $(SHOW)'CLEAN *.cmx *.o' 641 | $(HIDE)rm -f $(NATIVEFILES) 642 | $(HIDE)rm -f $(CMOFILES:%.cmo=%.cmx) 643 | .PHONY: archclean 644 | 645 | 646 | # Compilation rules ########################################################### 647 | 648 | $(MLIFILES:.mli=.cmi): %.cmi: %.mli 649 | $(SHOW)'CAMLC -c $<' 650 | $(HIDE)$(TIMER) $(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $< 651 | 652 | $(MLGFILES:.mlg=.ml): %.ml: %.mlg 653 | $(SHOW)'COQPP $<' 654 | $(HIDE)$(COQPP) $< 655 | 656 | # Stupid hack around a deficient syntax: we cannot concatenate two expansions 657 | $(filter %.cmo, $(MLFILES:.ml=.cmo) $(MLGFILES:.mlg=.cmo)): %.cmo: %.ml 658 | $(SHOW)'CAMLC -c $<' 659 | $(HIDE)$(TIMER) $(CAMLC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $< 660 | 661 | # Same hack 662 | $(filter %.cmx, $(MLFILES:.ml=.cmx) $(MLGFILES:.mlg=.cmx)): %.cmx: %.ml 663 | $(SHOW)'CAMLOPT -c $(FOR_PACK) $<' 664 | $(HIDE)$(TIMER) $(CAMLOPTC) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) $(FOR_PACK) $< 665 | 666 | 667 | $(MLLIBFILES:.mllib=.cmxs): %.cmxs: %.cmxa 668 | $(SHOW)'CAMLOPT -shared -o $@' 669 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \ 670 | -linkall -shared -o $@ $< 671 | 672 | $(MLLIBFILES:.mllib=.cma): %.cma: | %.mllib 673 | $(SHOW)'CAMLC -a -o $@' 674 | $(HIDE)$(TIMER) $(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^ 675 | 676 | $(MLLIBFILES:.mllib=.cmxa): %.cmxa: | %.mllib 677 | $(SHOW)'CAMLOPT -a -o $@' 678 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -a -o $@ $^ 679 | 680 | 681 | $(MLPACKFILES:.mlpack=.cmxs): %.cmxs: %.cmxa 682 | $(SHOW)'CAMLOPT -shared -o $@' 683 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \ 684 | -shared -linkall -o $@ $< 685 | 686 | $(MLPACKFILES:.mlpack=.cmxa): %.cmxa: %.cmx 687 | $(SHOW)'CAMLOPT -a -o $@' 688 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -a -o $@ $< 689 | 690 | $(MLPACKFILES:.mlpack=.cma): %.cma: %.cmo | %.mlpack 691 | $(SHOW)'CAMLC -a -o $@' 692 | $(HIDE)$(TIMER) $(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) -a -o $@ $^ 693 | 694 | $(MLPACKFILES:.mlpack=.cmo): %.cmo: | %.mlpack 695 | $(SHOW)'CAMLC -pack -o $@' 696 | $(HIDE)$(TIMER) $(CAMLLINK) $(CAMLDEBUG) $(CAMLFLAGS) -pack -o $@ $^ 697 | 698 | $(MLPACKFILES:.mlpack=.cmx): %.cmx: | %.mlpack 699 | $(SHOW)'CAMLOPT -pack -o $@' 700 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) -pack -o $@ $^ 701 | 702 | # This rule is for _CoqProject with no .mllib nor .mlpack 703 | $(filter-out $(MLLIBFILES:.mllib=.cmxs) $(MLPACKFILES:.mlpack=.cmxs) $(addsuffix .cmxs,$(PACKEDFILES)) $(addsuffix .cmxs,$(LIBEDFILES)),$(MLFILES:.ml=.cmxs) $(MLGFILES:.mlg=.cmxs)): %.cmxs: %.cmx 704 | $(SHOW)'[deprecated,use-mllib-or-mlpack] CAMLOPT -shared -o $@' 705 | $(HIDE)$(TIMER) $(CAMLOPTLINK) $(CAMLDEBUG) $(CAMLFLAGS) $(CAMLPKGS) \ 706 | -shared -o $@ $< 707 | 708 | ifneq (,$(TIMING)) 709 | TIMING_EXTRA = > $<.$(TIMING_EXT) 710 | else 711 | TIMING_EXTRA = 712 | endif 713 | 714 | $(VOFILES): %.vo: %.v 715 | $(SHOW)COQC $< 716 | $(HIDE)$(TIMER) $(COQC) $(COQDEBUG) $(TIMING_ARG) $(COQFLAGS) $(COQLIBS) $< $(TIMING_EXTRA) 717 | 718 | # FIXME ?merge with .vo / .vio ? 719 | $(GLOBFILES): %.glob: %.v 720 | $(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 721 | 722 | $(VFILES:.v=.vio): %.vio: %.v 723 | $(SHOW)COQC -vio $< 724 | $(HIDE)$(TIMER) $(COQC) -vio $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 725 | 726 | $(VFILES:.v=.vos): %.vos: %.v 727 | $(SHOW)COQC -vos $< 728 | $(HIDE)$(TIMER) $(COQC) -vos $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 729 | 730 | $(VFILES:.v=.vok): %.vok: %.v 731 | $(SHOW)COQC -vok $< 732 | $(HIDE)$(TIMER) $(COQC) -vok $(COQDEBUG) $(COQFLAGS) $(COQLIBS) $< 733 | 734 | $(addsuffix .timing.diff,$(VFILES)): %.timing.diff : %.before-timing %.after-timing 735 | $(SHOW)PYTHON TIMING-DIFF $*.{before,after}-timing 736 | $(HIDE)$(MAKE) --no-print-directory -f "$(SELF)" print-pretty-single-time-diff BEFORE=$*.before-timing AFTER=$*.after-timing TIME_OF_PRETTY_BUILD_FILE="$@" 737 | 738 | $(BEAUTYFILES): %.v.beautified: %.v 739 | $(SHOW)'BEAUTIFY $<' 740 | $(HIDE)$(TIMER) $(COQC) $(COQDEBUG) $(COQFLAGS) $(COQLIBS) -beautify $< 741 | 742 | $(TEXFILES): %.tex: %.v 743 | $(SHOW)'COQDOC -latex $<' 744 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex $< -o $@ 745 | 746 | $(GTEXFILES): %.g.tex: %.v 747 | $(SHOW)'COQDOC -latex -g $<' 748 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -latex -g $< -o $@ 749 | 750 | $(HTMLFILES): %.html: %.v %.glob 751 | $(SHOW)'COQDOC -html $<' 752 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -html $< -o $@ 753 | 754 | $(GHTMLFILES): %.g.html: %.v %.glob 755 | $(SHOW)'COQDOC -html -g $<' 756 | $(HIDE)$(COQDOC) $(COQDOCFLAGS) -html -g $< -o $@ 757 | 758 | # Dependency files ############################################################ 759 | 760 | ifndef MAKECMDGOALS 761 | -include $(ALLDFILES) 762 | else 763 | ifneq ($(filter-out archclean clean cleanall printenv make-pretty-timed make-pretty-timed-before make-pretty-timed-after print-pretty-timed print-pretty-timed-diff print-pretty-single-time-diff,$(MAKECMDGOALS)),) 764 | -include $(ALLDFILES) 765 | endif 766 | endif 767 | 768 | .SECONDARY: $(ALLDFILES) 769 | 770 | redir_if_ok = > "$@" || ( RV=$$?; rm -f "$@"; exit $$RV ) 771 | 772 | GENMLFILES:=$(MLGFILES:.mlg=.ml) 773 | $(addsuffix .d,$(ALLSRCFILES)): $(GENMLFILES) 774 | 775 | $(addsuffix .d,$(MLIFILES)): %.mli.d: %.mli 776 | $(SHOW)'CAMLDEP $<' 777 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 778 | 779 | $(addsuffix .d,$(MLGFILES)): %.mlg.d: %.ml 780 | $(SHOW)'CAMLDEP $<' 781 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 782 | 783 | $(addsuffix .d,$(MLFILES)): %.ml.d: %.ml 784 | $(SHOW)'CAMLDEP $<' 785 | $(HIDE)$(CAMLDEP) $(OCAMLLIBS) "$<" $(redir_if_ok) 786 | 787 | $(addsuffix .d,$(MLLIBFILES)): %.mllib.d: %.mllib 788 | $(SHOW)'OCAMLLIBDEP $<' 789 | $(HIDE)$(OCAMLLIBDEP) -c $(OCAMLLIBS) "$<" $(redir_if_ok) 790 | 791 | $(addsuffix .d,$(MLPACKFILES)): %.mlpack.d: %.mlpack 792 | $(SHOW)'OCAMLLIBDEP $<' 793 | $(HIDE)$(OCAMLLIBDEP) -c $(OCAMLLIBS) "$<" $(redir_if_ok) 794 | 795 | # If this makefile is created using a _CoqProject we have coqdep get 796 | # options from it. This avoids argument length limits for pathological 797 | # projects. Note that extra options might be on the command line. 798 | VDFILE_FLAGS:=$(if _CoqProject,-f _CoqProject,) $(CMDLINE_COQLIBS) $(CMDLINE_VFILES) 799 | 800 | $(VDFILE): $(VFILES) 801 | $(SHOW)'COQDEP VFILES' 802 | $(HIDE)$(COQDEP) -vos -dyndep var $(VDFILE_FLAGS) $(redir_if_ok) 803 | 804 | # Misc ######################################################################## 805 | 806 | byte: 807 | $(HIDE)$(MAKE) all "OPT:=-byte" -f "$(SELF)" 808 | .PHONY: byte 809 | 810 | opt: 811 | $(HIDE)$(MAKE) all "OPT:=-opt" -f "$(SELF)" 812 | .PHONY: opt 813 | 814 | # This is deprecated. To extend this makefile use 815 | # extension points and CoqMakefile.local 816 | printenv:: 817 | $(warning printenv is deprecated) 818 | $(warning write extensions in CoqMakefile.local or include CoqMakefile.conf) 819 | @echo 'LOCAL = $(LOCAL)' 820 | @echo 'COQLIB = $(COQLIB)' 821 | @echo 'DOCDIR = $(DOCDIR)' 822 | @echo 'OCAMLFIND = $(OCAMLFIND)' 823 | @echo 'HASNATDYNLINK = $(HASNATDYNLINK)' 824 | @echo 'SRC_SUBDIRS = $(SRC_SUBDIRS)' 825 | @echo 'COQ_SRC_SUBDIRS = $(COQ_SRC_SUBDIRS)' 826 | @echo 'OCAMLFIND = $(OCAMLFIND)' 827 | @echo 'PP = $(PP)' 828 | @echo 'COQFLAGS = $(COQFLAGS)' 829 | @echo 'COQLIB = $(COQLIBS)' 830 | @echo 'COQLIBINSTALL = $(COQLIBINSTALL)' 831 | @echo 'COQDOCINSTALL = $(COQDOCINSTALL)' 832 | .PHONY: printenv 833 | 834 | # Generate a .merlin file. If you need to append directives to this 835 | # file you can extend the merlin-hook target in CoqMakefile.local 836 | .merlin: 837 | $(SHOW)'FILL .merlin' 838 | $(HIDE)echo 'FLG $(COQMF_CAMLFLAGS)' > .merlin 839 | $(HIDE)echo 'B $(COQLIB)' >> .merlin 840 | $(HIDE)echo 'S $(COQLIB)' >> .merlin 841 | $(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \ 842 | echo 'B $(COQLIB)$(d)' >> .merlin;) 843 | $(HIDE)$(foreach d,$(COQ_SRC_SUBDIRS), \ 844 | echo 'S $(COQLIB)$(d)' >> .merlin;) 845 | $(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'B $(d)' >> .merlin;) 846 | $(HIDE)$(foreach d,$(SRC_SUBDIRS), echo 'S $(d)' >> .merlin;) 847 | $(HIDE)$(MAKE) merlin-hook -f "$(SELF)" 848 | .PHONY: merlin 849 | 850 | merlin-hook:: 851 | @# Extension point 852 | .PHONY: merlin-hook 853 | 854 | # prints all variables 855 | debug: 856 | $(foreach v,\ 857 | $(sort $(filter-out $(INITIAL_VARS) INITIAL_VARS,\ 858 | $(.VARIABLES))),\ 859 | $(info $(v) = $($(v)))) 860 | .PHONY: debug 861 | 862 | .DEFAULT_GOAL := all 863 | 864 | # Local Variables: 865 | # mode: makefile-gmake 866 | # End: 867 | -------------------------------------------------------------------------------- /CoqMakefile.conf: -------------------------------------------------------------------------------- 1 | # This configuration file was generated by running: 2 | # coq_makefile -f _CoqProject -o CoqMakefile 3 | 4 | 5 | ############################################################################### 6 | # # 7 | # Project files. # 8 | # # 9 | ############################################################################### 10 | 11 | COQMF_VFILES = Sequences.v Hoare.v Separation.v Seplog.v CSL.v Delay.v Monads.v 12 | COQMF_MLIFILES = 13 | COQMF_MLFILES = 14 | COQMF_MLGFILES = 15 | COQMF_MLPACKFILES = 16 | COQMF_MLLIBFILES = 17 | COQMF_CMDLINE_VFILES = 18 | 19 | ############################################################################### 20 | # # 21 | # Path directives (-I, -R, -Q). # 22 | # # 23 | ############################################################################### 24 | 25 | COQMF_OCAMLLIBS = 26 | COQMF_SRC_SUBDIRS = 27 | COQMF_COQLIBS = -R . CDF 28 | COQMF_COQLIBS_NOML = -R . CDF 29 | COQMF_CMDLINE_COQLIBS = 30 | 31 | ############################################################################### 32 | # # 33 | # Coq configuration. # 34 | # # 35 | ############################################################################### 36 | 37 | COQMF_LOCAL=0 38 | COQMF_COQLIB=/home/xleroy/.opam/4.11.1/lib/coq/ 39 | COQMF_DOCDIR=/home/xleroy/.opam/4.11.1/doc/ 40 | COQMF_OCAMLFIND=/home/xleroy/.opam/4.11.1/bin/ocamlfind 41 | COQMF_CAMLFLAGS=-thread -rectypes -w +a-4-9-27-41-42-44-45-48-58-67 -safe-string -strict-sequence 42 | COQMF_WARN=-warn-error +a-3 43 | COQMF_HASNATDYNLINK=true 44 | COQMF_COQ_SRC_SUBDIRS=config lib clib kernel library engine pretyping interp gramlib gramlib/.pack parsing proofs tactics toplevel printing ide stm vernac plugins/btauto plugins/cc plugins/derive plugins/extraction plugins/firstorder plugins/funind plugins/ltac plugins/micromega plugins/nsatz plugins/omega plugins/rtauto plugins/setoid_ring plugins/ssr plugins/ssrmatching plugins/ssrsearch plugins/syntax 45 | COQMF_WINDRIVE= 46 | 47 | ############################################################################### 48 | # # 49 | # Extra variables. # 50 | # # 51 | ############################################################################### 52 | 53 | COQMF_OTHERFLAGS = 54 | COQMF_INSTALLCOQDOCROOT = CDF 55 | -------------------------------------------------------------------------------- /Delay.v: -------------------------------------------------------------------------------- 1 | (** The [delay] type and Capretta's partiality monad *) 2 | 3 | (** The type [delay A] represents computations that produce a result 4 | of type [A] if they terminate, but can also diverge. *) 5 | 6 | CoInductive delay (A: Type) := 7 | | now : A -> delay A 8 | | later : delay A -> delay A. 9 | 10 | Arguments now [A]. 11 | Arguments later [A]. 12 | 13 | Lemma u_delay: 14 | forall {A: Type} (x: delay A), 15 | x = match x with now v => now v | later y => later y end. 16 | Proof. destruct x; auto. Qed. 17 | 18 | Ltac unroll_delay X := rewrite (u_delay X); cbn. 19 | 20 | (** The monad structure. *) 21 | 22 | Definition ret := now. 23 | 24 | CoFixpoint bind {A B: Type} (d: delay A) (f: A -> delay B) := 25 | match d with 26 | | now v => later (f v) 27 | | later d' => later (bind d' f) 28 | end. 29 | 30 | (** [safe Q d] means: if the computation [d] terminates, 31 | its value satisfies predicate [Q]. It's like a postcondition 32 | in a weak Hoare triple. *) 33 | 34 | CoInductive safe {A: Type} (Q: A -> Prop) : delay A -> Prop := 35 | | safe_now: forall a, Q a -> safe Q (now a) 36 | | safe_later: forall d, safe Q d -> safe Q (later d). 37 | 38 | Lemma safe_inv_now: forall {A: Type} (Q: A -> Prop) v, 39 | safe Q (now v) -> Q v. 40 | Proof. 41 | intros. inversion H. auto. 42 | Qed. 43 | 44 | Lemma safe_inv_later: forall {A: Type} (Q: A -> Prop) d, 45 | safe Q (later d) -> safe Q d. 46 | Proof. 47 | intros. inversion H. auto. 48 | Qed. 49 | 50 | Lemma safe_consequence: 51 | forall {A: Type} (Q1 Q2: A -> Prop), 52 | (forall a, Q1 a -> Q2 a) -> 53 | forall (d: delay A), safe Q1 d -> safe Q2 d. 54 | Proof. 55 | intros until Q2; intros IMP. cofix COINDHYP; destruct 1. 56 | - constructor; auto. 57 | - constructor; auto. 58 | Qed. 59 | 60 | Lemma safe_bind: 61 | forall {A B: Type} (Q1: A -> Prop) (Q2: B -> Prop) 62 | (d: delay A) (f: A -> delay B), 63 | safe Q1 d -> (forall v, Q1 v -> safe Q2 (f v)) -> safe Q2 (bind d f). 64 | Proof. 65 | intros until Q2. cofix COINDHYP; intros. 66 | unroll_delay (bind d f). destruct d. 67 | - apply safe_inv_now in H. constructor. apply H0. auto. 68 | - apply safe_inv_later in H. constructor. apply COINDHYP; auto. 69 | Qed. 70 | -------------------------------------------------------------------------------- /Makefile: -------------------------------------------------------------------------------- 1 | include CoqMakefile 2 | 3 | CoqMakefile: 4 | coq_makefile -f _CoqProject -o CoqMakefile 5 | 6 | documentation: 7 | coq2html -base CDF -short-names -no-css -d docs *.glob *.v 8 | -------------------------------------------------------------------------------- /Monads.v: -------------------------------------------------------------------------------- 1 | (** Hoare monads and Dijkstra monads *) 2 | 3 | From Coq Require Import Program ZArith. 4 | From CDF Require Hoare Separation Delay. 5 | 6 | Ltac inv H := inversion H; clear H; subst. 7 | 8 | (** * Hoare monads *) 9 | 10 | (** ** The generic interface *) 11 | 12 | Module Type HOAREMONAD. 13 | Parameter PRE: Type. 14 | Definition POST (A: Type) := A -> PRE. 15 | Parameter M: forall (P: PRE) (A: Type) (Q: POST A), Type. 16 | Parameter ret: 17 | forall {A: Type} (Q: POST A) (v: A), M (Q v) A Q. 18 | Parameter bind: 19 | forall {A B: Type} (P: PRE) (Q: POST A) (R: POST B), 20 | M P A Q -> (forall (v: A), M (Q v) B R) -> M P B R. 21 | Parameter implies: PRE -> PRE -> Prop. 22 | Parameter consequence_pre: 23 | forall {A: Type} (P P': PRE) (Q: POST A), 24 | implies P' P -> M P A Q -> M P' A Q. 25 | Parameter consequence_post: 26 | forall {A: Type} (P: PRE) (Q Q': POST A), 27 | (forall v, implies (Q v) (Q' v))-> M P A Q -> M P A Q'. 28 | End HOAREMONAD. 29 | 30 | (** ** The Hoare monad of pure computations *) 31 | 32 | Module HPure <: HOAREMONAD. 33 | 34 | Definition PRE := Prop. 35 | Definition POST (A: Type) := A -> PRE. 36 | Definition M (P: PRE) (A: Type) (Q: POST A) := P -> { a : A | Q a }. 37 | 38 | Definition ret {A: Type} (Q: POST A) (v: A) : M (Q v) A Q := 39 | fun p => exist _ v p. 40 | 41 | Definition bind {A B: Type} 42 | (P: PRE) (Q: POST A) (R: POST B) 43 | (a: M P A Q) (f: forall (v: A), M (Q v) B R) : M P B R := 44 | fun p => let '(exist _ b r) := a p in f b r. 45 | 46 | Definition implies (P P': PRE) := P -> P'. 47 | 48 | Definition consequence_pre 49 | {A: Type} (P P': PRE) (Q: POST A) (IMP: implies P' P) (m: M P A Q) : M P' A Q := 50 | fun p' => m (IMP p'). 51 | 52 | Program Definition consequence_post 53 | {A: Type} (P: PRE) (Q Q': POST A) (IMP: forall v, implies (Q v) (Q' v)) (m: M P A Q) : M P A Q' := 54 | fun p => m p. 55 | Next Obligation. 56 | destruct (m p); cbn. apply IMP; auto. 57 | Qed. 58 | 59 | End HPure. 60 | 61 | (** ** The Hoare monad of possibly-diverging computations *) 62 | 63 | Module HDiv <: HOAREMONAD. 64 | 65 | Import Delay. 66 | 67 | Definition PRE := Prop. 68 | Definition POST (A: Type) := A -> PRE. 69 | Definition M (P: PRE) (A: Type) (Q: POST A) := P -> { d : delay A | safe Q d }. 70 | 71 | Program Definition ret {A: Type} (Q: POST A) (v: A) : M (Q v) A Q := 72 | fun _ => now v. 73 | Next Obligation. 74 | constructor; auto. 75 | Qed. 76 | 77 | Section BIND. 78 | 79 | Context {A B: Type} 80 | (P: PRE) (Q: POST A) (R: POST B) 81 | (a: M P A Q) (f: forall (v: A), M (Q v) B R). 82 | 83 | CoFixpoint bind_aux (d: delay A) : safe Q d -> delay B := 84 | match d with 85 | | now v => fun SAFE => later (proj1_sig (f v (safe_inv_now _ _ SAFE))) 86 | | later d => fun SAFE => later (bind_aux d (safe_inv_later _ _ SAFE)) 87 | end. 88 | 89 | Lemma safe_bind_aux: 90 | forall (d: delay A) (s: safe Q d), safe R (bind_aux d s). 91 | Proof. 92 | cofix COH; intros. unroll_delay (bind_aux d s); destruct d. 93 | - destruct (f a0 (safe_inv_now Q a0 s)) as (d1 & s1); cbn. constructor; auto. 94 | - constructor. apply COH. 95 | Qed. 96 | 97 | Definition bind : M P B R := 98 | fun p => let '(exist _ d s) := a p in exist _ (bind_aux d s) (safe_bind_aux d s). 99 | 100 | End BIND. 101 | 102 | Definition implies (P P': PRE) := P -> P'. 103 | 104 | Definition consequence_pre 105 | {A: Type} (P P': PRE) (Q: POST A) (IMP: implies P' P) (m: M P A Q) : M P' A Q := 106 | fun p' => m (IMP p'). 107 | 108 | Program Definition consequence_post 109 | {A: Type} (P: PRE) (Q Q': POST A) (IMP: forall v, implies (Q v) (Q' v)) (m: M P A Q) : M P A Q' := 110 | fun p => m p. 111 | Next Obligation. 112 | destruct (m p); cbn. apply safe_consequence with Q; auto. 113 | Qed. 114 | 115 | (** A specific operation of the DIV monad: a [repeat...until] unbounded loop. 116 | [iter f x] is such that 117 | 118 | - [iter f x = y] if [f x] terminates with [inr y] 119 | - [iter f x = iter f x'] if [f x] terminates with [inl x']. 120 | *) 121 | 122 | Section ITER. 123 | 124 | Context {A B: Type} (P: A -> PRE) (Q: POST B). 125 | 126 | Let R : POST (A + B) := fun ab => match ab with inl a => P a | inr b => Q b end. 127 | 128 | Context (f: forall (a: A), M (P a) (A + B) R). 129 | 130 | CoFixpoint iter_aux (d: delay (A + B)) : safe R d -> delay B := 131 | match d with 132 | | now (inl a) => fun SAFE => let '(exist _ d' s') := f a (safe_inv_now _ _ SAFE) in later (iter_aux d' s') 133 | | now (inr b) => fun SAFE => now b 134 | | later d => fun SAFE => later (iter_aux d (safe_inv_later _ _ SAFE)) 135 | end. 136 | 137 | Lemma safe_iter_aux: 138 | forall (d: delay (A + B)) (s: safe R d), safe Q (iter_aux d s). 139 | Proof. 140 | cofix COINDHYP; intros. unroll_delay (iter_aux d s); destruct d as [[a | b] | d]. 141 | - destruct f as [ab' s']. constructor. apply COINDHYP. 142 | - apply safe_inv_now in s. constructor. auto. 143 | - constructor. apply COINDHYP. 144 | Qed. 145 | 146 | Program Definition iter (x: A) : M (P x) B Q := 147 | fun p => iter_aux (now (inl x)) _. 148 | Next Obligation. 149 | constructor. auto. 150 | Qed. 151 | Next Obligation. 152 | apply safe_iter_aux. 153 | Qed. 154 | 155 | End ITER. 156 | 157 | End HDiv. 158 | 159 | (** ** The Hoare monad for mutable state *) 160 | 161 | Module HST <: HOAREMONAD. 162 | 163 | Import Separation. 164 | 165 | Definition PRE := heap -> Prop. 166 | Definition POST (A: Type) := A -> PRE. 167 | 168 | Definition M (P: PRE) (A: Type) (Q: POST A) : Type := 169 | forall h, P h -> { a_h : A * heap | Q (fst a_h) (snd a_h) }. 170 | 171 | Program Definition ret {A: Type} (Q: POST A) (v: A) : M (Q v) A Q 172 | := fun h _ => (v, h). 173 | 174 | Program Definition bind {A B: Type} 175 | (P: PRE) (Q: POST A) (R: POST B) 176 | (a: M P A Q) (f: forall (v: A), M (Q v) B R) : M P B R := 177 | fun h p => let '(v, h') := a h p in f v h' _. 178 | Next Obligation. 179 | destruct (a h p) as [[v1 h1] p1]; simpl in *. inv Heq_anonymous. auto. 180 | Qed. 181 | 182 | Definition implies (P P': PRE) := P -->> P'. 183 | 184 | Definition consequence_pre 185 | {A: Type} (P P': PRE) (Q: POST A) (IMP: implies P' P) (m: M P A Q) : M P' A Q := 186 | fun h p' => m h (IMP h p'). 187 | 188 | Program Definition consequence_post 189 | {A: Type} (P: PRE) (Q Q': POST A) (IMP: forall v, implies (Q v) (Q' v)) (m: M P A Q) : M P A Q' := 190 | fun h p => m h p. 191 | Next Obligation. 192 | destruct (m h p) as [[v h'] q]; cbn. apply IMP; auto. 193 | Qed. 194 | 195 | Definition consequence 196 | {A: Type} (P P': PRE) (Q Q': POST A) 197 | (IMP1: implies P' P) (IMP2: forall v, implies (Q v) (Q' v)) (m: M P A Q) : M P' A Q' := 198 | consequence_pre _ _ _ IMP1 (consequence_post _ _ _ IMP2 m). 199 | 200 | Program Definition get (l: addr) : 201 | forall v R, M (contains l v ** R) Z (fun v' => (v' = v) //\\ contains l v ** R) := 202 | fun v R h p => match h l with Some v' => (v', h) | None => _ end. 203 | Next Obligation. 204 | split; auto. 205 | destruct p as (h1 & h2 & p1 & p2 & D & U). 206 | rewrite U in Heq_anonymous. cbn in Heq_anonymous. 207 | rewrite ! p1, hupdate_same in Heq_anonymous. 208 | congruence. 209 | Qed. 210 | Next Obligation. 211 | exfalso. 212 | destruct p as (h1 & h2 & p1 & p2 & D & U). 213 | rewrite U in Heq_anonymous. cbn in Heq_anonymous. 214 | rewrite ! p1, hupdate_same in Heq_anonymous. 215 | congruence. 216 | Qed. 217 | 218 | Program Definition set (l: addr) (v: Z) : 219 | forall R, M (valid l ** R) unit (fun _ => contains l v ** R) := 220 | fun R h p => (tt, hupdate l v h). 221 | Next Obligation. 222 | destruct p as (h1 & h2 & p1 & p2 & D & U). destruct p1 as (v0 & p1). rewrite p1 in *. 223 | exists (hupdate l v hempty), h2. 224 | split. red; auto. 225 | split. auto. 226 | split. red; intros x; cbn. specialize (D x); cbn in D. 227 | destruct (Z.eq_dec l x); intuition congruence. 228 | rewrite U. apply heap_extensionality; intros x. cbn. destruct (Z.eq_dec l x); auto. 229 | Qed. 230 | 231 | End HST. 232 | 233 | (** ** The Hoare monad for mutable state in separation logic style *) 234 | 235 | Module HSep. 236 | 237 | Import Separation. 238 | 239 | Definition PRE := heap -> Prop. 240 | Definition POST (A: Type) := A -> PRE. 241 | 242 | Definition M (P: PRE) (A: Type) (Q: POST A) : Type := 243 | forall (R: assertion), HST.M (P ** R) A (fun v => Q v ** R). 244 | 245 | Definition ret {A: Type} (Q: POST A) (v: A) : M (Q v) A Q := 246 | fun R => HST.ret (fun v => Q v ** R) v. 247 | 248 | Definition bind {A B: Type} 249 | (P: assertion) (Q: A -> assertion) (R: B -> assertion) 250 | (a: M P A Q) (f: forall (v: A), M (Q v) B R) : M P B R := 251 | fun F => HST.bind (P ** F) (fun v => Q v ** F) (fun v => R v ** F) 252 | (a F) (fun v => f v F). 253 | 254 | Definition implies (P P': PRE) := P -->> P'. 255 | 256 | Program Definition consequence_pre 257 | {A: Type} (P P': PRE) (Q: POST A) (IMP: implies P' P) (m: M P A Q) : M P' A Q := 258 | fun R => HST.consequence_pre (P ** R) (P' ** R) (fun v => Q v ** R) _ (m R). 259 | Next Obligation. 260 | apply sepconj_imp_l; auto. 261 | Qed. 262 | 263 | Program Definition consequence_post 264 | {A: Type} (P: PRE) (Q Q': POST A) (IMP: forall v, implies (Q v) (Q' v)) (m: M P A Q) : M P A Q' := 265 | fun R => HST.consequence_post (P ** R) (fun v => Q v ** R) (fun v => Q' v ** R) _ (m R). 266 | Next Obligation. 267 | apply sepconj_imp_l. apply IMP. 268 | Qed. 269 | 270 | Definition consequence 271 | {A: Type} (P P': PRE) (Q Q': POST A) 272 | (IMP1: implies P' P) (IMP2: forall v, implies (Q v) (Q' v)) (m: M P A Q) : M P' A Q' := 273 | consequence_pre _ _ _ IMP1 (consequence_post _ _ _ IMP2 m). 274 | 275 | Program Definition frame 276 | {A: Type} (R: PRE) (P: PRE) (Q: POST A) (m: M P A Q) : M (P ** R) A (fun v => Q v ** R) := 277 | fun R' => HST.consequence _ _ _ _ _ _ (m (R ** R')). 278 | Next Obligation. 279 | rewrite sepconj_assoc. hnf; auto. 280 | Qed. 281 | Next Obligation. 282 | rewrite sepconj_assoc. hnf; auto. 283 | Qed. 284 | 285 | Program Definition get (l: addr) : 286 | forall v, M (contains l v) Z (fun v' => (v' = v) //\\ contains l v) := 287 | fun v R => HST.consequence_post _ _ _ _ (HST.get l v R). 288 | Next Obligation. 289 | rewrite lift_pureconj. hnf; auto. 290 | Qed. 291 | 292 | Program Definition set (l: addr) (v: Z) : 293 | M (valid l) unit (fun _ => contains l v) := 294 | fun R => HST.set l v R. 295 | 296 | End HSep. 297 | 298 | (** * Dijkstra monads *) 299 | 300 | (** ** The generic interface *) 301 | 302 | Module Type DIJKSTRAMONAD. 303 | Parameter PRE: Type. 304 | Parameter POST: forall (A: Type), Type. 305 | Definition TRANSF (A: Type) := POST A -> PRE. 306 | Parameter M: forall (A: Type) (W: TRANSF A), Type. 307 | Parameter RET: forall {A: Type} (v: A), TRANSF A. 308 | Parameter ret: forall {A: Type} (v: A), M A (RET v). 309 | Parameter BIND: 310 | forall {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B), TRANSF B. 311 | Parameter bind: 312 | forall {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B), 313 | M A W1 -> (forall (v: A), M B (W2 v)) -> M B (BIND W1 W2). 314 | End DIJKSTRAMONAD. 315 | 316 | (** ** The Dijkstra monad of pure computations *) 317 | 318 | Module DPure <: DIJKSTRAMONAD. 319 | 320 | Definition PRE := Prop. 321 | Definition POST (A: Type) := A -> PRE. 322 | Definition TRANSF (A: Type) := POST A -> PRE. 323 | 324 | Definition M (A: Type) (W: TRANSF A) : Type := 325 | forall Q, W Q -> { r: A | Q r}. 326 | 327 | Definition RET {A: Type} (v: A) : TRANSF A := fun Q => Q v. 328 | 329 | Definition ret {A: Type} (v: A) : M A (RET v) := 330 | fun Q p => exist _ v p. 331 | 332 | Definition BIND {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) : TRANSF B := 333 | fun Q => W1 (fun x => W2 x Q). 334 | 335 | Definition bind {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) 336 | (m: M A W1) (f: forall (v: A), M B (W2 v)) : M B (BIND W1 W2) := 337 | fun Q p => let '(exist _ x q) := m _ p in f x Q q. 338 | 339 | End DPure. 340 | 341 | (** ** The Dijkstra monad of diverging computations *) 342 | 343 | Module DDiv <: DIJKSTRAMONAD. 344 | 345 | Definition PRE := Prop. 346 | Definition POST (A: Type) := A -> PRE. 347 | Definition TRANSF (A: Type) := POST A -> PRE. 348 | 349 | Definition M (A: Type) (W: TRANSF A) : Type := forall Q, HDiv.M (W Q) A Q. 350 | 351 | Definition RET {A: Type} (v: A) : TRANSF A := fun Q => Q v. 352 | 353 | Definition ret {A: Type} (v: A) : M A (RET v) := fun Q => HDiv.ret Q v. 354 | 355 | Definition BIND {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) : TRANSF B := 356 | fun Q => W1 (fun x => W2 x Q). 357 | 358 | Definition bind {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) 359 | (m: M A W1) (f: forall (v: A), M B (W2 v)) : M B (BIND W1 W2) := 360 | fun Q => HDiv.bind (BIND W1 W2 Q) (fun x => W2 x Q) Q (m (fun x => W2 x Q)) (fun x => f x Q). 361 | 362 | End DDiv. 363 | 364 | (** Lifting pure computations to the [DDiv] monad. *) 365 | 366 | Definition DIV_of_PURE {A: Type} (W: DPure.TRANSF A) : DDiv.TRANSF A := W. 367 | 368 | Definition div_of_pure {A: Type} {W: DPure.TRANSF A} (m: DPure.M A W) 369 | : DDiv.M A (DIV_of_PURE W) 370 | := fun Q p => let '(exist _ v q) := m Q p in (DDiv.ret v Q q). 371 | 372 | (** ** The Dijkstra monad of mutable state *) 373 | 374 | Module DST <: DIJKSTRAMONAD. 375 | 376 | Import Separation. 377 | 378 | Definition PRE := heap -> Prop. 379 | Definition POST (A: Type) := A -> PRE. 380 | Definition TRANSF (A: Type) := POST A -> PRE. 381 | 382 | Definition M (A: Type) (W: TRANSF A) : Type := forall Q, HST.M (W Q) A Q. 383 | 384 | Definition RET {A: Type} (v: A) : TRANSF A := fun Q => Q v. 385 | 386 | Definition ret {A: Type} (v: A) : M A (RET v) := fun Q => HST.ret Q v. 387 | 388 | Definition BIND {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) : TRANSF B := 389 | fun Q => W1 (fun x => W2 x Q). 390 | 391 | Definition bind {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) 392 | (m: M A W1) (f: forall (v: A), M B (W2 v)) : M B (BIND W1 W2) := 393 | fun Q => HST.bind (BIND W1 W2 Q) (fun x => W2 x Q) Q (m (fun x => W2 x Q)) (fun x => f x Q). 394 | 395 | Definition GET (l: addr) : TRANSF Z := 396 | fun Q (h: heap) => match h l with Some v => Q v h | None => False end. 397 | 398 | Program Definition get (l: addr) : M Z (GET l) := 399 | fun Q h p => match h l with Some v => (v, h) | None => _ end. 400 | Next Obligation. 401 | red in p. rewrite <- Heq_anonymous in p. auto. 402 | Qed. 403 | Next Obligation. 404 | exfalso. red in p. rewrite <- Heq_anonymous in p. auto. 405 | Qed. 406 | 407 | Definition SET (l: addr) (v: Z) : TRANSF unit := 408 | fun Q (h: heap) => h l <> None /\ Q tt (hupdate l v h). 409 | 410 | Program Definition set (l: addr) (v: Z) : M unit (SET l v) := 411 | fun Q h p => (tt, hupdate l v h). 412 | Next Obligation. 413 | apply p. 414 | Qed. 415 | 416 | End DST. 417 | 418 | (** Lifting pure computations to the [DST] monad. *) 419 | 420 | Definition ST_of_PURE {A: Type} (W: DPure.TRANSF A) : DST.TRANSF A := 421 | fun (Q: DST.POST A) h => W (fun a => Q a h). 422 | 423 | Program Definition st_of_pure {A: Type} {W: DPure.TRANSF A} (m: DPure.M A W) 424 | : DST.M A (ST_of_PURE W) 425 | := fun Q h p => (m (fun a => Q a h) p, h). 426 | Next Obligation. 427 | destruct m; cbn. auto. 428 | Qed. 429 | 430 | (** ** The Dijkstra monad of exceptions *) 431 | 432 | Parameter exn: Type. 433 | 434 | Module DExn <: DIJKSTRAMONAD. 435 | 436 | Definition PRE := Prop. 437 | Definition POST (A: Type) := A + exn -> PRE. 438 | Definition TRANSF (A: Type) := POST A -> PRE. 439 | 440 | Definition M (A: Type) (W: TRANSF A) : Type := 441 | forall Q, W Q -> { r: A + exn | Q r}. 442 | 443 | Definition RET {A: Type} (v: A) : TRANSF A := fun Q => Q (inl v). 444 | 445 | Definition ret {A: Type} (v: A) : M A (RET v) := 446 | fun Q p => exist _ (inl v) p. 447 | 448 | Definition BIND {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) : TRANSF B := 449 | fun Q => W1 (fun x => match x with inl v => W2 v Q | inr e => Q (inr e) end). 450 | 451 | Program Definition bind {A B: Type} (W1: TRANSF A) (W2: A -> TRANSF B) 452 | (m: M A W1) (f: forall (v: A), M B (W2 v)) : M B (BIND W1 W2) := 453 | fun Q p => match m _ p with inl v => f v Q _ | inr e => inr e end. 454 | Next Obligation. 455 | red in p. destruct m as [r q]. cbn in Heq_anonymous. subst r. auto. 456 | Qed. 457 | Next Obligation. 458 | red in p. destruct m as [r q]. cbn in Heq_anonymous. subst r. auto. 459 | Qed. 460 | 461 | Definition RAISE (A: Type) (e: exn) : TRANSF A := fun Q => Q (inr e). 462 | 463 | Definition raise (A: Type) (e: exn) : M A (RAISE A e) := 464 | fun Q p => exist _ (inr e) p. 465 | 466 | End DExn. 467 | 468 | (** Lifting pure computations to the [DExn] monad. *) 469 | 470 | Definition EXN_of_PURE {A: Type} (W: DPure.TRANSF A) : DExn.TRANSF A := 471 | fun (Q: DExn.POST A) => W (fun a => Q (inl a)). 472 | 473 | Program Definition exn_of_pure {A: Type} {W: DPure.TRANSF A} (m: DPure.M A W) 474 | : DExn.M A (EXN_of_PURE W) 475 | := fun Q p => inl (proj1_sig (m (fun a => Q (inl a)) p)). 476 | Next Obligation. 477 | destruct m; cbn. auto. 478 | Qed. 479 | 480 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | # Program logics: the companion Coq development 2 | 3 | This repository contains Coq sources for the course ["Program logics"](https://www.college-de-france.fr/site/xavier-leroy/course-2020-2021.htm) given by Xavier Leroy at Collège de France in 2021. 4 | 5 | This is work in progress. 6 | 7 | An HTML pretty-printing of the commented sources is also available: 8 | 9 | 1. Variables and loops: Hoare logics 10 | * Module [Hoare](https://xavierleroy.org/cdf-program-logics/CDF.Hoare.html): Hoare logics for the imperative language IMP. 11 | * Library [Sequences](https://xavierleroy.org/cdf-program-logics/CDF.Sequences.html): definitions and properties of reduction sequences. 12 | 13 | 2. Pointers and data structures: separation logic 14 | * Module [Seplog](https://xavierleroy.org/cdf-program-logics/CDF.Seplog.html): separation logic for the functional/imperative language PTR. 15 | * Library [Separation](https://xavierleroy.org/cdf-program-logics/CDF.Separation.html): definitions and properties of heaps and separation logic assertions. 16 | 17 | 3. Shared-memory concurrency: concurrent separation logic 18 | * Module [CSL](https://xavierleroy.org/cdf-program-logics/CDF.CSL.html): concurrent separation logic for the PTR language + parallel and atomic constructs. 19 | 20 | 4. Hoare monads and Dijkstra monads 21 | * Module [Monads](https://xavierleroy.org/cdf-program-logics/CDF.Monads.html) 22 | * Library [Delay](https://xavierleroy.org/cdf-program-logics/CDF.Delay.html) (coinductive definition of possibly nonterminating computations). 23 | -------------------------------------------------------------------------------- /Separation.v: -------------------------------------------------------------------------------- 1 | (** Heaps and heap assertions for separation logics. *) 2 | 3 | From Coq Require Import ZArith Lia Bool String List. 4 | From Coq Require Import FunctionalExtensionality PropExtensionality. 5 | 6 | Local Open Scope Z_scope. 7 | 8 | (** * 1. Memory heaps *) 9 | 10 | (** A memory heap is a partial function from addresses (memory locations) 11 | to values, with a finite domain. *) 12 | 13 | Definition addr := Z. 14 | 15 | Record heap : Type := { 16 | contents :> addr -> option Z; 17 | isfinite : exists i, forall j, i <= j -> contents j = None 18 | }. 19 | 20 | Lemma heap_extensionality: 21 | forall (h1 h2: heap), 22 | (forall l, h1 l = h2 l) -> h1 = h2. 23 | Proof. 24 | intros. destruct h1 as [c1 fin1], h2 as [c2 fin2]. 25 | assert (c1 = c2) by (apply functional_extensionality; auto). 26 | subst c2. f_equal. apply proof_irrelevance. 27 | Qed. 28 | 29 | (** The empty heap. *) 30 | 31 | Program Definition hempty : heap := 32 | {| contents := fun l => None |}. 33 | Next Obligation. 34 | exists 0; auto. 35 | Qed. 36 | 37 | (** The heap that contains [v] at address [l] and is equal to [h] on 38 | all other addresses. *) 39 | 40 | Program Definition hupdate (l: addr) (v: Z) (h: heap) : heap := 41 | {| contents := fun l' => if Z.eq_dec l l' then Some v else h l' |}. 42 | Next Obligation. 43 | destruct (isfinite h) as (i & fin). 44 | exists (Z.max i (l + 1)); intros. 45 | destruct (Z.eq_dec l j). lia. apply fin; lia. 46 | Qed. 47 | 48 | Lemma hupdate_same: forall l v h, (hupdate l v h) l = Some v. 49 | Proof. 50 | intros; cbn. destruct (Z.eq_dec l l); congruence. 51 | Qed. 52 | 53 | Lemma hupdate_other: forall l v h l', l <> l' -> (hupdate l v h) l' = h l'. 54 | Proof. 55 | intros; cbn. destruct (Z.eq_dec l l'); congruence. 56 | Qed. 57 | 58 | (** The heap [h] after deallocating address [l]. *) 59 | 60 | Program Definition hfree (l: addr) (h: heap) : heap := 61 | {| contents := fun l' => if Z.eq_dec l l' then None else h l' |}. 62 | Next Obligation. 63 | destruct (isfinite h) as (i & fin). 64 | exists i; intros. destruct (Z.eq_dec l j); auto. 65 | Qed. 66 | 67 | (** The heap [h] where addresses [l, ..., l + sz - 1] are initialized to 0. *) 68 | 69 | Fixpoint hinit (l: addr) (sz: nat) (h: heap) : heap := 70 | match sz with O => h | S sz => hupdate l 0 (hinit (l + 1) sz h) end. 71 | 72 | Lemma hinit_inside: 73 | forall h sz l l', l <= l' < l + Z.of_nat sz -> hinit l sz h l' = Some 0. 74 | Proof. 75 | induction sz; intros; cbn. 76 | - lia. 77 | - destruct (Z.eq_dec l l'); auto. apply IHsz. lia. 78 | Qed. 79 | 80 | Lemma hinit_outside: 81 | forall h sz l l', l' < l \/ l + Z.of_nat sz <= l' -> hinit l sz h l' = h l'. 82 | Proof. 83 | induction sz; intros; cbn. 84 | - auto. 85 | - destruct (Z.eq_dec l l'). lia. apply IHsz; lia. 86 | Qed. 87 | 88 | (** The disjoint union of two heaps. *) 89 | 90 | Definition hdisjoint (h1 h2: heap) : Prop := 91 | forall l, h1 l = None \/ h2 l = None. 92 | 93 | Lemma hdisjoint_sym: 94 | forall h1 h2, hdisjoint h1 h2 <-> hdisjoint h2 h1. 95 | Proof. 96 | unfold hdisjoint; intros; split; intros H l; specialize (H l); tauto. 97 | Qed. 98 | 99 | Program Definition hunion (h1 h2: heap) : heap := 100 | {| contents := fun l => if h1 l then h1 l else h2 l |}. 101 | Next Obligation. 102 | destruct (isfinite h1) as (i1 & fin1), (isfinite h2) as (i2 & fin2). 103 | exists (Z.max i1 i2); intros. rewrite fin1, fin2 by lia. auto. 104 | Qed. 105 | 106 | Lemma hunion_comm: 107 | forall h1 h2, hdisjoint h1 h2 -> hunion h2 h1 = hunion h1 h2. 108 | Proof. 109 | intros; apply heap_extensionality; intros; cbn. 110 | specialize (H l). destruct (h1 l), (h2 l); intuition congruence. 111 | Qed. 112 | 113 | Lemma hunion_assoc: 114 | forall h1 h2 h3, hunion (hunion h1 h2) h3 = hunion h1 (hunion h2 h3). 115 | Proof. 116 | intros; apply heap_extensionality; intros; cbn. destruct (h1 l); auto. 117 | Qed. 118 | 119 | Lemma hunion_empty: 120 | forall h, hunion hempty h = h. 121 | Proof. 122 | intros; apply heap_extensionality; intros; cbn. auto. 123 | Qed. 124 | 125 | Lemma hdisjoint_union_l: 126 | forall h1 h2 h3, 127 | hdisjoint (hunion h1 h2) h3 <-> hdisjoint h1 h3 /\ hdisjoint h2 h3. 128 | Proof. 129 | unfold hdisjoint, hunion; cbn; intros. split. 130 | - intros D; split; intros l; destruct (D l); auto; destruct (h1 l); auto. 131 | discriminate. 132 | - intros [D1 D2] l. destruct (D2 l); auto. destruct (h1 l) eqn:H1; auto. destruct (D1 l); auto. congruence. 133 | Qed. 134 | 135 | Lemma hdisjoint_union_r: 136 | forall h1 h2 h3, 137 | hdisjoint h1 (hunion h2 h3) <-> hdisjoint h1 h2 /\ hdisjoint h1 h3. 138 | Proof. 139 | intros. rewrite hdisjoint_sym. rewrite hdisjoint_union_l. 140 | rewrite (hdisjoint_sym h2), (hdisjoint_sym h3). tauto. 141 | Qed. 142 | 143 | (** Disjointness between three heaps. *) 144 | 145 | Definition hdisj3 (h1 h2 h3: heap) := 146 | hdisjoint h1 h2 /\ hdisjoint h1 h3 /\ hdisjoint h2 h3. 147 | 148 | (** A tactic to prove disjointness statements. *) 149 | 150 | Ltac HDISJ := 151 | match goal with 152 | | [ H: hdisj3 _ _ _ |- _ ] => 153 | destruct H as (? & ? & ?); HDISJ 154 | | [ H: hdisjoint (hunion _ _) _ |- _ ] => 155 | apply hdisjoint_union_l in H; destruct H; HDISJ 156 | | [ H: hdisjoint _ (hunion _ _) |- _ ] => 157 | apply hdisjoint_union_r in H; destruct H; HDISJ 158 | | [ |- hdisj3 _ _ _ ] => 159 | split; [|split]; HDISJ 160 | | [ |- hdisjoint (hunion _ _) _ ] => 161 | apply hdisjoint_union_l; split; HDISJ 162 | | [ |- hdisjoint _ (hunion _ _) ] => 163 | apply hdisjoint_union_r; split; HDISJ 164 | | [ |- hdisjoint hempty _ ] => 165 | red; auto 166 | | [ |- hdisjoint _ hempty ] => 167 | red; auto 168 | | [ |- hdisjoint _ _ ] => 169 | assumption || (apply hdisjoint_sym; assumption) || idtac 170 | | _ => idtac 171 | end. 172 | 173 | Lemma hunion_invert_r: 174 | forall h1 h2 h, 175 | hunion h h1 = hunion h h2 -> hdisjoint h h1 -> hdisjoint h h2 -> h1 = h2. 176 | Proof. 177 | intros. apply heap_extensionality; intros l. 178 | assert (hunion h h1 l = hunion h h2 l) by congruence. 179 | cbn in H2. specialize (H0 l); specialize (H1 l). destruct (h l); intuition congruence. 180 | Qed. 181 | 182 | Lemma hunion_invert_l: 183 | forall h1 h2 h, 184 | hunion h1 h = hunion h2 h -> hdisjoint h1 h -> hdisjoint h2 h -> h1 = h2. 185 | Proof. 186 | intros. apply hunion_invert_r with h. 187 | rewrite <- ! (hunion_comm h) by HDISJ. auto. 188 | HDISJ. HDISJ. 189 | Qed. 190 | 191 | (** * 2. Assertions for separation logic *) 192 | 193 | Definition assertion : Type := heap -> Prop. 194 | 195 | (** Implication (entailment). *) 196 | 197 | Definition aimp (P Q: assertion) : Prop := 198 | forall h, P h -> Q h. 199 | 200 | Notation "P -->> Q" := (aimp P Q) (at level 95, no associativity). 201 | 202 | (** Quantification. *) 203 | 204 | Definition aexists {A: Type} (P: A -> assertion) : assertion := 205 | fun h => exists a: A, P a h. 206 | 207 | Definition aforall {A: Type} (P: A -> assertion) : assertion := 208 | fun h => forall a: A, P a h. 209 | 210 | (** The assertion "the heap is empty". *) 211 | 212 | Definition emp : assertion := 213 | fun h => h = hempty. 214 | 215 | (** The pure assertion: "the heap is empty and P holds". *) 216 | 217 | Definition pure (P: Prop) : assertion := 218 | fun h => P /\ h = hempty. 219 | 220 | (** The assertion "address [l] contains value [v]". 221 | The domain of the heap must be the singleton [{l}]. *) 222 | 223 | Definition contains (l: addr) (v: Z) : assertion := 224 | fun h => h = hupdate l v hempty. 225 | 226 | (** The assertion "address [l] is valid" (i.e. in the domain of the heap). *) 227 | 228 | Definition valid (l: addr) : assertion := aexists (contains l). 229 | 230 | (** The separating conjunction. *) 231 | 232 | Definition sepconj (P Q: assertion) : assertion := 233 | fun h => exists h1 h2, P h1 234 | /\ Q h2 235 | /\ hdisjoint h1 h2 /\ h = hunion h1 h2. 236 | 237 | Notation "P ** Q" := (sepconj P Q) (at level 60, right associativity). 238 | 239 | (** The conjunction of a simple assertion and a general assertion. *) 240 | 241 | Definition pureconj (P: Prop) (Q: assertion) : assertion := 242 | fun h => P /\ Q h. 243 | 244 | Notation "P //\\ Q" := (pureconj P Q) (at level 60, right associativity). 245 | 246 | (** Plain conjunction and disjunction. *) 247 | 248 | Definition aand (P Q: assertion) : assertion := 249 | fun h => P h /\ Q h. 250 | Definition aor (P Q: assertion) : assertion := 251 | fun h => P h \/ Q h. 252 | 253 | (** Extensional equality between assertions. *) 254 | 255 | Lemma assertion_extensionality: 256 | forall (P Q: assertion), 257 | (forall h, P h <-> Q h) -> P = Q. 258 | Proof. 259 | intros. apply functional_extensionality; intros h. 260 | apply propositional_extensionality. auto. 261 | Qed. 262 | 263 | (** ** Properties of separating conjunction *) 264 | 265 | Lemma sepconj_comm: forall P Q, P ** Q = Q ** P. 266 | Proof. 267 | assert (forall P Q h, (P ** Q) h -> (Q ** P) h). 268 | { intros P Q h (h1 & h2 & P1 & Q2 & EQ & DISJ); subst h. 269 | exists h2, h1; intuition auto. 270 | apply hdisjoint_sym; auto. 271 | symmetry; apply hunion_comm; auto. } 272 | intros; apply assertion_extensionality; intros; split; auto. 273 | Qed. 274 | 275 | Lemma sepconj_assoc: forall P Q R, (P ** Q) ** R = P ** (Q ** R). 276 | Proof. 277 | intros; apply assertion_extensionality; intros; split. 278 | - intros (hx & h3 & (h1 & h2 & P1 & Q2 & EQ & DISJ) & R3 & EQ' & DISJ'). subst hx h. 279 | rewrite hdisjoint_union_l in EQ'. rewrite hunion_assoc. 280 | exists h1, (hunion h2 h3); intuition auto. 281 | exists h2, h3; intuition auto. 282 | rewrite hdisjoint_union_r; tauto. 283 | - intros (h1 & hx & P1 & (h2 & h3 & Q2 & R3 & EQ & DISJ) & EQ' & DISJ'). subst hx h. 284 | rewrite hdisjoint_union_r in EQ'. rewrite <- hunion_assoc. 285 | exists (hunion h1 h2), h3; intuition auto. 286 | exists h1, h2; intuition auto. 287 | rewrite hdisjoint_union_l; tauto. 288 | Qed. 289 | 290 | Lemma sepconj_emp: forall P, emp ** P = P. 291 | Proof. 292 | intros; apply assertion_extensionality; intros; split. 293 | - intros (h1 & h2 & EMP1 & P2 & EQ & DISJ). red in EMP1. subst h h1. 294 | rewrite hunion_empty; auto. 295 | - intros. exists hempty, h; intuition auto. 296 | red; auto. 297 | red; auto. 298 | rewrite hunion_empty; auto. 299 | Qed. 300 | 301 | Lemma sepconj_imp_l: forall P Q R, 302 | (P -->> Q) -> (P ** R -->> Q ** R). 303 | Proof. 304 | intros P Q R IMP h (h1 & h2 & P1 & Q2 & D & U). 305 | exists h1, h2; intuition auto. 306 | Qed. 307 | 308 | Lemma sepconj_imp_r: forall P Q R, 309 | (P -->> Q) -> (R ** P -->> R ** Q). 310 | Proof. 311 | intros. rewrite ! (sepconj_comm R). apply sepconj_imp_l; auto. 312 | Qed. 313 | 314 | (** ** Other useful logical properties *) 315 | 316 | Lemma pure_sep: forall P Q, 317 | pure (P /\ Q) = pure P ** pure Q. 318 | Proof. 319 | intros. apply assertion_extensionality; intros. 320 | unfold pure, sepconj. split. 321 | - intros ((A & B) & C). subst h. exists hempty, hempty; intuition auto. 322 | intro; auto. 323 | rewrite hunion_empty; auto. 324 | - intros (h1 & h2 & (PP & E1) & (QQ & E2) & C & D). 325 | subst. rewrite hunion_empty; auto. 326 | Qed. 327 | 328 | Lemma pureconj_sepconj: forall P Q, 329 | pure P ** Q = P //\\ Q. 330 | Proof. 331 | intros. apply assertion_extensionality; intros. 332 | unfold pure, sepconj, pureconj; split. 333 | - intros (h1 & h2 & (A & B) & C & D & E). 334 | split. auto. subst. rewrite hunion_empty. auto. 335 | - intros (A & B). exists hempty, h; intuition auto. 336 | intro l. auto. 337 | rewrite hunion_empty; auto. 338 | Qed. 339 | 340 | Lemma lift_pureconj: forall P Q R, (P //\\ Q) ** R = P //\\ (Q ** R). 341 | Proof. 342 | intros. rewrite <- ! pureconj_sepconj. rewrite sepconj_assoc. auto. 343 | Qed. 344 | 345 | Lemma lift_aexists: forall (A: Type) (P: A -> assertion) Q, 346 | aexists P ** Q = aexists (fun x => P x ** Q). 347 | Proof. 348 | intros; apply assertion_extensionality; intros; split. 349 | - intros (h1 & h2 & (a & P1) & Q2 & DISJ & EQ). 350 | exists a, h1, h2; auto. 351 | - intros (a & h1 & h2 & P1 & Q2 & DISJ & EQ). 352 | exists h1, h2; intuition auto. exists a; auto. 353 | Qed. 354 | 355 | Lemma sepconj_swap3: forall R P Q, P ** Q ** R = R ** P ** Q. 356 | Proof. 357 | intros. rewrite <- sepconj_assoc, sepconj_comm. auto. 358 | Qed. 359 | 360 | Lemma sepconj_swap4: forall S P Q R, P ** Q ** R ** S = S ** P ** Q ** R. 361 | Proof. 362 | intros. rewrite <- sepconj_assoc, sepconj_swap3, sepconj_assoc. auto. 363 | Qed. 364 | 365 | Lemma sepconj_pick2: forall Q P R, P ** Q ** R = Q ** P ** R. 366 | Proof. 367 | intros. rewrite (sepconj_comm Q), <- sepconj_assoc, sepconj_comm. auto. 368 | Qed. 369 | 370 | Lemma sepconj_pick3: forall R P Q S, P ** Q ** R ** S = R ** P ** Q ** S. 371 | Proof. 372 | intros. rewrite (sepconj_pick2 R), (sepconj_pick2 P). auto. 373 | Qed. 374 | 375 | (** ** Magic wand *) 376 | 377 | Definition wand (P Q: assertion) : assertion := 378 | fun h => forall h', hdisjoint h h' -> P h' -> Q (hunion h h'). 379 | 380 | Notation "P --* Q" := (wand P Q) (at level 70, right associativity). 381 | 382 | Lemma wand_intro: forall P Q R, 383 | P ** Q -->> R -> P -->> Q --* R. 384 | Proof. 385 | intros P Q R IMP h Ph h' DISJ Qh'. 386 | apply IMP. exists h, h'; auto. 387 | Qed. 388 | 389 | Lemma wand_cancel: forall P Q, 390 | P ** (P --* Q) -->> Q. 391 | Proof. 392 | intros P Q h (h1 & h2 & Ph1 & Wh2 & D & U). subst h. 393 | assert (D': hdisjoint h2 h1) by (apply hdisjoint_sym; auto). 394 | rewrite hunion_comm by auto. apply Wh2; auto. 395 | Qed. 396 | 397 | Lemma wand_charact: forall P Q, 398 | (P --*Q) = (aexists (fun R => (P ** R -->> Q) //\\ R)). 399 | Proof. 400 | intros P Q; apply assertion_extensionality; intros h; split. 401 | - intros W. exists (P --* Q). split; auto. apply wand_cancel. 402 | - intros (R & A & B) h' D Ph'. 403 | assert (D': hdisjoint h' h) by (apply hdisjoint_sym; auto). 404 | rewrite hunion_comm by auto. apply A. exists h', h; auto. 405 | Qed. 406 | 407 | Lemma wand_equiv: forall P Q R, 408 | (P -->> (Q --* R)) <-> (P ** Q -->> R). 409 | Proof. 410 | intros; split; intros H. 411 | - intros h (h1 & h2 & Ph1 & Wh2 & D & U). subst h. 412 | apply H; auto. 413 | - apply wand_intro; auto. 414 | Qed. 415 | 416 | Lemma wand_imp_l: forall P P' Q, 417 | (P' -->> P) -> (P --* Q -->> P' --* Q). 418 | Proof. 419 | intros. intros h W h' DISJ P'h'. apply W; auto. 420 | Qed. 421 | 422 | Lemma wand_imp_r: forall P Q Q', 423 | (Q -->> Q') -> (P --* Q -->> P --* Q'). 424 | Proof. 425 | intros. intros h W h' DISJ Ph'. apply H; apply W; auto. 426 | Qed. 427 | 428 | Lemma wand_idem: forall P, 429 | emp -->> P --* P. 430 | Proof. 431 | intros P h E. rewrite E. red. intros. rewrite hunion_empty. auto. 432 | Qed. 433 | 434 | Lemma wand_pure_l: forall (P: Prop) Q, 435 | P -> (pure P --* Q) = Q. 436 | Proof. 437 | intros P Q PP. apply assertion_extensionality. intros h; split. 438 | - intros W. rewrite <- hunion_empty, hunion_comm by HDISJ. apply W. HDISJ. split; auto. 439 | - intros Qh h' DISJ (Ph' & E). subst h'. rewrite hunion_comm, hunion_empty by HDISJ. auto. 440 | Qed. 441 | 442 | Lemma wand_curry: forall P Q R, 443 | (P ** Q --* R) = (P --* Q --* R). 444 | Proof. 445 | intros; apply assertion_extensionality; intros h; split. 446 | - intros W h1 D1 Ph1 h2 D2 Qh2. rewrite hunion_assoc. apply W. HDISJ. exists h1, h2; intuition auto. HDISJ. 447 | - intros W h' D (h1 & h2 & Ph1 & Qh2 & D12 & U12). subst h'. 448 | rewrite <- hunion_assoc. apply W. HDISJ. auto. HDISJ. auto. 449 | Qed. 450 | 451 | Lemma wand_star: forall P Q R, 452 | ((P --* Q) ** R ) -->> (P --* (Q ** R)). 453 | Proof. 454 | intros; intros h (h1 & h2 & W1 & R2 & D & U). subst h. intros h' D' Ph'. 455 | exists (hunion h1 h'), h2; intuition auto. 456 | apply W1; auto. HDISJ. 457 | HDISJ. 458 | rewrite ! hunion_assoc. f_equal. apply hunion_comm. HDISJ. 459 | Qed. 460 | 461 | (** ** Precise assertions *) 462 | 463 | (** An assertion is precise if "it unambiguously carves out an area of the heap" 464 | (in the words of Gotsman, Berdine, Cook, 2011). *) 465 | 466 | Definition precise (P: assertion) : Prop := 467 | forall h1 h2 h1' h2', 468 | hdisjoint h1 h2 -> hdisjoint h1' h2' -> hunion h1 h2 = hunion h1' h2' -> 469 | P h1 -> P h1' -> h1 = h1'. 470 | 471 | (** A parameterized assertion is precise if, in addition, the parameter 472 | is uniquely determined as well. *) 473 | 474 | Definition param_precise {X: Type} (P: X -> assertion) : Prop := 475 | forall x1 x2 h1 h2 h1' h2', 476 | hdisjoint h1 h2 -> hdisjoint h1' h2' -> hunion h1 h2 = hunion h1' h2' -> 477 | P x1 h1 -> P x2 h1' -> x1 = x2 /\ h1 = h1'. 478 | 479 | Remark param_precise_precise: 480 | forall (X: Type) (P: X -> assertion), 481 | param_precise P -> forall x, precise (P x). 482 | Proof. 483 | intros; red; intros. edestruct (H x x h1 h2 h1' h2'); eauto. 484 | Qed. 485 | 486 | Remark precise_param_precise: 487 | forall P, precise P -> param_precise (fun _ : unit => P). 488 | Proof. 489 | intros; red; intros. split. destruct x1, x2; auto. eauto. 490 | Qed. 491 | 492 | Lemma pure_precise: forall P, 493 | precise (pure P). 494 | Proof. 495 | unfold pure; intros; red; intros. destruct H2, H3. congruence. 496 | Qed. 497 | 498 | Lemma pure_param_precise: forall (X: Type) (P: X -> Prop), 499 | (forall x1 x2, P x1 -> P x2 -> x1 = x2) -> 500 | param_precise (fun x => pure (P x)). 501 | Proof. 502 | unfold pure; intros; red; intros. destruct H3, H4. split. auto. congruence. 503 | Qed. 504 | 505 | Lemma contains_param_precise: forall l, 506 | param_precise (fun v => contains l v). 507 | Proof. 508 | unfold contains; intros; red; intros. 509 | assert (E: hunion h1 h2 l = hunion h1' h2' l) by congruence. 510 | cbn in E. subst h1 h1'. rewrite ! hupdate_same in E. 511 | replace x2 with x1 by congruence. auto. 512 | Qed. 513 | 514 | Lemma contains_precise: forall l v, 515 | precise (contains l v). 516 | Proof. 517 | intros. apply param_precise_precise. apply contains_param_precise. 518 | Qed. 519 | 520 | Lemma aexists_precise: forall (X: Type) (P: X -> assertion), 521 | param_precise P -> precise (aexists P). 522 | Proof. 523 | intros; red; intros. destruct H3 as (x1 & P1), H4 as (x2 & P2). 524 | eapply H; eauto. 525 | Qed. 526 | 527 | Lemma valid_precise: forall l, 528 | precise (valid l). 529 | Proof. 530 | intros. apply aexists_precise. apply contains_param_precise. 531 | Qed. 532 | 533 | Lemma sepconj_param_precise: forall (X: Type) (P Q: X -> assertion), 534 | param_precise P -> (forall x, precise (Q x)) -> 535 | param_precise (fun x => P x ** Q x). 536 | Proof. 537 | intros; red; intros. 538 | destruct H4 as (h3 & h4 & P3 & Q4 & D & E). 539 | destruct H5 as (h3' & h4' & P3' & Q4' & D' & E'). 540 | subst h1 h1'. 541 | assert (x1 = x2 /\ h3 = h3'). 542 | { apply H with (hunion h4 h2) (hunion h4' h2'); auto. HDISJ. HDISJ. 543 | rewrite <- ! hunion_assoc. auto. } 544 | destruct H4. subst x2. 545 | assert (h4 = h4'). 546 | { apply (H0 x1) with (hunion h3 h2) (hunion h3' h2'); auto. HDISJ. HDISJ. 547 | rewrite <- ! hunion_assoc. 548 | rewrite (hunion_comm h3) by HDISJ. 549 | rewrite (hunion_comm h3') by HDISJ. 550 | auto. 551 | } 552 | subst; auto. 553 | Qed. 554 | 555 | Lemma sepconj_precise: forall P Q, 556 | precise P -> precise Q -> precise (P ** Q). 557 | Proof. 558 | intros. 559 | assert (param_precise (fun _ : unit => P ** Q)). 560 | { apply sepconj_param_precise. apply precise_param_precise; auto. auto. } 561 | apply param_precise_precise in H1. auto. exact tt. 562 | Qed. 563 | 564 | (** Distributivity laws for precise assertions. *) 565 | 566 | Lemma sepconj_and_distr_1: forall P1 P2 Q, 567 | aand P1 P2 ** Q -->> aand (P1 ** Q) (P2 ** Q). 568 | Proof. 569 | intros P1 P2 Q h (h1 & h2 & (P11 & P21) & Q2 & D & E); split; exists h1, h2; auto. 570 | Qed. 571 | 572 | Lemma sepconj_and_distr_2: forall P1 P2 Q, 573 | precise Q -> 574 | aand (P1 ** Q) (P2 ** Q) -->> aand P1 P2 ** Q. 575 | Proof. 576 | intros P1 P2 Q PQ. 577 | rewrite (sepconj_comm P1), (sepconj_comm P2). 578 | intros h ((h1 & h2 & Q1 & P12 & D & E) & (h1' & h2' & Q1' & P22 & D' & E')). 579 | assert (h1 = h1'). 580 | { apply PQ with h2 h2'; auto. HDISJ. HDISJ. congruence. } 581 | subst h1'. 582 | assert (h2 = h2'). 583 | { apply hunion_invert_r with h1; auto; congruence. } 584 | subst h2'. 585 | unfold aand; exists h2, h1; intuition auto. HDISJ. rewrite hunion_comm by HDISJ; auto. 586 | Qed. 587 | 588 | (** Self-conjunction law for precise assertions. *) 589 | 590 | Lemma sepconj_self: forall P, 591 | precise P -> 592 | P ** P -->> P. 593 | Proof. 594 | intros. intros h (h1 & h2 & P1 & P2 & D & E). 595 | assert (h1 = h2). { apply H with h2 h1; auto. HDISJ. apply hunion_comm; HDISJ. } 596 | subst h2. 597 | assert (h = h1). { apply heap_extensionality; intros l. rewrite E; cbn. destruct (h1 l); auto. } 598 | congruence. 599 | Qed. 600 | -------------------------------------------------------------------------------- /Seplog.v: -------------------------------------------------------------------------------- 1 | (** Separation logic. *) 2 | 3 | From Coq Require Import ZArith Lia Bool List Program.Equality. 4 | From Coq Require Import FunctionalExtensionality PropExtensionality. 5 | From CDF Require Import Sequences Separation. 6 | 7 | Local Open Scope Z_scope. 8 | 9 | (** * 1. A language of pointers *) 10 | 11 | (** We now define a small programming language to work with pointers to 12 | mutable state. The language has variables, but these variables are 13 | immutable. This in unlike IMP but like ML: mutable variables are 14 | expressed as immutable pointers (references) to mutable state. *) 15 | 16 | (** As in ML too, we blur the distinction between expressions and commands. 17 | Every command returns a value, which we take to be an integer, 18 | in addition to possibly performing effects. *) 19 | 20 | (** We use higher-order abstract syntax to represent commands in this 21 | language. With first-order abstract syntax, a "let" binding 22 | [let x = a in b] would be represented using the constructor 23 | << 24 | LET: forall (x: ident) (a b: com), com 25 | >> 26 | With higher-order syntax, we use a Coq function [fun x => b] to 27 | represent the binding of [x] inside [b]: 28 | << 29 | LET: forall (a: com) (b: Z -> com), com 30 | >> 31 | As a benefit, we can use any Coq expression of type [Z] as a 32 | pure command of the language, making it unnecessary to define 33 | syntax and semantics for a specific expression language. 34 | *) 35 | 36 | CoInductive com: Type := 37 | | PURE (x: Z) (**r command without effects *) 38 | | LET (c: com) (f: Z -> com) (**r sequencing of commands *) 39 | | IFTHENELSE (b: Z) (c1 c2: com) (**r conditional *) 40 | | ALLOC (sz: nat) (**r allocate [sz] words of storage *) 41 | | GET (l: addr) (**r dereference a pointer *) 42 | | SET (l: addr) (v: Z) (**r assign through a pointer *) 43 | | FREE (l: addr) (**r free one word of storage *) 44 | | PICK (n: Z). (**r pick a number between 0 and [n] *) 45 | 46 | (** Some derived forms. *) 47 | 48 | Definition SKIP: com := PURE 0. 49 | 50 | Definition SEQ (c1 c2: com) := LET c1 (fun _ => c2). 51 | 52 | Definition EITHER (c1 c2: com) := LET (PICK 2) (fun n => IFTHENELSE n c1 c2). 53 | 54 | (** Reduction semantics. *) 55 | 56 | Inductive red: com * heap -> com * heap -> Prop := 57 | | red_pick: forall n i h, 58 | 0 <= i < n -> 59 | red (PICK n, h) (PURE i, h) 60 | | red_let_done: forall x f h, 61 | red (LET (PURE x) f, h) (f x, h) 62 | | red_let_step: forall c f h c' h', 63 | red (c, h) (c', h') -> 64 | red (LET c f, h) (LET c' f, h') 65 | | red_ifthenelse: forall b c1 c2 h, 66 | red (IFTHENELSE b c1 c2, h) ((if b =? 0 then c2 else c1), h) 67 | | red_alloc: forall sz (h: heap) l, 68 | (forall i, l <= i < l + Z.of_nat sz -> h i = None) -> 69 | l <> 0 -> 70 | red (ALLOC sz, h) (PURE l, hinit l sz h) 71 | | red_get: forall l (h: heap) v, 72 | h l = Some v -> 73 | red (GET l, h) (PURE v, h) 74 | | red_set: forall l v (h: heap), 75 | h l <> None -> 76 | red (SET l v, h) (SKIP, hupdate l v h) 77 | | red_free: forall l (h: heap), 78 | h l <> None -> 79 | red (FREE l, h) (SKIP, hfree l h). 80 | 81 | (** Absence of run-time errors. [immsafe c h] holds if [c / h] is not 82 | going to abort immediately on a run-time error, such as dereferencing 83 | an invalid pointer. *) 84 | 85 | Inductive immsafe: com * heap -> Prop := 86 | | immsafe_pure: forall v h, 87 | immsafe (PURE v, h) 88 | | immsafe_let: forall c f h, 89 | immsafe (c, h) -> immsafe (LET c f, h) 90 | | immsafe_ifthenelse: forall b c1 c2 h, 91 | immsafe (IFTHENELSE b c1 c2, h) 92 | | immsafe_alloc: forall sz (h: heap) l, 93 | l <> 0 -> (forall i, l <= i < l + Z.of_nat sz -> h i = None) -> 94 | immsafe (ALLOC sz, h) 95 | | immsafe_get: forall l (h: heap), 96 | h l <> None -> immsafe (GET l, h) 97 | | immsafe_set: forall l v (h: heap), 98 | h l <> None -> immsafe (SET l v, h) 99 | | immsafe_free: forall l (h: heap), 100 | h l <> None -> immsafe (FREE l, h) 101 | | immsafe_pick: forall n h, 102 | immsafe (PICK n, h). 103 | 104 | (** * 2. The rules of separation logic *) 105 | 106 | Definition precond := assertion. 107 | Definition postcond := Z -> assertion. 108 | 109 | (** ** 2.1. Semantic definition of strong triples *) 110 | 111 | (** Instead of axiomatizing the rules of separation logic, then prove 112 | their soundness against the operational semantics, we define 113 | triples [ ⦃ P ⦄ c ⦃ Q ⦄ ] directly in terms of the 114 | operational semantics, then show the rules of separation logic as 115 | lemmas about these semantic triples. 116 | 117 | Note: the way triples are defined below, they are strong triples 118 | that guarantee termination. However, we write them with braces 119 | instead of brackets, for consistency with the third lecture 120 | and with the literature on separation logic. 121 | *) 122 | 123 | (** [safe c h Q] holds if [c] started in [h] always terminates without errors, 124 | and when it terminates with value [v], the postcondition [Q v] holds 125 | of the final heap. *) 126 | 127 | Inductive safe: com -> heap -> postcond -> Prop := 128 | | safe_done: forall v h (Q: postcond), 129 | Q v h -> 130 | safe (PURE v) h Q 131 | | safe_step: forall c h Q, 132 | match c with PURE _ => False | _ => True end -> 133 | immsafe (c, h) -> 134 | (forall c' h', red (c, h) (c', h') -> safe c' h' Q) -> 135 | safe c h Q. 136 | 137 | (** We define semantic triples like we did for Hoare logic: *) 138 | 139 | Definition triple (P: precond) (c: com) (Q: postcond) := 140 | forall h, P h -> safe c h Q. 141 | 142 | Notation "⦃ P ⦄ c ⦃ Q ⦄" := (triple P c Q) (at level 90, c at next level). 143 | 144 | (** ** 2.2. The frame rule *) 145 | 146 | (** The frame rule is valid because the operational semantics has nice 147 | properties with respect to heap extension: if a command is safe 148 | in a small heap, it is safe in a bigger heap, and any reduction 149 | from the bigger heap is simulated by a reduction from the smaller heap. *) 150 | 151 | Ltac inv H := inversion H; clear H; subst. 152 | 153 | Lemma immsafe_frame: forall h' c h, 154 | immsafe (c, h) -> hdisjoint h h' -> immsafe (c, hunion h h'). 155 | Proof. 156 | intros h' c h IMM; dependent induction IMM; intros DISJ. 157 | - constructor. 158 | - constructor; auto. 159 | - constructor. 160 | - destruct (isfinite (hunion h h')) as [l' FIN]. 161 | apply immsafe_alloc with (Z.max 1 l'). 162 | lia. 163 | intros. apply FIN. lia. 164 | - constructor. cbn. destruct (h l); congruence. 165 | - constructor. cbn. destruct (h l); congruence. 166 | - constructor. cbn. destruct (h l); congruence. 167 | - constructor. 168 | Qed. 169 | 170 | Lemma red_frame: forall h2 c h1 c' h', 171 | red (c, hunion h1 h2) (c', h') -> 172 | immsafe (c, h1) -> 173 | hdisjoint h1 h2 -> 174 | exists h1', red (c, h1) (c', h1') /\ hdisjoint h1' h2 /\ h' = hunion h1' h2. 175 | Proof. 176 | intros until h'; intros RED; dependent induction RED; intros IMM DISJ; inv IMM. 177 | - exists h1; intuition auto. constructor; auto. 178 | - exists h1; intuition auto. constructor; auto. 179 | - edestruct IHRED as (h1' & R & D & U); eauto. 180 | exists h1'; intuition auto. constructor; auto. 181 | - exists h1; intuition auto. constructor; auto. 182 | - exists (hinit l sz h1); intuition auto. 183 | + constructor; auto. intros. apply H in H1. cbn in H1. destruct (h1 i); congruence. 184 | + red; cbn; intros i. 185 | assert (EITHER: l <= i < l + Z.of_nat sz \/ (i < l \/ l + Z.of_nat sz <= i)) by lia. 186 | destruct EITHER. 187 | * right. apply H in H1. cbn in H1. destruct (h1 i), (h2 i); congruence. 188 | * rewrite hinit_outside by auto. apply DISJ. 189 | + apply heap_extensionality; intros i; cbn. 190 | assert (EITHER: l <= i < l + Z.of_nat sz \/ (i < l \/ l + Z.of_nat sz <= i)) by lia. 191 | destruct EITHER. 192 | * rewrite ! hinit_inside by auto. auto. 193 | * rewrite ! hinit_outside by auto. auto. 194 | - exists h1; intuition auto. constructor; auto. cbn in H. destruct (h1 l); congruence. 195 | - exists (hupdate l v h1); intuition auto. 196 | + constructor; auto. 197 | + intros i; cbn. generalize (DISJ i). 198 | destruct (Z.eq_dec l i); intuition congruence. 199 | + apply heap_extensionality; intros i; cbn. 200 | destruct (Z.eq_dec l i); auto. 201 | - exists (hfree l h1); intuition auto. 202 | + constructor; auto. 203 | + intros i; cbn. generalize (DISJ i). 204 | destruct (Z.eq_dec l i); intuition congruence. 205 | + apply heap_extensionality; intros i; cbn. 206 | destruct (Z.eq_dec l i); auto. 207 | subst i. generalize (DISJ l). intuition. 208 | Qed. 209 | 210 | Lemma safe_frame: 211 | forall (R: assertion) h', R h' -> 212 | forall c h Q, 213 | safe c h Q -> hdisjoint h h' -> safe c (hunion h h') (fun v => Q v ** R). 214 | Proof. 215 | induction 2; intros DISJ. 216 | - constructor. exists h, h'; auto. 217 | - constructor. auto. apply immsafe_frame; auto. 218 | intros. edestruct red_frame as (h1' & RED1 & D & U); eauto. subst h'0. 219 | apply H3; auto. 220 | Qed. 221 | 222 | Lemma triple_frame: forall P c Q R, 223 | ⦃ P ⦄ c ⦃ Q ⦄ -> 224 | ⦃ P ** R ⦄ c ⦃ fun v => Q v ** R ⦄. 225 | Proof. 226 | intros P c Q R TR h (h1 & h2 & P1 & R2 & D & U). subst h. 227 | apply safe_frame; auto. 228 | Qed. 229 | 230 | (** ** 2.3. The "small rules" for heap operations *) 231 | 232 | Lemma triple_get: forall l v, 233 | ⦃ contains l v ⦄ GET l ⦃ fun v' => (v' = v) //\\ contains l v ⦄. 234 | Proof. 235 | intros l v h P. 236 | assert (L: h l = Some v). 237 | { red in P. subst h. apply hupdate_same. } 238 | constructor; auto. 239 | - constructor. congruence. 240 | - intros c' h' RED. inv RED. constructor. split; auto; congruence. 241 | Qed. 242 | 243 | Lemma triple_set: forall l v, 244 | ⦃ valid l ⦄ SET l v ⦃ fun _ => contains l v ⦄. 245 | Proof. 246 | intros l v h (v0 & P). 247 | assert (L: h l = Some v0). 248 | { red in P; subst h; apply hupdate_same. } 249 | constructor; auto. 250 | - constructor. congruence. 251 | - intros c' h' RED. inv RED. constructor. 252 | red in P; subst h. 253 | red. apply heap_extensionality; intros l'; cbn. 254 | destruct (Z.eq_dec l l'); auto. 255 | Qed. 256 | 257 | Fixpoint valid_N (l: addr) (sz: nat) : assertion := 258 | match sz with O => emp | S sz => valid l ** valid_N (l + 1) sz end. 259 | 260 | Remark valid_N_init: forall sz l, 261 | (valid_N l sz) (hinit l sz hempty). 262 | Proof. 263 | induction sz as [ | sz]; intros l; cbn. 264 | - red; auto. 265 | - exists (hupdate l 0 hempty), (hinit (l + 1) sz hempty); intuition auto. 266 | + exists 0. red; auto. 267 | + intros x. unfold hupdate, hempty; cbn. destruct (Z.eq_dec l x); auto. 268 | right. rewrite hinit_outside by lia. auto. 269 | + apply heap_extensionality; intros x. cbn. destruct (Z.eq_dec l x); auto. 270 | Qed. 271 | 272 | Lemma triple_alloc: forall sz, 273 | ⦃ emp ⦄ 274 | ALLOC sz 275 | ⦃ fun l => (l <> 0) //\\ valid_N l sz ⦄. 276 | Proof. 277 | intros sz h P. red in P. subst h. 278 | constructor; auto. 279 | - apply immsafe_alloc with 1; intros. 280 | + lia. 281 | + auto. 282 | - intros c' h' RED; inv RED. constructor. 283 | split; auto. apply valid_N_init; auto. 284 | Qed. 285 | 286 | Lemma triple_free: forall l, 287 | ⦃ valid l ⦄ 288 | FREE l 289 | ⦃ fun _ => emp ⦄. 290 | Proof. 291 | intros l h (v0 & P). red in P. 292 | assert (L: h l = Some v0). 293 | { subst h. apply hupdate_same. } 294 | constructor; auto. 295 | - constructor. congruence. 296 | - intros c' h' RED; inv RED. constructor. 297 | red. apply heap_extensionality; intros x. cbn. 298 | destruct (Z.eq_dec l x); auto. 299 | Qed. 300 | 301 | (** ** 2.4. Properties of the [safe] predicate *) 302 | 303 | Lemma safe_pure: forall v h Q, 304 | safe (PURE v) h Q -> Q v h. 305 | Proof. 306 | intros. inv H. 307 | - auto. 308 | - contradiction. 309 | Qed. 310 | 311 | Lemma safe_red: forall c h Q c' h', 312 | safe c h Q -> red (c, h) (c', h') -> safe c' h' Q. 313 | Proof. 314 | intros. inv H. 315 | - inv H0. 316 | - eauto. 317 | Qed. 318 | 319 | Lemma safe_immsafe: forall c h Q, 320 | safe c h Q -> immsafe (c, h). 321 | Proof. 322 | intros. inv H. 323 | - constructor. 324 | - auto. 325 | Qed. 326 | 327 | Lemma safe_let: forall (Q R: postcond) f, 328 | (forall v h', Q v h' -> safe (f v) h' R) -> 329 | forall c h, 330 | safe c h Q -> 331 | safe (LET c f) h R. 332 | Proof. 333 | intros Q R f POST. induction 1. 334 | - constructor; auto. 335 | + constructor. constructor. 336 | + intros c' h' RED; inv RED. apply POST; auto. inv H1. 337 | - constructor; auto. 338 | + constructor; auto. 339 | + intros c' h' RED; inv RED. contradiction. eauto. 340 | Qed. 341 | 342 | Lemma safe_consequence: forall (Q Q': postcond), 343 | (forall v, Q v -->> Q' v) -> 344 | forall c h, safe c h Q -> safe c h Q'. 345 | Proof. 346 | intros Q Q' IMP. induction 1. 347 | - apply safe_done. apply IMP. assumption. 348 | - apply safe_step; auto. 349 | Qed. 350 | 351 | (** ** 2.5. Rules for control structures *) 352 | 353 | Lemma triple_pure: forall P v (Q: postcond), 354 | P -->> Q v -> 355 | ⦃ P ⦄ PURE v ⦃ Q ⦄. 356 | Proof. 357 | intros; intros h Ph. constructor. apply H; auto. 358 | Qed. 359 | 360 | Lemma triple_let: 361 | forall c f (P: precond) (Q R: postcond), 362 | ⦃ P ⦄ c ⦃ Q ⦄ -> 363 | (forall v, ⦃ Q v ⦄ f v ⦃ R ⦄) -> 364 | ⦃ P ⦄ LET c f ⦃ R ⦄. 365 | Proof. 366 | intros until R; intros HR1 HR2 h Ph. 367 | apply safe_let with Q. apply HR2. apply HR1. auto. 368 | Qed. 369 | 370 | Lemma triple_ifthenelse: forall b c1 c2 P Q, 371 | ⦃ (b <> 0) //\\ P ⦄ c1 ⦃ Q ⦄ -> 372 | ⦃ (b = 0) //\\ P ⦄ c2 ⦃ Q ⦄ -> 373 | ⦃ P ⦄ IFTHENELSE b c1 c2 ⦃ Q ⦄. 374 | Proof. 375 | intros until Q; intros HR1 HR2 h Ph. constructor; auto. 376 | - constructor. 377 | - intros c' h' RED; inv RED. destruct (Z.eqb_spec b 0). 378 | + apply HR2. split; auto. 379 | + apply HR1. split; auto. 380 | Qed. 381 | 382 | Lemma triple_consequence: forall P P' c Q' Q, 383 | ⦃ P' ⦄ c ⦃ Q' ⦄ -> 384 | P -->> P' -> (forall v, Q' v -->> Q v) -> 385 | ⦃ P ⦄ c ⦃ Q ⦄. 386 | Proof. 387 | intros; red; intros. apply safe_consequence with Q'; auto. 388 | Qed. 389 | 390 | Lemma triple_pick: forall n, 391 | ⦃ emp ⦄ 392 | PICK n 393 | ⦃ fun i => pure (0 <= i < n) ⦄. 394 | Proof. 395 | intros n h Ph. constructor; auto. 396 | - constructor. 397 | - intros c' h' RED; inv RED. constructor. split; auto. 398 | Qed. 399 | 400 | (** ** 2.6. Useful derived rules *) 401 | 402 | (** The following rules are heavily used in the examples of section 3. *) 403 | 404 | Lemma triple_consequence_pre: forall P P' c Q, 405 | ⦃ P' ⦄ c ⦃ Q ⦄ -> 406 | P -->> P' -> 407 | ⦃ P ⦄ c ⦃ Q ⦄. 408 | Proof. 409 | intros. apply triple_consequence with P' Q; auto. intros; red; auto. 410 | Qed. 411 | 412 | Lemma triple_consequence_post: forall P c Q Q', 413 | ⦃ P ⦄ c ⦃ Q' ⦄ -> 414 | (forall v, Q' v -->> Q v) -> 415 | ⦃ P ⦄ c ⦃ Q ⦄. 416 | Proof. 417 | intros. apply triple_consequence with P Q'; auto. red; auto. 418 | Qed. 419 | 420 | Lemma triple_lift_pure: forall (P: Prop) P' c Q, 421 | (P -> ⦃ P' ⦄ c ⦃ Q ⦄) -> 422 | ⦃ P //\\ P' ⦄ c ⦃ Q ⦄. 423 | Proof. 424 | intros. intros h [P1 P2]. apply H; auto. 425 | Qed. 426 | 427 | Lemma triple_lift_exists: forall (X: Type) (P: X -> assertion) c Q, 428 | (forall x, ⦃ P x ⦄ c ⦃ Q ⦄) -> 429 | ⦃ aexists P ⦄ c ⦃ Q ⦄. 430 | Proof. 431 | intros. intros h (x & Px). apply (H x); auto. 432 | Qed. 433 | 434 | Lemma triple_ifthen: forall b c1 c2 P Q, 435 | b <> 0 -> ⦃ P ⦄ c1 ⦃ Q ⦄ -> 436 | ⦃ P ⦄ IFTHENELSE b c1 c2 ⦃ Q ⦄. 437 | Proof. 438 | intros. apply triple_ifthenelse; apply triple_lift_pure; intros. 439 | - auto. 440 | - lia. 441 | Qed. 442 | 443 | Lemma triple_ifelse: forall b c1 c2 P Q, 444 | b = 0 -> ⦃ P ⦄ c2 ⦃ Q ⦄ -> 445 | ⦃ P ⦄ IFTHENELSE b c1 c2 ⦃ Q ⦄. 446 | Proof. 447 | intros. apply triple_ifthenelse; apply triple_lift_pure; intros. 448 | - lia. 449 | - auto. 450 | Qed. 451 | 452 | Lemma unroll_com: forall c, 453 | c = match c with 454 | | PURE x => PURE x 455 | | LET c f => LET c f 456 | | IFTHENELSE b c1 c2 => IFTHENELSE b c1 c2 457 | | ALLOC sz => ALLOC sz 458 | | GET l => GET l 459 | | SET l v => SET l v 460 | | FREE l => FREE l 461 | | PICK n => PICK n 462 | end. 463 | Proof. 464 | destruct c; auto. 465 | Qed. 466 | 467 | (** * 3. Singly-linked lists *) 468 | 469 | (** ** Representation predicate *) 470 | 471 | (** Here is a separation logic assertion that describes the in-memory 472 | representation of a list. 473 | - [a] is the pointer to the list head (or 0 if the list is empty). 474 | - [l] is the Coq list of the list elements. 475 | *) 476 | 477 | Fixpoint list_at (a: addr) (l: list Z) : assertion := 478 | match l with 479 | | nil => (a = 0) //\\ emp 480 | | h :: t => (a <> 0) //\\ aexists (fun a' => contains a h ** contains (a + 1) a' ** list_at a' t) 481 | end. 482 | 483 | (** ** The "cons" operation *) 484 | 485 | Definition list_cons (n: Z) (a: addr) : com := 486 | LET (ALLOC 2) (fun a' => SEQ (SET a' n) (SEQ (SET (a' + 1) a) (PURE a'))). 487 | 488 | Lemma list_cons_correct: forall a n l, 489 | ⦃ list_at a l ⦄ 490 | list_cons n a 491 | ⦃ fun a' => list_at a' (n :: l) ⦄. 492 | Proof. 493 | intros. eapply triple_let. 494 | rewrite <- sepconj_emp at 1. apply triple_frame. apply triple_alloc. 495 | intros b; simpl. rewrite lift_pureconj, ! sepconj_assoc, sepconj_emp. 496 | apply triple_lift_pure; intros H1. 497 | eapply triple_let. apply triple_frame. apply triple_set. simpl; intros _. 498 | eapply triple_let. rewrite sepconj_pick2. 499 | apply triple_frame. apply triple_set. simpl; intros _. 500 | rewrite sepconj_pick2. 501 | apply triple_pure. intros h A. split. auto. exists a; auto. 502 | Qed. 503 | 504 | (** ** Computing the length of a list *) 505 | 506 | (** Taking advantage of the coinductive nature of type [com], 507 | we use infinite commands to represent loops and tail-recursive functions. *) 508 | 509 | CoFixpoint list_length_rec (a: addr) (len: Z) : com := 510 | IFTHENELSE a (LET (GET (a + 1)) (fun t => list_length_rec t (len + 1))) (PURE len). 511 | 512 | Definition list_length (a: addr) : com := list_length_rec a 0. 513 | 514 | (** Normally we would write 515 | << 516 | len = 0; 517 | while (a != 0) { a = get (a + 1); len = len + 1; } 518 | >> 519 | With the coinductive definition, we write the equivalent infinite command 520 | << 521 | if (a == 0) return 0; else { 522 | a1 = get (a + 1); 523 | if (a1 == 0) return 1; else { 524 | a2 = get (a1 + 1); 525 | if (a2 == 0) return 2; else ... 526 | >> 527 | *) 528 | 529 | Lemma list_length_rec_correct: forall l a len, 530 | ⦃ list_at a l ⦄ 531 | list_length_rec a len 532 | ⦃ fun len' => (len' = len + Z.of_nat (List.length l)) //\\ list_at a l ⦄. 533 | Proof. 534 | Local Opaque Z.of_nat. 535 | induction l as [ | h t]; intros; rewrite (unroll_com (list_length_rec a len)); cbn. 536 | - apply triple_lift_pure; intro H1. 537 | apply triple_ifelse; auto. 538 | apply triple_pure. intros h H2. split. lia. split; auto. 539 | - apply triple_lift_pure; intro H1. 540 | apply triple_lift_exists; intros a'. 541 | apply triple_ifthen; auto. 542 | eapply triple_let. 543 | rewrite sepconj_pick2. apply triple_frame. apply triple_get. simpl. 544 | intros a''. rewrite lift_pureconj. apply triple_lift_pure; intros H3. subst a''. 545 | rewrite sepconj_swap3. 546 | eapply triple_consequence_post. 547 | apply triple_frame. apply IHt. intros len'; simpl. 548 | rewrite lift_pureconj. rewrite <- sepconj_swap3, sepconj_pick2. 549 | intros h1 (A & B). split. lia. split. auto. exists a'; auto. 550 | Qed. 551 | 552 | Corollary list_length_correct: forall l a, 553 | ⦃ list_at a l ⦄ 554 | list_length a 555 | ⦃ fun len => (len = Z.of_nat (length l)) //\\ list_at a l ⦄. 556 | Proof. 557 | intros. apply list_length_rec_correct. 558 | Qed. 559 | 560 | (** ** Concatenating two lists in-place *) 561 | 562 | (** In loop notation: 563 | << 564 | if (l1 == 0) return l2; else { 565 | t = get(l1 + 1); 566 | while (get (t + 1) != 0) t = get (t + 1); 567 | set (t + 1, l2); 568 | return l1; 569 | } 570 | >> 571 | *) 572 | 573 | CoFixpoint list_concat_rec (a1 a2: addr) : com := 574 | LET (GET (a1 + 1)) (fun t => IFTHENELSE t (list_concat_rec t a2) (SET (a1 + 1) a2)). 575 | 576 | Definition list_concat (a1 a2: addr) : com := 577 | IFTHENELSE a1 (SEQ (list_concat_rec a1 a2) (PURE a1)) (PURE a2). 578 | 579 | Lemma list_concat_rec_correct: forall l2 a2 l1 a1, 580 | a1 <> 0 -> 581 | ⦃ list_at a1 l1 ** list_at a2 l2 ⦄ 582 | list_concat_rec a1 a2 583 | ⦃ fun _ => list_at a1 (l1 ++ l2) ⦄. 584 | Proof. 585 | induction l1 as [ | h1 t1]; intros; rewrite (unroll_com (list_concat_rec a1 a2)); simpl. 586 | - rewrite lift_pureconj. apply triple_lift_pure; intros. lia. 587 | - rewrite lift_pureconj. apply triple_lift_pure. intros H1. 588 | rewrite lift_aexists. apply triple_lift_exists. intros a'. 589 | rewrite sepconj_assoc. 590 | eapply triple_let. 591 | + rewrite sepconj_assoc, sepconj_pick2. apply triple_frame. apply triple_get. 592 | + intros t. simpl. 593 | rewrite lift_pureconj. apply triple_lift_pure. intros H2; subst t. 594 | apply triple_ifthenelse. 595 | * apply triple_lift_pure. intros H2. 596 | rewrite <- sepconj_assoc, sepconj_comm. 597 | eapply triple_consequence_post. apply triple_frame. apply IHt1. auto. 598 | simpl. intros _. rewrite sepconj_pick2, sepconj_swap3. 599 | intros h P. split; auto. exists a'; auto. 600 | * apply triple_lift_pure. intros H2. 601 | eapply triple_consequence_post. 602 | apply triple_frame. 603 | eapply triple_consequence_pre. apply triple_set. 604 | intros h P; exists a'; auto. 605 | simpl. intros _. rewrite sepconj_pick2, sepconj_pick3. 606 | destruct t1; simpl. 607 | ** rewrite lift_pureconj, sepconj_emp. 608 | intros h (A & B). split; auto. exists a2; auto. 609 | ** rewrite lift_pureconj. intros h (A & B). lia. 610 | Qed. 611 | 612 | Lemma list_concat_correct: forall l1 a1 l2 a2, 613 | ⦃ list_at a1 l1 ** list_at a2 l2 ⦄ 614 | list_concat a1 a2 615 | ⦃ fun a => list_at a (l1 ++ l2) ⦄. 616 | Proof. 617 | intros. unfold list_concat. apply triple_ifthenelse. 618 | - apply triple_lift_pure; intros H1. 619 | eapply triple_let. apply list_concat_rec_correct; auto. 620 | simpl. intros _. apply triple_pure. red; auto. 621 | - apply triple_lift_pure; intros H1. 622 | destruct l1; simpl. 623 | + apply triple_pure. rewrite lift_pureconj, sepconj_emp. intros h (A & B); auto. 624 | + rewrite lift_pureconj. apply triple_lift_pure. intros; lia. 625 | Qed. 626 | 627 | (** ** List reversal in place *) 628 | 629 | (** In loop notation: 630 | << 631 | p = 0; 632 | while (l != 0) { 633 | n = get (l + 1); 634 | set (l + 1, p); 635 | p = l; 636 | l = n; 637 | } 638 | return p; 639 | >> 640 | *) 641 | 642 | CoFixpoint list_rev_rec (a p: addr) : com := 643 | IFTHENELSE a 644 | (LET (GET (a + 1)) (fun n => 645 | SEQ (SET (a + 1) p) 646 | (list_rev_rec n a))) 647 | (PURE p). 648 | 649 | Definition list_rev (a: addr) : com := list_rev_rec a 0. 650 | 651 | Lemma list_rev_rec_correct: forall l a l' p, 652 | ⦃ list_at a l ** list_at p l' ⦄ 653 | list_rev_rec a p 654 | ⦃ fun x => list_at x (List.rev_append l l') ⦄. 655 | Proof. 656 | induction l as [ | hd l]; intros; rewrite (unroll_com (list_rev_rec a p)); simpl. 657 | - rewrite lift_pureconj, sepconj_emp. apply triple_lift_pure; intros H1. 658 | apply triple_ifelse; auto. apply triple_pure. red; auto. 659 | - rewrite lift_pureconj; apply triple_lift_pure; intros H1. 660 | rewrite lift_aexists; apply triple_lift_exists; intros a'. 661 | apply triple_ifthen; auto. 662 | eapply triple_let. 663 | rewrite ! sepconj_assoc, sepconj_pick2. 664 | apply triple_frame. apply triple_get. intros a''. simpl. 665 | rewrite lift_pureconj. apply triple_lift_pure. intros H3. subst a''. 666 | eapply triple_let. 667 | apply triple_frame. eapply triple_consequence_pre. 668 | apply triple_set. 669 | intros h P; exists a'; auto. 670 | simpl. intros _. 671 | rewrite sepconj_pick2, sepconj_pick3. 672 | eapply triple_consequence_pre. 673 | apply IHl. 674 | simpl. apply sepconj_imp_r. intros h A. split; auto. exists p; auto. 675 | Qed. 676 | 677 | Lemma list_rev_correct: forall a l, 678 | ⦃ list_at a l ⦄ 679 | list_rev a 680 | ⦃ fun x => list_at x (List.rev l) ⦄. 681 | Proof. 682 | intros. rewrite List.rev_alt. 683 | eapply triple_consequence_pre. apply list_rev_rec_correct. 684 | simpl. rewrite sepconj_comm, lift_pureconj, sepconj_emp. 685 | intros h A; split; auto. 686 | Qed. 687 | 688 | (** * 4. An alternate definition of separation logic triples *) 689 | 690 | Module AlternateSeplog. 691 | 692 | (** For some languages, the frame property for reductions (lemma 693 | [red_frame] above) does not hold, e.g. because allocations are 694 | deterministic. Or maybe we do not want to prove the [red_frame] 695 | lemma. 696 | 697 | In this case, not all is lost: we can define our separation 698 | triples [ ⦃ P ⦄ c ⦃ Q ⦄ ] as Hoare triples plus framing. *) 699 | 700 | Definition Hoare (P: precond) (c: com) (Q: postcond) : Prop := 701 | forall h, P h -> safe c h Q. 702 | 703 | Definition triple (P: precond) (c: com) (Q: postcond) := 704 | forall (R: assertion), Hoare (P ** R) c (fun v => Q v ** R). 705 | 706 | Notation "⦃ P ⦄ c ⦃ Q ⦄" := (triple P c Q) (at level 90, c at next level). 707 | 708 | (** This definition validates the frame rule. *) 709 | 710 | Lemma triple_frame: forall P c Q R, 711 | ⦃ P ⦄ c ⦃ Q ⦄ -> 712 | ⦃ P ** R ⦄ c ⦃ fun v => Q v ** R ⦄. 713 | Proof. 714 | intros P c Q R TR R'. rewrite sepconj_assoc. 715 | replace (fun v => (Q v ** R) ** R') with (fun v => Q v ** (R ** R')). 716 | apply TR. 717 | apply functional_extensionality; intros. rewrite sepconj_assoc; auto. 718 | Qed. 719 | 720 | (** It also validates the "small rules" for heap operations. *) 721 | 722 | Lemma triple_get: forall l v, 723 | ⦃ contains l v ⦄ GET l ⦃ fun v' => (v' = v) //\\ contains l v ⦄. 724 | Proof. 725 | intros l v R h (h1 & h2 & H1 & H2 & D & U). 726 | assert (L1: h1 l = Some v). 727 | { red in H1. subst h1. apply hupdate_same. } 728 | assert (L: h l = Some v). 729 | { intros. rewrite U; simpl. rewrite L1; auto. } 730 | constructor; auto. 731 | - constructor. congruence. 732 | - intros c' h' RED. inv RED. constructor. 733 | exists h1, h2. unfold pureconj. intuition congruence. 734 | Qed. 735 | 736 | Lemma triple_set: forall l v, 737 | ⦃ valid l ⦄ SET l v ⦃ fun _ => contains l v ⦄. 738 | Proof. 739 | intros l v R h (h1 & h2 & H1 & H2 & D & U). 740 | destruct H1 as (v0 & H1). red in H1. 741 | assert (L1: h1 l = Some v0). 742 | { subst h1; apply hupdate_same. } 743 | assert (L: h l = Some v0). 744 | { rewrite U; cbn. rewrite L1; auto. } 745 | constructor; auto. 746 | - constructor. congruence. 747 | - intros c' h' RED. inv RED. constructor. 748 | exists (hupdate l v hempty), h2. 749 | split. red. auto. 750 | split. auto. 751 | split. intro l'. specialize (D l'). cbn in *. destruct D; auto. destruct (Z.eq_dec l l'); auto. congruence. 752 | apply heap_extensionality; intros l'; cbn. destruct (Z.eq_dec l l'); auto. 753 | Qed. 754 | 755 | Remark valid_N_init: 756 | forall (R: assertion) sz l h, 757 | R h -> 758 | (forall i, l <= i < l + Z.of_nat sz -> h i = None) -> 759 | (valid_N l sz ** R) (hinit l sz h). 760 | Proof. 761 | induction sz as [ | sz]; intros l h Rh EMPTY; cbn. 762 | - rewrite sepconj_emp. auto. 763 | - rewrite sepconj_assoc. exists (hupdate l 0 hempty), (hinit (l + 1) sz h). 764 | split. exists 0. red; auto. 765 | split. apply IHsz. auto. intros. apply EMPTY. lia. 766 | split. intros x. unfold hupdate, hempty; cbn. destruct (Z.eq_dec l x); auto. 767 | right. rewrite hinit_outside by lia. apply EMPTY; lia. 768 | apply heap_extensionality; intros x. cbn. destruct (Z.eq_dec l x); auto. 769 | Qed. 770 | 771 | Lemma triple_alloc: forall sz, 772 | ⦃ emp ⦄ 773 | ALLOC sz 774 | ⦃ fun l => (l <> 0) //\\ valid_N l sz ⦄. 775 | Proof. 776 | intros sz R h H. rewrite sepconj_emp in H. 777 | constructor; auto. 778 | - destruct (isfinite h) as (l0 & FIN). apply immsafe_alloc with (Z.max l0 1); intros. 779 | + lia. 780 | + apply FIN. lia. 781 | - intros c' h' RED; inv RED. constructor. 782 | rewrite lift_pureconj; split. auto. apply valid_N_init; auto. 783 | Qed. 784 | 785 | Lemma triple_free: forall l, 786 | ⦃ valid l ⦄ 787 | FREE l 788 | ⦃ fun _ => emp ⦄. 789 | Proof. 790 | intros l R h (h1 & h2 & H1 & H2 & D & U). 791 | destruct H1 as (v0 & H1). 792 | assert (L1: h1 l = Some v0). 793 | { rewrite H1. apply hupdate_same. } 794 | assert (L: h l = Some v0). 795 | { rewrite U; cbn. rewrite L1. auto. } 796 | constructor; auto. 797 | - constructor. congruence. 798 | - intros c' h' RED; inv RED. constructor. rewrite sepconj_emp. 799 | replace (hfree l (hunion h1 h2)) with h2; auto. 800 | apply heap_extensionality; intros x. generalize (D x); rewrite H1; cbn. 801 | destruct (Z.eq_dec l x); auto. intuition congruence. 802 | Qed. 803 | 804 | (** The rules for control structures are also valid. 805 | Proof plan: first show Hoare-style rules for the [Hoare] triple, 806 | then frame by an arbitrary [R] to obtain the separation triple. *) 807 | 808 | Lemma Hoare_pure: forall P v (Q: postcond), 809 | P -->> Q v -> 810 | Hoare P (PURE v) Q. 811 | Proof. 812 | intros; intros h Ph. constructor. apply H; auto. 813 | Qed. 814 | 815 | Lemma triple_pure: forall P v (Q: postcond), 816 | P -->> Q v -> 817 | ⦃ P ⦄ PURE v ⦃ Q ⦄. 818 | Proof. 819 | intros; intros R. apply Hoare_pure. apply sepconj_imp_l; auto. 820 | Qed. 821 | 822 | Lemma Hoare_let: 823 | forall c f (P: precond) (Q R: postcond), 824 | Hoare P c Q -> 825 | (forall v, Hoare (Q v) (f v) R) -> 826 | Hoare P (LET c f) R. 827 | Proof. 828 | intros until R; intros HR1 HR2 h Ph. 829 | apply safe_let with Q. apply HR2. apply HR1. auto. 830 | Qed. 831 | 832 | Lemma triple_let: 833 | forall c f (P: precond) (Q R: postcond), 834 | ⦃ P ⦄ c ⦃ Q ⦄ -> 835 | (forall v, ⦃ Q v ⦄ f v ⦃ R ⦄) -> 836 | ⦃ P ⦄ LET c f ⦃ R ⦄. 837 | Proof. 838 | intros c f P Q R TR1 TR2 R'. 839 | apply Hoare_let with (fun v => Q v ** R'). 840 | apply TR1. 841 | intros. apply TR2. 842 | Qed. 843 | 844 | Lemma Hoare_ifthenelse: forall b c1 c2 P Q, 845 | Hoare ((b <> 0) //\\ P) c1 Q -> 846 | Hoare ((b = 0) //\\ P) c2 Q -> 847 | Hoare P (IFTHENELSE b c1 c2) Q. 848 | Proof. 849 | intros until Q; intros HR1 HR2 h Ph. constructor; auto. 850 | - constructor. 851 | - intros c' h' RED; inv RED. destruct (Z.eqb_spec b 0). 852 | + apply HR2. split; auto. 853 | + apply HR1. split; auto. 854 | Qed. 855 | 856 | Lemma triple_ifthenelse: forall b c1 c2 P Q, 857 | ⦃ (b <> 0) //\\ P ⦄ c1 ⦃ Q ⦄ -> 858 | ⦃ (b = 0) //\\ P ⦄ c2 ⦃ Q ⦄ -> 859 | ⦃ P ⦄ IFTHENELSE b c1 c2 ⦃ Q ⦄. 860 | Proof. 861 | intros b c1 c2 P Q TR1 TR2 R. 862 | apply Hoare_ifthenelse; rewrite <- lift_pureconj; auto. 863 | Qed. 864 | 865 | Lemma Hoare_consequence: forall P P' c Q' Q, 866 | Hoare P' c Q' -> 867 | P -->> P' -> (forall v, Q' v -->> Q v) -> 868 | Hoare P c Q. 869 | Proof. 870 | intros; red; intros. apply safe_consequence with Q'; auto. 871 | Qed. 872 | 873 | Lemma triple_consequence: forall P P' c Q' Q, 874 | ⦃ P' ⦄ c ⦃ Q' ⦄ -> 875 | P -->> P' -> (forall v, Q' v -->> Q v) -> 876 | ⦃ P ⦄ c ⦃ Q ⦄. 877 | Proof. 878 | intros; red; intros. apply Hoare_consequence with (P' ** R) (fun v => Q' v ** R). 879 | apply H. 880 | apply sepconj_imp_l; auto. 881 | intros; apply sepconj_imp_l; auto. 882 | Qed. 883 | 884 | Lemma Hoare_pick: forall P n, 885 | Hoare P (PICK n) (fun i => (0 <= i < n) //\\ P). 886 | Proof. 887 | intros P n h Ph. constructor; auto. 888 | - constructor. 889 | - intros c' h' RED; inv RED. constructor. split; auto. 890 | Qed. 891 | 892 | Lemma triple_pick: forall n, 893 | ⦃ emp ⦄ 894 | PICK n 895 | ⦃ fun i => pure (0 <= i < n) ⦄. 896 | Proof. 897 | intros; intros R. rewrite sepconj_emp. eapply Hoare_consequence with (P' := R). apply Hoare_pick. 898 | red; auto. 899 | intros; red; intros. rewrite pureconj_sepconj. auto. 900 | Qed. 901 | 902 | End AlternateSeplog. 903 | 904 | (** * 5. Ramification *) 905 | 906 | (** Assume we have a triple [{P'} c {Q'}] and we want to conclude [{P} c {Q}]. 907 | In general, we need to frame the former triple by an appropriate [R], 908 | then use the consequence rule to conclude. *) 909 | 910 | Lemma triple_frame_consequence: forall R P c Q P' Q', 911 | ⦃ P ⦄ c ⦃ Q ⦄ -> 912 | P' -->> P ** R -> 913 | (forall v, Q v ** R -->> Q' v) -> 914 | ⦃ P' ⦄ c ⦃ Q' ⦄. 915 | Proof. 916 | intros. apply triple_consequence with (P ** R) (fun v => Q v ** R); auto. apply triple_frame; auto. 917 | Qed. 918 | 919 | (** This rule still needs the user to guess the framing predicate [R]. 920 | An alternate presentation uses the magic wand instead. 921 | This approach is called "ramification" in the literature. *) 922 | 923 | Lemma triple_ramification: forall P c Q P' Q', 924 | ⦃ P ⦄ c ⦃ Q ⦄ -> 925 | P' -->> P ** (aforall (fun v => Q v --* Q' v)) -> 926 | ⦃ P' ⦄ c ⦃ Q' ⦄. 927 | Proof. 928 | intros. eapply triple_frame_consequence with (R := aforall (fun v => Q v --* Q' v)). 929 | eassumption. 930 | assumption. 931 | intros v h (h1 & h2 & Q1 & W2 & D & U). 932 | apply (wand_cancel (Q v)). exists h1, h2; auto. 933 | Qed. 934 | 935 | (** * 6. Weakest preconditions *) 936 | 937 | (** ** 6.1. Definition and characterization *) 938 | 939 | (** Here is one possible definition of the weakest precondition for 940 | command [c] with postcondition [Q]. *) 941 | 942 | Definition wp (c: com) (Q: postcond) : precond := 943 | aexists (fun P => ⦃ P ⦄ c ⦃ Q ⦄ //\\ P). 944 | 945 | (** What matters about [wp c Q] is that it is a precondition... *) 946 | 947 | Lemma wp_precond: forall c Q, 948 | ⦃ wp c Q ⦄ c ⦃ Q ⦄. 949 | Proof. 950 | intros c Q h (P & T & C). apply T. auto. 951 | Qed. 952 | 953 | (** ... and it is implied by any other precondition. *) 954 | 955 | Lemma wp_weakest: forall P c Q, 956 | ⦃ P ⦄ c ⦃ Q ⦄ -> 957 | P -->> wp c Q. 958 | Proof. 959 | intros P c Q T h Ph. exists P; split; auto. 960 | Qed. 961 | 962 | (** This leads to the following alternate definition of triples in terms 963 | of weakest preconditions. *) 964 | 965 | Corollary wp_equiv: forall P c Q, 966 | ⦃ P ⦄ c ⦃ Q ⦄ <-> (P -->> wp c Q). 967 | Proof. 968 | intros; split; intros. 969 | - apply wp_weakest; auto. 970 | - apply triple_consequence_pre with (wp c Q); auto using wp_precond. 971 | Qed. 972 | 973 | (** Here is another definition of the weakest precondition, using the 974 | operational semantics directly. *) 975 | 976 | Definition wp' (c: com) (Q: postcond) : precond := 977 | fun h => safe c h Q. 978 | 979 | Lemma wp'_precond: forall c Q, 980 | ⦃ wp' c Q ⦄ c ⦃ Q ⦄. 981 | Proof. 982 | intros c Q h SAFE. apply SAFE. 983 | Qed. 984 | 985 | Lemma wp'_weakest: forall P c Q, 986 | ⦃ P ⦄ c ⦃ Q ⦄ -> 987 | P -->> wp' c Q. 988 | Proof. 989 | intros; intros h Ph. apply H. auto. 990 | Qed. 991 | 992 | (** ** 6.2. Structural rules for weakest preconditions *) 993 | 994 | Lemma wp_consequence: forall (Q Q': postcond) c, 995 | (forall v, Q v -->> Q' v) -> 996 | wp c Q -->> wp c Q'. 997 | Proof. 998 | intros. apply wp_weakest. apply triple_consequence_post with Q; auto using wp_precond. 999 | Qed. 1000 | 1001 | Lemma wp_frame: forall R c Q, 1002 | wp c Q ** R -->> wp c (fun v => Q v ** R). 1003 | Proof. 1004 | intros. apply wp_weakest. apply triple_frame. apply wp_precond. 1005 | Qed. 1006 | 1007 | Corollary wp_frame_consequence: forall R Q c Q', 1008 | (forall v, Q v ** R -->> Q' v) -> 1009 | wp c Q ** R -->> wp c Q'. 1010 | Proof. 1011 | intros; red; intros. apply wp_consequence with (fun v => Q v ** R). assumption. 1012 | apply wp_frame; auto. 1013 | Qed. 1014 | 1015 | Corollary wp_ramification: forall c Q Q', 1016 | wp c Q ** aforall (fun v => Q v --* Q' v) -->> wp c Q'. 1017 | Proof. 1018 | intros. apply wp_frame_consequence. 1019 | intros v h (h1 & h2 & A & B & D & U). apply (wand_cancel (Q v)). exists h1, h2; auto. 1020 | Qed. 1021 | 1022 | (** ** 6.3. Weakest precondition rules for our language of pointers *) 1023 | 1024 | Lemma wp_pure: forall (Q: postcond) v, 1025 | Q v -->> wp (PURE v) Q. 1026 | Proof. 1027 | intros. apply wp_weakest. apply triple_pure. red; auto. 1028 | Qed. 1029 | 1030 | Lemma wp_let: forall c f Q, 1031 | wp c (fun v => wp (f v) Q) -->> wp (LET c f) Q. 1032 | Proof. 1033 | intros. apply wp_weakest. eapply triple_let. 1034 | apply wp_precond. 1035 | intros. apply wp_precond. 1036 | Qed. 1037 | 1038 | Lemma wp_ifthenelse: forall b c1 c2 Q, 1039 | (if b =? 0 then wp c2 Q else wp c1 Q) -->> wp (IFTHENELSE b c1 c2) Q. 1040 | Proof. 1041 | intros. apply wp_weakest. apply triple_ifthenelse. 1042 | - apply triple_consequence_pre with (wp c1 Q). apply wp_precond. 1043 | intros h (A & B). rewrite <- Z.eqb_neq in A. rewrite A in B. auto. 1044 | - apply triple_consequence_pre with (wp c2 Q). apply wp_precond. 1045 | intros h (A & B). subst b. auto. 1046 | Qed. 1047 | 1048 | Lemma wp_alloc: forall sz Q, 1049 | aforall (fun l => (l <> 0) //\\ valid_N l sz --* Q l) -->> wp (ALLOC sz) Q. 1050 | Proof. 1051 | intros; red; intros. 1052 | apply wp_ramification with (Q := fun l => (l <> 0) //\\ valid_N l sz). 1053 | apply sepconj_imp_l with emp. 1054 | apply wp_weakest. apply triple_alloc. 1055 | rewrite sepconj_emp. assumption. 1056 | Qed. 1057 | 1058 | Lemma wp_get: forall l v Q, 1059 | contains l v ** (contains l v --* Q v) -->> wp (GET l) Q. 1060 | Proof. 1061 | intros. 1062 | assert (W: contains l v -->> wp (GET l) (fun v' => (v' = v) //\\ contains l v)). 1063 | { apply wp_weakest. apply triple_get. } 1064 | intros; red; intros. 1065 | eapply wp_ramification. eapply sepconj_imp_l. eexact W. 1066 | eapply sepconj_imp_r. 2: eexact H. 1067 | intros h' H' v' h'' D (A & B). subst v'. apply H'; auto. 1068 | Qed. 1069 | 1070 | Lemma wp_set: forall l v Q, 1071 | valid l ** aforall (fun v' => (contains l v --* Q v')) -->> wp (SET l v) Q. 1072 | Proof. 1073 | intros. 1074 | assert (W: valid l -->> wp (SET l v) (fun _ => contains l v)). 1075 | { apply wp_weakest. apply triple_set. } 1076 | intros; red; intros. 1077 | eapply wp_ramification. eapply sepconj_imp_l. eexact W. 1078 | eapply sepconj_imp_r. 2: eexact H. 1079 | red; auto. 1080 | Qed. 1081 | 1082 | Corollary wp_set': forall l v Q, 1083 | valid l ** (contains l v --* Q) -->> wp (SET l v) (fun _ => Q). 1084 | Proof. 1085 | intros; red; intros. apply wp_set. eapply sepconj_imp_r; eauto. 1086 | intros h' H' v'. auto. 1087 | Qed. 1088 | 1089 | Lemma wp_free: forall l Q, 1090 | valid l ** aforall (fun v' => Q v') -->> wp (FREE l) Q. 1091 | Proof. 1092 | intros. 1093 | assert (W: valid l -->> wp (FREE l) (fun _ => emp)). 1094 | { apply wp_weakest. apply triple_free. } 1095 | intros; red; intros. 1096 | eapply wp_ramification. eapply sepconj_imp_l. eexact W. 1097 | eapply sepconj_imp_r. 2: eexact H. 1098 | red; intros. intros v h' D E. rewrite E in *. rewrite hunion_comm, hunion_empty by HDISJ. 1099 | apply H0. 1100 | Qed. 1101 | 1102 | Corollary wp_free': forall l Q, 1103 | valid l ** Q -->> wp (FREE l) (fun _ => Q). 1104 | Proof. 1105 | intros; red; intros. apply wp_free. eapply sepconj_imp_r; eauto. 1106 | intros h' H' v'. auto. 1107 | Qed. 1108 | 1109 | Lemma wp_pick: forall n Q, 1110 | aforall (fun i => pure (0 <= i < n) --* Q i) -->> wp (PICK n) Q. 1111 | Proof. 1112 | intros. 1113 | assert (W: emp -->> wp (PICK n) (fun i => pure (0 <= i < n))). 1114 | { apply wp_weakest. apply triple_pick. } 1115 | intros; red; intros. 1116 | eapply wp_ramification. eapply sepconj_imp_l. eexact W. 1117 | eapply sepconj_imp_r. 2: rewrite sepconj_emp; eexact H. 1118 | red; auto. 1119 | Qed. 1120 | 1121 | 1122 | 1123 | -------------------------------------------------------------------------------- /Sequences.v: -------------------------------------------------------------------------------- 1 | (** A library of operators over relations, 2 | to define transition sequences and their properties. *) 3 | 4 | Set Implicit Arguments. 5 | 6 | Section SEQUENCES. 7 | 8 | Variable A: Type. (**r the type of states *) 9 | Variable R: A -> A -> Prop. (**r the transition relation between states *) 10 | 11 | (** ** Finite sequences of transitions *) 12 | 13 | (** Zero, one or several transitions: reflexive transitive closure of [R]. *) 14 | 15 | Inductive star: A -> A -> Prop := 16 | | star_refl: forall a, 17 | star a a 18 | | star_step: forall a b c, 19 | R a b -> star b c -> star a c. 20 | 21 | Lemma star_one: 22 | forall (a b: A), R a b -> star a b. 23 | Proof. 24 | eauto using star. 25 | Qed. 26 | 27 | Lemma star_trans: 28 | forall (a b: A), star a b -> forall c, star b c -> star a c. 29 | Proof. 30 | induction 1; eauto using star. 31 | Qed. 32 | 33 | (** One or several transitions: transitive closure of [R]. *) 34 | 35 | Inductive plus: A -> A -> Prop := 36 | | plus_left: forall a b c, 37 | R a b -> star b c -> plus a c. 38 | 39 | Lemma plus_one: 40 | forall a b, R a b -> plus a b. 41 | Proof. 42 | eauto using star, plus. 43 | Qed. 44 | 45 | Lemma plus_star: 46 | forall a b, 47 | plus a b -> star a b. 48 | Proof. 49 | intros. inversion H. eauto using star. 50 | Qed. 51 | 52 | Lemma plus_star_trans: 53 | forall a b c, plus a b -> star b c -> plus a c. 54 | Proof. 55 | intros. inversion H. eauto using plus, star_trans. 56 | Qed. 57 | 58 | Lemma star_plus_trans: 59 | forall a b c, star a b -> plus b c -> plus a c. 60 | Proof. 61 | intros. inversion H0. inversion H; eauto using plus, star, star_trans. 62 | Qed. 63 | 64 | Lemma plus_right: 65 | forall a b c, star a b -> R b c -> plus a c. 66 | Proof. 67 | eauto using star_plus_trans, plus_one. 68 | Qed. 69 | 70 | (** Absence of transitions from a state. *) 71 | 72 | Definition irred (a: A) : Prop := forall b, ~(R a b). 73 | 74 | (** A variant of [star] that counts the number of transitions. *) 75 | 76 | Inductive starN: nat -> A -> A -> Prop := 77 | | starN_refl: forall a, 78 | starN O a a 79 | | starN_step: forall n a b c, 80 | R a b -> starN n b c -> starN (S n) a c. 81 | 82 | Lemma starN_star: 83 | forall n a b, starN n a b -> star a b. 84 | Proof. 85 | induction 1; econstructor; eauto. 86 | Qed. 87 | 88 | Lemma star_starN: 89 | forall a b, star a b -> exists n, starN n a b. 90 | Proof. 91 | induction 1. 92 | - exists O; constructor. 93 | - destruct IHstar as (n & Sn). exists (S n); econstructor; eauto. 94 | Qed. 95 | 96 | (** ** Infinite transition sequences *) 97 | 98 | (** It is easy to characterize the fact that all transition sequences starting 99 | from a state [a] are infinite: it suffices to say that any finite sequence 100 | starting from [a] can always be extended by one more transition. *) 101 | 102 | Definition all_seq_inf (a: A) : Prop := 103 | forall b, star a b -> exists c, R b c. 104 | 105 | (** However, this is not the notion we are trying to characterize: the case 106 | where there exists an infinite sequence of transitions starting from [a], 107 | [a --> a1 --> a2 --> ... -> aN -> ...] 108 | leaving open the possibility that there exists finite sequences 109 | starting from [a]. 110 | 111 | Example: consider [A = nat] and [R] such that [R 0 0] and [R 0 1]. 112 | [all_seq_inf 0] does not hold, because a sequence [0 -->* 1] cannot 113 | be extended. Yet, [R] admits an infinite sequence, namely 114 | [0 --> 0 --> ...]. 115 | 116 | Another attempt would be to represent the sequence of states 117 | [a0 --> a1 --> a2 --> ... -> aN -> ...] explicitly, as a function 118 | [f: nat -> A] such that [f i] is the [i]-th state [ai] of the sequence. *) 119 | 120 | Definition infseq_with_function (a: A) : Prop := 121 | exists f: nat -> A, f 0 = a /\ forall i, R (f i) (f (1 + i)). 122 | 123 | (** This is a correct characterization of the existence of an infinite 124 | sequence of reductions. However, it is inconvenient to work with 125 | this definition in Coq's constructive logic: in most use cases, the 126 | function [f] is not computable and therefore cannot be defined in Coq. 127 | 128 | However, we do not really need the function [f]: its codomain [X] is 129 | all we need! What matters is the existence of a set [X] such as 130 | [a] is in [X], and 131 | every [b] in [X] can make a transition to an element of [X]. 132 | This suffices to prove the existence of an infinite sequence of transitions 133 | starting from [a]. 134 | *) 135 | 136 | Definition infseq (a: A) : Prop := 137 | exists X: A -> Prop, 138 | X a /\ (forall a1, X a1 -> exists a2, R a1 a2 /\ X a2). 139 | 140 | (** This definition is essentially a coinduction principle. 141 | Let us show some expected properties. For instance: if relation [R] 142 | contains a cycle, an infinite sequence exists. *) 143 | 144 | Remark cycle_infseq: 145 | forall a, R a a -> infseq a. 146 | Proof. 147 | intros. exists (fun b => b = a); split. 148 | auto. 149 | intros. subst a1. exists a; auto. 150 | Qed. 151 | 152 | (** Mon generally: if all sequences from [a] are infinite, there exists one 153 | infinite sequence starting in [a]. *) 154 | 155 | Lemma infseq_if_all_seq_inf: 156 | forall a, all_seq_inf a -> infseq a. 157 | Proof. 158 | intros a0 ALL0. 159 | exists all_seq_inf; split; auto. 160 | intros a1 ALL1. destruct (ALL1 a1) as [a2 R12]. constructor. 161 | exists a2; split; auto. 162 | intros a3 S23. destruct (ALL1 a3) as [a4 R23]. apply star_step with a2; auto. 163 | exists a4; auto. 164 | Qed. 165 | 166 | (** Likewise, the characterization [infseq_with_function] based on functions 167 | implies [infseq]. *) 168 | 169 | Lemma infseq_from_function: 170 | forall a, infseq_with_function a -> infseq a. 171 | Proof. 172 | intros a0 (f & F0 & Fn). exists (fun a => exists i, f i = a); split. 173 | - exists 0; auto. 174 | - intros a1 (i1 & F1). subst a1. exists (f (1 + i1)); split; auto. exists (1 + i1); auto. 175 | Qed. 176 | 177 | (** An "inversion lemma" for [infseq]: if [infseq a], i.e. there exists 178 | an infinite sequence starting in [a], then [a] can transition to a state [b] 179 | that satisfies [infseq b]. *) 180 | 181 | Lemma infseq_inv: 182 | forall a, infseq a -> exists b, R a b /\ infseq b. 183 | Proof. 184 | intros a (X & Xa & XP). destruct (XP a Xa) as (b & Rab & Xb). 185 | exists b; split; auto. exists X; auto. 186 | Qed. 187 | 188 | (** A very useful coinduction principle considers a set [X] where for 189 | every [a] in [X], we can make one *or several* transitions to reach 190 | a state [b] that belongs to [X]. *) 191 | 192 | Lemma infseq_coinduction_principle: 193 | forall (X: A -> Prop), 194 | (forall a, X a -> exists b, plus a b /\ X b) -> 195 | forall a, X a -> infseq a. 196 | Proof. 197 | intros X H a0 Xa0. 198 | exists (fun a => exists b, star a b /\ X b); split. 199 | - exists a0; auto using star_refl. 200 | - intros a1 (a2 & S12 & X2). inversion S12; subst. 201 | + destruct (H a2 X2) as (a3 & P23 & X3). inversion P23; subst. 202 | exists b; split; auto. exists a3; auto. 203 | + exists b; split; auto. exists a2; auto. 204 | Qed. 205 | 206 | (** ** Determinism properties for functional transition relations. *) 207 | 208 | (** A transition relation is functional if every state can transition 209 | to at most one other state. *) 210 | 211 | Hypothesis R_functional: 212 | forall a b c, R a b -> R a c -> b = c. 213 | 214 | (** Uniqueness of finite transition sequences. *) 215 | 216 | Lemma star_star_inv: 217 | forall a b, star a b -> forall c, star a c -> star b c \/ star c b. 218 | Proof. 219 | induction 1; intros. 220 | - auto. 221 | - inversion H1; subst. 222 | + right. eauto using star. 223 | + assert (b = b0) by (eapply R_functional; eauto). subst b0. 224 | apply IHstar; auto. 225 | Qed. 226 | 227 | Lemma finseq_unique: 228 | forall a b b', 229 | star a b -> irred b -> 230 | star a b' -> irred b' -> 231 | b = b'. 232 | Proof. 233 | intros. destruct (star_star_inv H H1). 234 | - inversion H3; subst. auto. elim (H0 _ H4). 235 | - inversion H3; subst. auto. elim (H2 _ H4). 236 | Qed. 237 | 238 | (** A state cannot both diverge and terminate on an irreducible state. *) 239 | 240 | Lemma infseq_inv': 241 | forall a b, R a b -> infseq a -> infseq b. 242 | Proof. 243 | intros a b Rab Ia. 244 | destruct (infseq_inv Ia) as (b' & Rab' & Xb'). 245 | assert (b' = b) by (eapply R_functional; eauto). 246 | subst b'. auto. 247 | Qed. 248 | 249 | Lemma infseq_star_inv: 250 | forall a b, star a b -> infseq a -> infseq b. 251 | Proof. 252 | induction 1; intros. 253 | - auto. 254 | - apply IHstar. apply infseq_inv' with a; auto. 255 | Qed. 256 | 257 | Lemma infseq_finseq_excl: 258 | forall a b, 259 | star a b -> irred b -> infseq a -> False. 260 | Proof. 261 | intros. 262 | destruct (@infseq_inv b) as (c & Rbc & _). eapply infseq_star_inv; eauto. 263 | apply (H0 c); auto. 264 | Qed. 265 | 266 | (** If there exists an infinite sequence of transitions from [a], 267 | all sequences of transitions arising from [a] are infinite. *) 268 | 269 | Lemma infseq_all_seq_inf: 270 | forall a, infseq a -> all_seq_inf a. 271 | Proof. 272 | intros. unfold all_seq_inf. intros. 273 | destruct (@infseq_inv b) as (c & Rbc & _). eapply infseq_star_inv; eauto. 274 | exists c; auto. 275 | Qed. 276 | 277 | End SEQUENCES. 278 | 279 | 280 | 281 | -------------------------------------------------------------------------------- /_CoqProject: -------------------------------------------------------------------------------- 1 | -R . CDF 2 | Sequences.v 3 | Hoare.v 4 | Separation.v 5 | Seplog.v 6 | CSL.v 7 | Delay.v 8 | Monads.v 9 | -------------------------------------------------------------------------------- /docs/.nojekyll: -------------------------------------------------------------------------------- https://raw.githubusercontent.com/xavierleroy/cdf-program-logics/5c16d588e702810688180e80576d1a7b2e65085c/docs/.nojekyll -------------------------------------------------------------------------------- /docs/CDF.Delay.html: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 5 | 6 | 7 | Module Delay 8 | 9 | 10 | 11 | 12 | 13 | 14 |

Module Delay

15 |
16 |
The delay type and Capretta's partiality monad
17 |
18 |
The type delay A represents computations that produce a result 19 | of type A if they terminate, but can also diverge.
20 |
21 | CoInductive delay (A: Type) :=
22 |   | now : A -> delay A
23 |   | later : delay A -> delay A.
24 |
25 | Arguments now [A].
26 | Arguments later [A].
27 |
28 | Lemma u_delay:
29 |   forall {A: Type} (x: delay A),
30 |   x = match x with now v => now v | later y => later y end.
31 |
Proof.
32 |
33 | destruct x; auto. Qed.
34 |
35 | Ltac unroll_delay X := rewrite (u_delay X); cbn.
36 |
37 |
The monad structure.
38 |
39 | Definition ret := now.
40 |
41 | CoFixpoint bind {A B: Type} (d: delay A) (f: A -> delay B) :=
42 |   match d with
43 |   | now v => later (f v)
44 |   | later d' => later (bind d' f)
45 |   end.
46 |
47 |
safe Q d means: if the computation d terminates, 48 | its value satisfies predicate Q. It's like a postcondition 49 | in a weak Hoare triple.
50 |
51 | CoInductive safe {A: Type} (Q: A -> Prop) : delay A -> Prop :=
52 |   | safe_now: forall a, Q a -> safe Q (now a)
53 |   | safe_later: forall d, safe Q d -> safe Q (later d).
54 |
55 | Lemma safe_inv_now: forall {A: Type} (Q: A -> Prop) v,
56 |   safe Q (now v) -> Q v.
57 |
Proof.
58 |
59 |   intros. inversion H. auto.
60 | Qed.
61 |
62 | Lemma safe_inv_later: forall {A: Type} (Q: A -> Prop) d,
63 |   safe Q (later d) -> safe Q d.
64 |
Proof.
65 |
66 |   intros. inversion H. auto.
67 | Qed.
68 |
69 | Lemma safe_consequence:
70 |   forall {A: Type} (Q1 Q2: A -> Prop),
71 |   (forall a, Q1 a -> Q2 a) ->
72 |   forall (d: delay A), safe Q1 d -> safe Q2 d.
73 |
Proof.
74 |
75 |   intros until Q2; intros IMP. cofix COINDHYP; destruct 1.
76 | - constructor; auto.
77 | - constructor; auto.
78 | Qed.
79 |
80 | Lemma safe_bind:
81 |   forall {A B: Type} (Q1: A -> Prop) (Q2: B -> Prop)
82 |          (d: delay A) (f: A -> delay B),
83 |   safe Q1 d -> (forall v, Q1 v -> safe Q2 (f v)) -> safe Q2 (bind d f).
84 |
Proof.
85 |
86 |   intros until Q2. cofix COINDHYP; intros.
87 |   unroll_delay (bind d f). destruct d.
88 | - apply safe_inv_now in H. constructor. apply H0. auto.
89 | - apply safe_inv_later in H. constructor. apply COINDHYP; auto.
90 | Qed.
91 | 92 |
93 | 94 | 95 | 96 | -------------------------------------------------------------------------------- /docs/CDF.Sequences.html: -------------------------------------------------------------------------------- 1 | 2 | 3 | 4 | 5 | 6 | 7 | Module Sequences 8 | 9 | 10 | 11 | 12 | 13 | 14 |

Module Sequences

15 |
16 |
A library of operators over relations, 17 | to define transition sequences and their properties.
18 |
19 | Set Implicit Arguments.
20 |
21 | Section SEQUENCES.
22 |
23 | Variable A: Type. (* the type of states *)
24 | Variable R: A -> A -> Prop. (* the transition relation between states *)
25 |
26 |

Finite sequences of transitions

27 |
28 |
Zero, one or several transitions: reflexive transitive closure of R.
29 |
30 | Inductive star: A -> A -> Prop :=
31 |   | star_refl: forall a,
32 |       star a a
33 |   | star_step: forall a b c,
34 |       R a b -> star b c -> star a c.
35 |
36 | Lemma star_one:
37 |   forall (a b: A), R a b -> star a b.
38 |
Proof.
39 |
40 |   eauto using star.
41 | Qed.
42 |
43 | Lemma star_trans:
44 |   forall (a b: A), star a b -> forall c, star b c -> star a c.
45 |
Proof.
46 |
47 |   induction 1; eauto using star.
48 | Qed.
49 |
50 |
One or several transitions: transitive closure of R.
51 |
52 | Inductive plus: A -> A -> Prop :=
53 |   | plus_left: forall a b c,
54 |       R a b -> star b c -> plus a c.
55 |
56 | Lemma plus_one:
57 |   forall a b, R a b -> plus a b.
58 |
Proof.
59 |
60 |   eauto using star, plus.
61 | Qed.
62 |
63 | Lemma plus_star:
64 |   forall a b,
65 |   plus a b -> star a b.
66 |
Proof.
67 |
68 |   intros. inversion H. eauto using star.
69 | Qed.
70 |
71 | Lemma plus_star_trans:
72 |   forall a b c, plus a b -> star b c -> plus a c.
73 |
Proof.
74 |
75 |   intros. inversion H. eauto using plus, star_trans.
76 | Qed.
77 |
78 | Lemma star_plus_trans:
79 |   forall a b c, star a b -> plus b c -> plus a c.
80 |
Proof.
81 |
82 |   intros. inversion H0. inversion H; eauto using plus, star, star_trans.
83 | Qed.
84 |
85 | Lemma plus_right:
86 |   forall a b c, star a b -> R b c -> plus a c.
87 |
Proof.
88 |
89 |   eauto using star_plus_trans, plus_one.
90 | Qed.
91 |
92 |
Absence of transitions from a state.
93 |
94 | Definition irred (a: A) : Prop := forall b, ~(R a b).
95 |
96 |
A variant of star that counts the number of transitions.
97 |
98 | Inductive starN: nat -> A -> A -> Prop :=
99 |   | starN_refl: forall a,
100 |       starN O a a
101 |   | starN_step: forall n a b c,
102 |       R a b -> starN n b c -> starN (S n) a c.
103 |
104 | Lemma starN_star:
105 |   forall n a b, starN n a b -> star a b.
106 |
Proof.
107 |
108 |   induction 1; econstructor; eauto.
109 | Qed.
110 |
111 | Lemma star_starN:
112 |   forall a b, star a b -> exists n, starN n a b.
113 |
Proof.
114 |
115 |   induction 1.
116 | - exists O; constructor.
117 | - destruct IHstar as (n & Sn). exists (S n); econstructor; eauto.
118 | Qed.
119 |
120 |

Infinite transition sequences

121 |
122 |
It is easy to characterize the fact that all transition sequences starting 123 | from a state a are infinite: it suffices to say that any finite sequence 124 | starting from a can always be extended by one more transition.
125 |
126 | Definition all_seq_inf (a: A) : Prop :=
127 |   forall b, star a b -> exists c, R b c.
128 |
129 |
However, this is not the notion we are trying to characterize: the case 130 | where there exists an infinite sequence of transitions starting from a, 131 | a --> a1 --> a2 --> ... -> aN -> ... 132 | leaving open the possibility that there exists finite sequences 133 | starting from a. 134 | Example: consider A = nat and R such that R 0 0 and R 0 1. 135 | all_seq_inf 0 does not hold, because a sequence 0 -->* 1 cannot 136 | be extended. Yet, R admits an infinite sequence, namely 137 | 0 --> 0 --> .... 138 | Another attempt would be to represent the sequence of states 139 | a0 --> a1 --> a2 --> ... -> aN -> ... explicitly, as a function 140 | f: nat -> A such that f i is the i-th state ai of the sequence.
141 |
142 | Definition infseq_with_function (a: A) : Prop :=
143 |   exists f: nat -> A, f 0 = a /\ forall i, R (f i) (f (1 + i)).
144 |
145 |
This is a correct characterization of the existence of an infinite 146 | sequence of reductions. However, it is inconvenient to work with 147 | this definition in Coq's constructive logic: in most use cases, the 148 | function f is not computable and therefore cannot be defined in Coq. 149 | However, we do not really need the function f: its codomain X is 150 | all we need! What matters is the existence of a set X such as 151 | a is in X, and 152 | every b in X can make a transition to an element of X. 153 | This suffices to prove the existence of an infinite sequence of transitions 154 | starting from a. 155 |
156 |
157 | Definition infseq (a: A) : Prop :=
158 |   exists X: A -> Prop,
159 |   X a /\ (forall a1, X a1 -> exists a2, R a1 a2 /\ X a2).
160 |
161 |
This definition is essentially a coinduction principle. 162 | Let us show some expected properties. For instance: if relation R 163 | contains a cycle, an infinite sequence exists.
164 |
165 | Remark cycle_infseq:
166 |   forall a, R a a -> infseq a.
167 |
Proof.
168 |
169 |   intros. exists (fun b => b = a); split.
170 |   auto.
171 |   intros. subst a1. exists a; auto.
172 | Qed.
173 |
174 |
Mon generally: if all sequences from a are infinite, there exists one 175 | infinite sequence starting in a.
176 |
177 | Lemma infseq_if_all_seq_inf:
178 |   forall a, all_seq_inf a -> infseq a.
179 |
Proof.
180 |
181 |   intros a0 ALL0.
182 |   exists all_seq_inf; split; auto.
183 |   intros a1 ALL1. destruct (ALL1 a1) as [a2 R12]. constructor.
184 |   exists a2; split; auto.
185 |   intros a3 S23. destruct (ALL1 a3) as [a4 R23]. apply star_step with a2; auto.
186 |   exists a4; auto.
187 | Qed.
188 |
189 |
Likewise, the characterization infseq_with_function based on functions 190 | implies infseq.
191 |
192 | Lemma infseq_from_function:
193 |   forall a, infseq_with_function a -> infseq a.
194 |
Proof.
195 |
196 |   intros a0 (f & F0 & Fn). exists (fun a => exists i, f i = a); split.
197 | - exists 0; auto.
198 | - intros a1 (i1 & F1). subst a1. exists (f (1 + i1)); split; auto. exists (1 + i1); auto.
199 | Qed.
200 |
201 |
An "inversion lemma" for infseq: if infseq a, i.e. there exists 202 | an infinite sequence starting in a, then a can transition to a state b 203 | that satisfies infseq b.
204 |
205 | Lemma infseq_inv:
206 |   forall a, infseq a -> exists b, R a b /\ infseq b.
207 |
Proof.
208 |
209 |   intros a (X & Xa & XP). destruct (XP a Xa) as (b & Rab & Xb).
210 |   exists b; split; auto. exists X; auto.
211 | Qed.
212 |
213 |
A very useful coinduction principle considers a set X where for 214 | every a in X, we can make one *or several* transitions to reach 215 | a state b that belongs to X.
216 |
217 | Lemma infseq_coinduction_principle:
218 |   forall (X: A -> Prop),
219 |   (forall a, X a -> exists b, plus a b /\ X b) ->
220 |   forall a, X a -> infseq a.
221 |
Proof.
222 |
223 |   intros X H a0 Xa0.
224 |   exists (fun a => exists b, star a b /\ X b); split.
225 | - exists a0; auto using star_refl.
226 | - intros a1 (a2 & S12 & X2). inversion S12; subst.
227 |   + destruct (H a2 X2) as (a3 & P23 & X3). inversion P23; subst.
228 |     exists b; split; auto. exists a3; auto.
229 |   + exists b; split; auto. exists a2; auto.
230 | Qed.
231 |
232 |

Determinism properties for functional transition relations.

233 |
234 |
A transition relation is functional if every state can transition 235 | to at most one other state.
236 |
237 | Hypothesis R_functional:
238 |   forall a b c, R a b -> R a c -> b = c.
239 |
240 |
Uniqueness of finite transition sequences.
241 |
242 | Lemma star_star_inv:
243 |   forall a b, star a b -> forall c, star a c -> star b c \/ star c b.
244 |
Proof.
245 |
246 |   induction 1; intros.
247 | - auto.
248 | - inversion H1; subst.
249 | + right. eauto using star.
250 | + assert (b = b0) by (eapply R_functional; eauto). subst b0.
251 |   apply IHstar; auto.
252 | Qed.
253 |
254 | Lemma finseq_unique:
255 |   forall a b b',
256 |   star a b -> irred b ->
257 |   star a b' -> irred b' ->
258 |   b = b'.
259 |
Proof.
260 |
261 |   intros. destruct (star_star_inv H H1).
262 | - inversion H3; subst. auto. elim (H0 _ H4).
263 | - inversion H3; subst. auto. elim (H2 _ H4).
264 | Qed.
265 |
266 |
A state cannot both diverge and terminate on an irreducible state.
267 |
268 | Lemma infseq_inv':
269 |   forall a b, R a b -> infseq a -> infseq b.
270 |
Proof.
271 |
272 |   intros a b Rab Ia.
273 |   destruct (infseq_inv Ia) as (b' & Rab' & Xb').
274 |   assert (b' = b) by (eapply R_functional; eauto).
275 |   subst b'. auto.
276 | Qed.
277 |
278 | Lemma infseq_star_inv:
279 |   forall a b, star a b -> infseq a -> infseq b.
280 |
Proof.
281 |
282 |   induction 1; intros.
283 | - auto.
284 | - apply IHstar. apply infseq_inv' with a; auto.
285 | Qed.
286 |
287 | Lemma infseq_finseq_excl:
288 |   forall a b,
289 |   star a b -> irred b -> infseq a -> False.
290 |
Proof.
291 |
292 |   intros.
293 |   destruct (@infseq_inv b) as (c & Rbc & _). eapply infseq_star_inv; eauto.
294 |   apply (H0 c); auto.
295 | Qed.
296 |
297 |
If there exists an infinite sequence of transitions from a, 298 | all sequences of transitions arising from a are infinite.
299 |
300 | Lemma infseq_all_seq_inf:
301 |   forall a, infseq a -> all_seq_inf a.
302 |
Proof.
303 |
304 |   intros. unfold all_seq_inf. intros.
305 |   destruct (@infseq_inv b) as (c & Rbc & _). eapply infseq_star_inv; eauto.
306 |   exists c; auto.
307 | Qed.
308 |
309 | End SEQUENCES.
310 |
311 |
312 |
313 | 314 |
315 | 316 | 317 | 318 | -------------------------------------------------------------------------------- /docs/coq2html.css: -------------------------------------------------------------------------------- 1 | 2 | /* Classes: 3 | h1.title the title of the page 4 | div.coq encloses all generated body 5 | div.doc contents of (** *) comments 6 | div.footer footer 7 | div.togglescript "Proof." line 8 | div.proofscript contents of proof script 9 | span.docright contents of (**r *) comments 10 | span.bracket contents of [ ] within comments 11 | span.comment contents of (* *) comments 12 | span.kwd Coq keyword 13 | span.tactic Coq tactic 14 | span.id any other identifier 15 | */ 16 | 17 | body { 18 | color: black; 19 | background: white; 20 | } 21 | 22 | h1.title { 23 | font-size: 2em; 24 | text-align: center 25 | } 26 | 27 | h1 { 28 | font-size: 1.5em; 29 | } 30 | h2 { 31 | font-size: 1.17em; 32 | } 33 | h3 { 34 | font-size: 1em; 35 | } 36 | 37 | h1, h2, h3 { 38 | font-family: sans-serif; 39 | margin-left: -5%; 40 | } 41 | 42 | div.coq { 43 | margin-left: 15%; 44 | margin-right: 5%; 45 | font-family: monospace, DejaVu Sans; 46 | } 47 | 48 | div.doc { 49 | margin-left: -5%; 50 | margin-top: 0.2em; 51 | margin-bottom: 0.5em; 52 | font-family: serif; 53 | } 54 | 55 | div.toggleproof { 56 | font-size: 0.8em; 57 | text-decoration: underline; 58 | } 59 | 60 | div.toggleproof:hover { 61 | cursor: pointer; 62 | } 63 | 64 | div.proofscript { 65 | font-size: 0.8em; 66 | } 67 | 68 | div.footer { 69 | margin-top: 1em; 70 | margin-bottom: 1em; 71 | font-size: 0.8em; 72 | font-style: italic; 73 | } 74 | 75 | span.docright { 76 | position: absolute; 77 | left: 50%; 78 | width: 50%; 79 | font-family: serif; 80 | } 81 | 82 | span.bracket { 83 | font-family: monospace; 84 | color: #008000; 85 | } 86 | 87 | span.kwd { 88 | color: #cf1d1d; 89 | } 90 | 91 | span.comment { 92 | color: #008000; 93 | } 94 | 95 | div.doc pre { 96 | color: Maroon; 97 | } 98 | 99 | a:visited {color : #416DFF; text-decoration : none; } 100 | a:link {color : #416DFF; text-decoration : none; } 101 | a:hover {text-decoration : none; } 102 | a:active {text-decoration : none; } 103 | -------------------------------------------------------------------------------- /docs/coq2html.js: -------------------------------------------------------------------------------- 1 | 2 | function toggleDisplay(id) 3 | { 4 | var elt = document.getElementById(id); 5 | if (elt.style.display == 'none') { 6 | elt.style.display = 'block'; 7 | } else { 8 | elt.style.display = 'none'; 9 | } 10 | } 11 | 12 | function hideAll(cls) 13 | { 14 | var testClass = new RegExp("(^|s)" + cls + "(s|$)"); 15 | var tag = tag || "*"; 16 | var elements = document.getElementsByTagName("div"); 17 | var current; 18 | var length = elements.length; 19 | for(var i=0; i