// RUN: %dafny /compile:0 /print:"%t.print" /dprint:"%t.dprint" "%s" > "%t" // RUN: %diff "%s.expect" "%t" class C { method M(x: T, u: U) returns (y: T) ensures x == y && u == u; { y := x; } function method F(x: X, u: U): bool { x == x && u == u } method Main(u: U) { var t := F(3,u) && F(this,u); var kz := M(t,u); var a := G(); assert kz && (a || !a); } method G() returns (a: Y) { } } class SetTest { method Add(s: set, x: T) returns (t: set) ensures t == s + {x}; { t := s + {x}; } method Good() { var s := {2, 5, 3}; var t := Add(s, 7); assert {5,7,2,3} == t; } method Bad() { var s := {2, 50, 3}; var t := Add(s, 7); assert {5,7,2,3} == t; // error } } class SequenceTest { method Add(s: seq, x: T) returns (t: seq) ensures t == s + [x]; { t := s + [x]; } method Good() { var s := [2, 5, 3]; var t := Add(s, 7); assert [2,5,3,7] == t; } method Bad() { var s := [2, 5, 3]; var t := Add(s, 7); assert [2,5,7,3] == t || [2,5,3] == t; // error } } // ------------------------- class CC { var x: T; method M(c: CC, z: T) returns (y: T) requires c != null; modifies this; ensures y == c.x && x == z; { x := c.x; x := z; y := c.x; } } class CClient { method Main() { var c := new CC; var k := c.x + 3; if (c.x == c.x) { k := k + 1; } var m := c.x; if (m == c.x) { k := k + 1; } c.x := 5; c.x := c.x; var z := c.M(c, 17); assert z == c.x; } } // ------------------------- function IsCelebrity(c: Person, people: set): bool requires c == c || c in people; { false } method FindCelebrity3(people: set, ghost c: int) requires IsCelebrity(c, people); // once upon a time, this caused the translator to produce bad Boogie { ghost var b: bool; b := IsCelebrity(c, people); b := F(c, people); } function F(c: int, people: set): bool requires IsCelebrity(c, people); { false } function RogerThat(g: G): G { g } function Cool(b: bool): bool { b } function Rockin'(g: G): G { var h := g; h } function Groovy(g: G, x: int): G { var h := g; if x == 80 then RogerThat(h) else [h][0] } method IsRogerCool(n: int) requires RogerThat(true); // once upon a time, this caused the translator to produce bad Boogie { if (*) { assert Cool(2 < 3 && n < n && n < n+1); // the error message here will peek into the argument of Cool } else if (*) { assert RogerThat(2 < 3 && n < n && n < n+1); // same here; cool, huh? } else if (*) { assert Rockin'(false); // error } else if (*) { assert Groovy(n < n, 80); // error } else if (*) { assert Groovy(n + 1 <= n, 81); // error } } method LoopyRoger(n: int) { var i := 0; while (i < n) invariant RogerThat(0 <= n ==> i <= n); { i := i + 1; } i := 0; while (i < n) invariant RogerThat(0 <= n ==> i <= n); // error: failure to maintain loop invariant { i := i + 2; } } // ---------------------- class TyKn_C { var x: T; function G(): T reads this; { x } method M() returns (t: T) { } } class TyKn_K { function F(): int { 176 } } method TyKn_Main(k0: TyKn_K) { var c := new TyKn_C; var k1: TyKn_K; assert k0 != null ==> k0.F() == 176; assert k1 != null ==> k1.F() == 176; assert c.x != null ==> c.x.F() == 176; // the Dafny encoding needs the canCall mechanism to verify this assert c.G() != null ==> c.G().F() == 176; // ditto var k2 := c.M(); assert k2 != null ==> k2.F() == 176; // the canCall mechanism does the trick here, but so does the encoding // via k2's where clause } // ------------------- there was once a bug in the handling of the following example module OneLayer { datatype wrap = Wrap(V) } module TwoLayers { import OneLayer datatype wrap2 = Wrap2(get: OneLayer.wrap) function F(w: wrap2) : OneLayer.wrap { match w case Wrap2(a) => a } function G(w: wrap2) : OneLayer.wrap { match w case Wrap2(a) => w.get } function H(w: wrap2) : OneLayer.wrap { w.get } } // --------------------------------------------------------------------- datatype List = Nil | Cons(T, List) predicate InList(x: T, xs: List) predicate Subset(xs: List, ys: List) { forall x :: InList(x, xs) ==> InList(x, ys) } ghost method ListLemma_T(xs: List, ys: List) requires forall x :: InList(x, xs) ==> InList(x, ys); { assert Subset(xs, ys); } ghost method ammeLtsiL_T(xs: List, ys: List) requires Subset(xs, ys); { assert forall x :: InList(x, xs) ==> InList(x, ys); } ghost method ListLemma_int(xs: List, ys: List) requires forall x :: InList(x, xs) ==> InList(x, ys); { assert Subset(xs, ys); } ghost method ammeLtsiL_int(xs: List, ys: List) requires Subset(xs, ys); { assert forall x :: InList(x, xs) ==> InList(x, ys); } // -------------- auto filled-in type arguments for collection types ------ function length(xs: List): nat { match xs case Nil => 0 case Cons(_, tail) => 1 + length(tail) } function elems(xs: List): set { match xs case Nil => {} case Cons(x, tail) => {x} + elems(tail) } function Card(s: set): nat { |s| } function Identity(s: set): set { s } function MultisetToSet(m: multiset): set { if |m| == 0 then {} else var x :| x in m; MultisetToSet(m - multiset{x}) + {x} } function SeqToSet(q: seq): set { if q == [] then {} else {q[0]} + SeqToSet(q[1..]) } datatype Pair = MkPair(0: T, 1: U) method IdentityMap(s: set) returns (m: map) { m := map[]; var s := s; while s != {} decreases s; { var p :| p in s; m, s := m[p.0 := p.1], s - {p}; } } // -------------- auto filled-in type arguments for array types ------ module ArrayTypeMagic { method M(a: array2) { } method F(b: array) returns (s: seq) requires b != null; { return b[..]; } datatype ArrayCubeTree = Leaf(array3) | Node(ArrayCubeTree, ArrayCubeTree) datatype AnotherACT = Leaf(array3) | Node(AnotherACT, AnotherACT) datatype OneMoreACT = Leaf(array3) | Node(OneMoreACT, OneMoreACT) function G(t: ArrayCubeTree): array3 { match t case Leaf(d) => d case Node(l, _) => G(l) } datatype Nat = Zero | Succ(Nat) datatype List = Nil | Cons(T, List) datatype Cell = Mk(T) datatype DList = Nil(Cell) | Cons(T, U, DList) }